Documentation

Mathlib.GroupTheory.GroupAction.Prod

Prod instances for additive and multiplicative actions #

This file defines instances for binary product of additive and multiplicative actions and provides scalar multiplication as a homomorphism from α × β to β.

Main declarations #

See also #

Porting notes #

The to_additive attribute can be used to generate both the smul and vadd lemmas from the corresponding pow lemmas, as explained on zulip here: https://leanprover.zulipchat.com/#narrow/near/316087838

This was not done as part of the port in order to stay as close as possible to the mathlib3 code.

instance Prod.vadd {M : Type u_1} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd M β] :
VAdd M (α × β)
Equations
  • Prod.vadd = { vadd := fun a p => (a +ᵥ p.fst, a +ᵥ p.snd) }
instance Prod.smul {M : Type u_1} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul M β] :
SMul M (α × β)
Equations
  • Prod.smul = { smul := fun a p => (a p.fst, a p.snd) }
@[simp]
theorem Prod.vadd_fst {M : Type u_1} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd M β] (a : M) (x : α × β) :
(a +ᵥ x).fst = a +ᵥ x.fst
@[simp]
theorem Prod.smul_fst {M : Type u_1} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul M β] (a : M) (x : α × β) :
(a x).fst = a x.fst
@[simp]
theorem Prod.vadd_snd {M : Type u_1} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd M β] (a : M) (x : α × β) :
(a +ᵥ x).snd = a +ᵥ x.snd
@[simp]
theorem Prod.smul_snd {M : Type u_1} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul M β] (a : M) (x : α × β) :
(a x).snd = a x.snd
@[simp]
theorem Prod.vadd_mk {M : Type u_1} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd M β] (a : M) (b : α) (c : β) :
a +ᵥ (b, c) = (a +ᵥ b, a +ᵥ c)
@[simp]
theorem Prod.smul_mk {M : Type u_1} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul M β] (a : M) (b : α) (c : β) :
a (b, c) = (a b, a c)
theorem Prod.vadd_def {M : Type u_1} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd M β] (a : M) (x : α × β) :
a +ᵥ x = (a +ᵥ x.fst, a +ᵥ x.snd)
theorem Prod.smul_def {M : Type u_1} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul M β] (a : M) (x : α × β) :
a x = (a x.fst, a x.snd)
@[simp]
theorem Prod.vadd_swap {M : Type u_1} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd M β] (a : M) (x : α × β) :
@[simp]
theorem Prod.smul_swap {M : Type u_1} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul M β] (a : M) (x : α × β) :
theorem Prod.smul_zero_mk {M : Type u_1} {β : Type u_6} [SMul M β] {α : Type u_7} [Monoid M] [AddMonoid α] [DistribMulAction M α] (a : M) (c : β) :
a (0, c) = (0, a c)
theorem Prod.smul_mk_zero {M : Type u_1} {α : Type u_5} [SMul M α] {β : Type u_7} [Monoid M] [AddMonoid β] [DistribMulAction M β] (a : M) (b : α) :
a (b, 0) = (a b, 0)
instance Prod.pow {E : Type u_4} {α : Type u_5} {β : Type u_6} [Pow α E] [Pow β E] :
Pow (α × β) E
Equations
  • Prod.pow = { pow := fun p c => (p.fst ^ c, p.snd ^ c) }
@[simp]
theorem Prod.pow_fst {E : Type u_4} {α : Type u_5} {β : Type u_6} [Pow α E] [Pow β E] (p : α × β) (c : E) :
(p ^ c).fst = p.fst ^ c
@[simp]
theorem Prod.pow_snd {E : Type u_4} {α : Type u_5} {β : Type u_6} [Pow α E] [Pow β E] (p : α × β) (c : E) :
(p ^ c).snd = p.snd ^ c
@[simp]
theorem Prod.pow_mk {E : Type u_4} {α : Type u_5} {β : Type u_6} [Pow α E] [Pow β E] (c : E) (a : α) (b : β) :
(a, b) ^ c = (a ^ c, b ^ c)
theorem Prod.pow_def {E : Type u_4} {α : Type u_5} {β : Type u_6} [Pow α E] [Pow β E] (p : α × β) (c : E) :
p ^ c = (p.fst ^ c, p.snd ^ c)
@[simp]
theorem Prod.pow_swap {E : Type u_4} {α : Type u_5} {β : Type u_6} [Pow α E] [Pow β E] (p : α × β) (c : E) :
Prod.swap (p ^ c) = Prod.swap p ^ c
instance Prod.vaddAssocClass {M : Type u_1} {N : Type u_2} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd M β] [VAdd N α] [VAdd N β] [VAdd M N] [VAddAssocClass M N α] [VAddAssocClass M N β] :
VAddAssocClass M N (α × β)
Equations
theorem Prod.vaddAssocClass.proof_1 {M : Type u_1} {N : Type u_2} {α : Type u_3} {β : Type u_4} [VAdd M α] [VAdd M β] [VAdd N α] [VAdd N β] [VAdd M N] [VAddAssocClass M N α] [VAddAssocClass M N β] :
VAddAssocClass M N (α × β)
instance Prod.isScalarTower {M : Type u_1} {N : Type u_2} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul M β] [SMul N α] [SMul N β] [SMul M N] [IsScalarTower M N α] [IsScalarTower M N β] :
IsScalarTower M N (α × β)
Equations
instance Prod.vaddCommClass {M : Type u_1} {N : Type u_2} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd M β] [VAdd N α] [VAdd N β] [VAddCommClass M N α] [VAddCommClass M N β] :
VAddCommClass M N (α × β)
Equations
theorem Prod.vaddCommClass.proof_1 {M : Type u_1} {N : Type u_2} {α : Type u_3} {β : Type u_4} [VAdd M α] [VAdd M β] [VAdd N α] [VAdd N β] [VAddCommClass M N α] [VAddCommClass M N β] :
VAddCommClass M N (α × β)
instance Prod.smulCommClass {M : Type u_1} {N : Type u_2} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul M β] [SMul N α] [SMul N β] [SMulCommClass M N α] [SMulCommClass M N β] :
SMulCommClass M N (α × β)
Equations
theorem Prod.isCentralVAdd.proof_1 {M : Type u_1} {α : Type u_2} {β : Type u_3} [VAdd M α] [VAdd M β] [VAdd Mᵃᵒᵖ α] [VAdd Mᵃᵒᵖ β] [IsCentralVAdd M α] [IsCentralVAdd M β] :
IsCentralVAdd M (α × β)
instance Prod.isCentralVAdd {M : Type u_1} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd M β] [VAdd Mᵃᵒᵖ α] [VAdd Mᵃᵒᵖ β] [IsCentralVAdd M α] [IsCentralVAdd M β] :
IsCentralVAdd M (α × β)
Equations
instance Prod.isCentralScalar {M : Type u_1} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul M β] [SMul Mᵐᵒᵖ α] [SMul Mᵐᵒᵖ β] [IsCentralScalar M α] [IsCentralScalar M β] :
Equations
instance Prod.faithfulVAddLeft {M : Type u_1} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd M β] [FaithfulVAdd M α] [Nonempty β] :
FaithfulVAdd M (α × β)
Equations
theorem Prod.faithfulVAddLeft.proof_1 {M : Type u_1} {α : Type u_2} {β : Type u_3} [VAdd M α] [VAdd M β] [FaithfulVAdd M α] [Nonempty β] :
FaithfulVAdd M (α × β)
abbrev Prod.faithfulVAddLeft.match_1 {β : Type u_1} (motive : Nonempty βProp) :
(x : Nonempty β) → ((b : β) → motive (_ : Nonempty β)) → motive x
Equations
Instances For
    instance Prod.faithfulSMulLeft {M : Type u_1} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul M β] [FaithfulSMul M α] [Nonempty β] :
    FaithfulSMul M (α × β)
    Equations
    theorem Prod.faithfulVAddRight.proof_1 {M : Type u_1} {α : Type u_2} {β : Type u_3} [VAdd M α] [VAdd M β] [Nonempty α] [FaithfulVAdd M β] :
    FaithfulVAdd M (α × β)
    instance Prod.faithfulVAddRight {M : Type u_1} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd M β] [Nonempty α] [FaithfulVAdd M β] :
    FaithfulVAdd M (α × β)
    Equations
    instance Prod.faithfulSMulRight {M : Type u_1} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul M β] [Nonempty α] [FaithfulSMul M β] :
    FaithfulSMul M (α × β)
    Equations
    theorem Prod.vaddCommClassBoth.proof_1 {M : Type u_1} {N : Type u_2} {P : Type u_3} [Add N] [Add P] [VAdd M N] [VAdd M P] [VAddCommClass M N N] [VAddCommClass M P P] :
    VAddCommClass M (N × P) (N × P)
    instance Prod.vaddCommClassBoth {M : Type u_1} {N : Type u_2} {P : Type u_3} [Add N] [Add P] [VAdd M N] [VAdd M P] [VAddCommClass M N N] [VAddCommClass M P P] :
    VAddCommClass M (N × P) (N × P)
    Equations
    instance Prod.smulCommClassBoth {M : Type u_1} {N : Type u_2} {P : Type u_3} [Mul N] [Mul P] [SMul M N] [SMul M P] [SMulCommClass M N N] [SMulCommClass M P P] :
    SMulCommClass M (N × P) (N × P)
    Equations
    instance Prod.isScalarTowerBoth {M : Type u_1} {N : Type u_2} {P : Type u_3} [Mul N] [Mul P] [SMul M N] [SMul M P] [IsScalarTower M N N] [IsScalarTower M P P] :
    IsScalarTower M (N × P) (N × P)
    Equations
    instance Prod.addAction {M : Type u_1} {α : Type u_5} {β : Type u_6} [AddMonoid M] [AddAction M α] [AddAction M β] :
    AddAction M (α × β)
    Equations
    • One or more equations did not get rendered due to their size.
    theorem Prod.addAction.proof_2 {M : Type u_3} {α : Type u_1} {β : Type u_2} [AddMonoid M] [AddAction M α] [AddAction M β] :
    ∀ (x x_1 : M) (x_2 : α × β), (x + x_1 +ᵥ x_2.fst, x + x_1 +ᵥ x_2.snd) = (x +ᵥ (x_1 +ᵥ x_2).fst, x +ᵥ (x_1 +ᵥ x_2).snd)
    abbrev Prod.addAction.match_1 {α : Type u_1} {β : Type u_2} (motive : α × βProp) :
    (x : α × β) → ((fst : α) → (snd : β) → motive (fst, snd)) → motive x
    Equations
    Instances For
      theorem Prod.addAction.proof_1 {M : Type u_3} {α : Type u_1} {β : Type u_2} [AddMonoid M] [AddAction M α] [AddAction M β] :
      ∀ (x : α × β), 0 +ᵥ x = x
      instance Prod.mulAction {M : Type u_1} {α : Type u_5} {β : Type u_6} [Monoid M] [MulAction M α] [MulAction M β] :
      MulAction M (α × β)
      Equations
      • One or more equations did not get rendered due to their size.
      instance Prod.smulZeroClass {R : Type u_7} {M : Type u_8} {N : Type u_9} [Zero M] [Zero N] [SMulZeroClass R M] [SMulZeroClass R N] :
      Equations
      instance Prod.distribSMul {R : Type u_7} {M : Type u_8} {N : Type u_9} [AddZeroClass M] [AddZeroClass N] [DistribSMul R M] [DistribSMul R N] :
      DistribSMul R (M × N)
      Equations
      instance Prod.distribMulAction {M : Type u_1} {N : Type u_2} {R : Type u_7} [Monoid R] [AddMonoid M] [AddMonoid N] [DistribMulAction R M] [DistribMulAction R N] :
      Equations
      • Prod.distribMulAction = let src := Prod.mulAction; let src_1 := Prod.distribSMul; DistribMulAction.mk (_ : ∀ (a : R), a 0 = 0) (_ : ∀ (a : R) (x y : M × N), a (x + y) = a x + a y)
      instance Prod.mulDistribMulAction {M : Type u_1} {N : Type u_2} {R : Type u_7} [Monoid R] [Monoid M] [Monoid N] [MulDistribMulAction R M] [MulDistribMulAction R N] :
      Equations
      • One or more equations did not get rendered due to their size.

      Scalar multiplication as a homomorphism #

      @[simp]
      theorem smulMulHom_apply {α : Type u_5} {β : Type u_6} [Monoid α] [Mul β] [MulAction α β] [IsScalarTower α β β] [SMulCommClass α β β] (a : α × β) :
      smulMulHom a = a.fst a.snd
      def smulMulHom {α : Type u_5} {β : Type u_6} [Monoid α] [Mul β] [MulAction α β] [IsScalarTower α β β] [SMulCommClass α β β] :
      α × β →ₙ* β

      Scalar multiplication as a multiplicative homomorphism.

      Equations
      • smulMulHom = { toFun := fun a => a.fst a.snd, map_mul' := (_ : ∀ (x x_1 : α × β), (x.fst * x_1.fst) (x.snd * x_1.snd) = x.fst x.snd * x_1.fst x_1.snd) }
      Instances For
        @[simp]
        theorem smulMonoidHom_apply {α : Type u_5} {β : Type u_6} [Monoid α] [MulOneClass β] [MulAction α β] [IsScalarTower α β β] [SMulCommClass α β β] :
        ∀ (a : α × β), smulMonoidHom a = MulHom.toFun smulMulHom a
        def smulMonoidHom {α : Type u_5} {β : Type u_6} [Monoid α] [MulOneClass β] [MulAction α β] [IsScalarTower α β β] [SMulCommClass α β β] :
        α × β →* β

        Scalar multiplication as a monoid homomorphism.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For