Documentation

Mathlib.Algebra.Group.Action.Defs

Definitions of group actions #

This file defines a hierarchy of group action type-classes on top of the previously defined notation classes SMul and its additive version VAdd:

The hierarchy is extended further by Module, defined elsewhere.

Also provided are typeclasses regarding the interaction of different group actions,

Notation #

Implementation details #

This file should avoid depending on other parts of GroupTheory, to avoid import cycles. More sophisticated lemmas belong in GroupTheory.GroupAction.

Tags #

group action

@[instance 910]
instance Mul.toSMul (α : Type u_9) [Mul α] :
SMul α α

See also Monoid.toMulAction and MulZeroClass.toSMulWithZero.

Equations
    @[instance 910]
    instance Add.toVAdd (α : Type u_9) [Add α] :
    VAdd α α

    See also AddMonoid.toAddAction

    Equations
      @[instance 910]
      instance Mul.toSMulMulOpposite (α : Type u_9) [Mul α] :

      Like Mul.toSMul, but multiplies on the right.

      See also Monoid.toOppositeMulAction and MonoidWithZero.toOppositeMulActionWithZero.

      Equations
        @[instance 910]
        instance Add.toVAddAddOpposite (α : Type u_9) [Add α] :

        Like Add.toVAdd, but adds on the right.

        See also AddMonoid.toOppositeAddAction.

        Equations
          @[simp]
          theorem smul_eq_mul {α : Type u_9} [Mul α] (a b : α) :
          a b = a * b
          @[simp]
          theorem vadd_eq_add {α : Type u_9} [Add α] (a b : α) :
          a +ᵥ b = a + b
          theorem op_smul_eq_mul {α : Type u_9} [Mul α] (a b : α) :
          theorem op_vadd_eq_add {α : Type u_9} [Add α] (a b : α) :
          @[simp]
          theorem MulOpposite.smul_eq_mul_unop {α : Type u_5} [Mul α] (a : αᵐᵒᵖ) (b : α) :
          a b = b * unop a
          @[simp]
          theorem AddOpposite.vadd_eq_add_unop {α : Type u_5} [Add α] (a : αᵃᵒᵖ) (b : α) :
          a +ᵥ b = b + unop a
          class AddAction (G : Type u_9) (P : Type u_10) [AddMonoid G] extends VAdd G P :
          Type (max u_10 u_9)

          Type class for additive monoid actions on types, with notation g +ᵥ p.

          The AddAction G P typeclass says that the additive monoid G acts additively on a type P. More precisely this means that the action satisfies the two axioms 0 +ᵥ p = p and (g₁ + g₂) +ᵥ p = g₁ +ᵥ (g₂ +ᵥ p). A mathematician might simply say that the additive monoid G acts on P.

          For example, if A is an additive group and X is a type, if a mathematician says say "let A act on the set X" they will usually mean [AddAction A X].

          • vadd : GPP
          • zero_vadd (p : P) : 0 +ᵥ p = p

            Zero is a neutral element for +ᵥ

          • add_vadd (g₁ g₂ : G) (p : P) : (g₁ + g₂) +ᵥ p = g₁ +ᵥ g₂ +ᵥ p

            Associativity of + and +ᵥ

          Instances
            class MulAction (α : Type u_9) (β : Type u_10) [Monoid α] extends SMul α β :
            Type (max u_10 u_9)

            Type class for monoid actions on types, with notation g • p.

            The MulAction G P typeclass says that the monoid G acts multiplicatively on a type P. More precisely this means that the action satisfies the two axioms 1 • p = p and (g₁ * g₂) • p = g₁ • (g₂ • p). A mathematician might simply say that the monoid G acts on P.

            For example, if G is a group and X is a type, if a mathematician says say "let G act on the set X" they will probably mean [AddAction G X].

            • smul : αββ
            • one_smul (b : β) : 1 b = b

              One is the neutral element for

            • mul_smul (x y : α) (b : β) : (x * y) b = x y b

              Associativity of and *

            Instances
              theorem MulAction.ext_iff {α : Type u_9} {β : Type u_10} {inst✝ : Monoid α} {x y : MulAction α β} :
              theorem AddAction.ext_iff {G : Type u_9} {P : Type u_10} {inst✝ : AddMonoid G} {x y : AddAction G P} :
              theorem AddAction.ext {G : Type u_9} {P : Type u_10} {inst✝ : AddMonoid G} {x y : AddAction G P} (vadd : VAdd.vadd = VAdd.vadd) :
              x = y
              theorem MulAction.ext {α : Type u_9} {β : Type u_10} {inst✝ : Monoid α} {x y : MulAction α β} (smul : SMul.smul = SMul.smul) :
              x = y

              Scalar tower and commuting actions #

              class VAddCommClass (M : Type u_9) (N : Type u_10) (α : Type u_11) [VAdd M α] [VAdd N α] :

              A typeclass mixin saying that two additive actions on the same space commute.

              • vadd_comm (m : M) (n : N) (a : α) : m +ᵥ n +ᵥ a = n +ᵥ m +ᵥ a

                +ᵥ is left commutative

              Instances
                class SMulCommClass (M : Type u_9) (N : Type u_10) (α : Type u_11) [SMul M α] [SMul N α] :

                A typeclass mixin saying that two multiplicative actions on the same space commute.

                • smul_comm (m : M) (n : N) (a : α) : m n a = n m a

                  is left commutative

                Instances
                  theorem SMulCommClass.symm (M : Type u_9) (N : Type u_10) (α : Type u_11) [SMul M α] [SMul N α] [SMulCommClass M N α] :

                  Commutativity of actions is a symmetric relation. This lemma can't be an instance because this would cause a loop in the instance search graph.

                  theorem VAddCommClass.symm (M : Type u_9) (N : Type u_10) (α : Type u_11) [VAdd M α] [VAdd N α] [VAddCommClass M N α] :

                  Commutativity of additive actions is a symmetric relation. This lemma can't be an instance because this would cause a loop in the instance search graph.

                  theorem Function.Injective.smulCommClass {M : Type u_1} {N : Type u_2} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul N α] [SMul M β] [SMul N β] [SMulCommClass M N β] {f : αβ} (hf : Injective f) (h₁ : ∀ (c : M) (x : α), f (c x) = c f x) (h₂ : ∀ (c : N) (x : α), f (c x) = c f x) :
                  theorem Function.Injective.vaddCommClass {M : Type u_1} {N : Type u_2} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd N α] [VAdd M β] [VAdd N β] [VAddCommClass M N β] {f : αβ} (hf : Injective f) (h₁ : ∀ (c : M) (x : α), f (c +ᵥ x) = c +ᵥ f x) (h₂ : ∀ (c : N) (x : α), f (c +ᵥ x) = c +ᵥ f x) :
                  theorem Function.Surjective.smulCommClass {M : Type u_1} {N : Type u_2} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul N α] [SMul M β] [SMul N β] [SMulCommClass M N α] {f : αβ} (hf : Surjective f) (h₁ : ∀ (c : M) (x : α), f (c x) = c f x) (h₂ : ∀ (c : N) (x : α), f (c x) = c f x) :
                  theorem Function.Surjective.vaddCommClass {M : Type u_1} {N : Type u_2} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd N α] [VAdd M β] [VAdd N β] [VAddCommClass M N α] {f : αβ} (hf : Surjective f) (h₁ : ∀ (c : M) (x : α), f (c +ᵥ x) = c +ᵥ f x) (h₂ : ∀ (c : N) (x : α), f (c +ᵥ x) = c +ᵥ f x) :
                  instance smulCommClass_self (M : Type u_9) (α : Type u_10) [CommMonoid M] [MulAction M α] :
                  instance vaddCommClass_self (M : Type u_9) (α : Type u_10) [AddCommMonoid M] [AddAction M α] :
                  class VAddAssocClass (M : Type u_9) (N : Type u_10) (α : Type u_11) [VAdd M N] [VAdd N α] [VAdd M α] :

                  An instance of VAddAssocClass M N α states that the additive action of M on α is determined by the additive actions of M on N and N on α.

                  • vadd_assoc (x : M) (y : N) (z : α) : (x +ᵥ y) +ᵥ z = x +ᵥ y +ᵥ z

                    Associativity of +ᵥ

                  Instances
                    class IsScalarTower (M : Type u_9) (N : Type u_10) (α : Type u_11) [SMul M N] [SMul N α] [SMul M α] :

                    An instance of IsScalarTower M N α states that the multiplicative action of M on α is determined by the multiplicative actions of M on N and N on α.

                    • smul_assoc (x : M) (y : N) (z : α) : (x y) z = x y z

                      Associativity of

                    Instances
                      @[simp]
                      theorem smul_assoc {α : Type u_5} {M : Type u_9} {N : Type u_10} [SMul M N] [SMul N α] [SMul M α] [IsScalarTower M N α] (x : M) (y : N) (z : α) :
                      (x y) z = x y z
                      @[simp]
                      theorem vadd_assoc {α : Type u_5} {M : Type u_9} {N : Type u_10} [VAdd M N] [VAdd N α] [VAdd M α] [VAddAssocClass M N α] (x : M) (y : N) (z : α) :
                      (x +ᵥ y) +ᵥ z = x +ᵥ y +ᵥ z
                      instance Semigroup.isScalarTower {α : Type u_5} [Semigroup α] :
                      IsScalarTower α α α
                      class IsCentralVAdd (M : Type u_9) (α : Type u_10) [VAdd M α] [VAdd Mᵃᵒᵖ α] :

                      A typeclass indicating that the right (aka AddOpposite) and left actions by M on α are equal, that is that M acts centrally on α. This can be thought of as a version of commutativity for +ᵥ.

                      • op_vadd_eq_vadd (m : M) (a : α) : AddOpposite.op m +ᵥ a = m +ᵥ a

                        The right and left actions of M on α are equal.

                      Instances
                        class IsCentralScalar (M : Type u_9) (α : Type u_10) [SMul M α] [SMul Mᵐᵒᵖ α] :

                        A typeclass indicating that the right (aka MulOpposite) and left actions by M on α are equal, that is that M acts centrally on α. This can be thought of as a version of commutativity for .

                        • op_smul_eq_smul (m : M) (a : α) : MulOpposite.op m a = m a

                          The right and left actions of M on α are equal.

                        Instances
                          theorem IsCentralScalar.unop_smul_eq_smul {M : Type u_9} {α : Type u_10} [SMul M α] [SMul Mᵐᵒᵖ α] [IsCentralScalar M α] (m : Mᵐᵒᵖ) (a : α) :
                          theorem IsCentralVAdd.unop_vadd_eq_vadd {M : Type u_9} {α : Type u_10} [VAdd M α] [VAdd Mᵃᵒᵖ α] [IsCentralVAdd M α] (m : Mᵃᵒᵖ) (a : α) :
                          @[instance 50]
                          instance SMulCommClass.op_left {M : Type u_1} {N : Type u_2} {α : Type u_5} [SMul M α] [SMul Mᵐᵒᵖ α] [IsCentralScalar M α] [SMul N α] [SMulCommClass M N α] :
                          @[instance 50]
                          instance VAddCommClass.op_left {M : Type u_1} {N : Type u_2} {α : Type u_5} [VAdd M α] [VAdd Mᵃᵒᵖ α] [IsCentralVAdd M α] [VAdd N α] [VAddCommClass M N α] :
                          @[instance 50]
                          instance SMulCommClass.op_right {M : Type u_1} {N : Type u_2} {α : Type u_5} [SMul M α] [SMul N α] [SMul Nᵐᵒᵖ α] [IsCentralScalar N α] [SMulCommClass M N α] :
                          @[instance 50]
                          instance VAddCommClass.op_right {M : Type u_1} {N : Type u_2} {α : Type u_5} [VAdd M α] [VAdd N α] [VAdd Nᵃᵒᵖ α] [IsCentralVAdd N α] [VAddCommClass M N α] :
                          @[instance 50]
                          instance IsScalarTower.op_left {M : Type u_1} {N : Type u_2} {α : Type u_5} [SMul M α] [SMul Mᵐᵒᵖ α] [IsCentralScalar M α] [SMul M N] [SMul Mᵐᵒᵖ N] [IsCentralScalar M N] [SMul N α] [IsScalarTower M N α] :
                          @[instance 50]
                          instance VAddAssocClass.op_left {M : Type u_1} {N : Type u_2} {α : Type u_5} [VAdd M α] [VAdd Mᵃᵒᵖ α] [IsCentralVAdd M α] [VAdd M N] [VAdd Mᵃᵒᵖ N] [IsCentralVAdd M N] [VAdd N α] [VAddAssocClass M N α] :
                          @[instance 50]
                          instance IsScalarTower.op_right {M : Type u_1} {N : Type u_2} {α : Type u_5} [SMul M α] [SMul M N] [SMul N α] [SMul Nᵐᵒᵖ α] [IsCentralScalar N α] [IsScalarTower M N α] :
                          @[instance 50]
                          instance VAddAssocClass.op_right {M : Type u_1} {N : Type u_2} {α : Type u_5} [VAdd M α] [VAdd M N] [VAdd N α] [VAdd Nᵃᵒᵖ α] [IsCentralVAdd N α] [VAddAssocClass M N α] :
                          def SMul.comp.smul {M : Type u_1} {N : Type u_2} {α : Type u_5} [SMul M α] (g : NM) (n : N) (a : α) :
                          α

                          Auxiliary definition for SMul.comp, MulAction.compHom, DistribMulAction.compHom, Module.compHom, etc.

                          Equations
                            Instances For
                              def VAdd.comp.vadd {M : Type u_1} {N : Type u_2} {α : Type u_5} [VAdd M α] (g : NM) (n : N) (a : α) :
                              α

                              Auxiliary definition for VAdd.comp, AddAction.compHom, etc.

                              Equations
                                Instances For
                                  @[reducible, inline]
                                  abbrev SMul.comp {M : Type u_1} {N : Type u_2} (α : Type u_5) [SMul M α] (g : NM) :
                                  SMul N α

                                  An action of M on α and a function N → M induces an action of N on α.

                                  Equations
                                    Instances For
                                      @[reducible, inline]
                                      abbrev VAdd.comp {M : Type u_1} {N : Type u_2} (α : Type u_5) [VAdd M α] (g : NM) :
                                      VAdd N α

                                      An additive action of M on α and a function N → M induces an additive action of N on α.

                                      Equations
                                        Instances For
                                          theorem SMul.comp.isScalarTower {M : Type u_1} {N : Type u_2} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul M β] [SMul α β] [IsScalarTower M α β] (g : NM) :

                                          Given a tower of scalar actions M → α → β, if we use SMul.comp to pull back both of M's actions by a map g : N → M, then we obtain a new tower of scalar actions N → α → β.

                                          This cannot be an instance because it can cause infinite loops whenever the SMul arguments are still metavariables.

                                          theorem VAdd.comp.vaddAssocClass {M : Type u_1} {N : Type u_2} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd M β] [VAdd α β] [VAddAssocClass M α β] (g : NM) :

                                          Given a tower of additive actions M → α → β, if we use SMul.comp to pull back both of M's actions by a map g : N → M, then we obtain a new tower of scalar actions N → α → β.

                                          This cannot be an instance because it can cause infinite loops whenever the SMul arguments are still metavariables.

                                          theorem SMul.comp.smulCommClass {M : Type u_1} {N : Type u_2} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul β α] [SMulCommClass M β α] (g : NM) :

                                          This cannot be an instance because it can cause infinite loops whenever the SMul arguments are still metavariables.

                                          theorem VAdd.comp.vaddCommClass {M : Type u_1} {N : Type u_2} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd β α] [VAddCommClass M β α] (g : NM) :

                                          This cannot be an instance because it can cause infinite loops whenever the VAdd arguments are still metavariables.

                                          theorem SMul.comp.smulCommClass' {M : Type u_1} {N : Type u_2} {α : Type u_5} {β : Type u_6} [SMul M α] [SMul β α] [SMulCommClass β M α] (g : NM) :

                                          This cannot be an instance because it can cause infinite loops whenever the SMul arguments are still metavariables.

                                          theorem VAdd.comp.vaddCommClass' {M : Type u_1} {N : Type u_2} {α : Type u_5} {β : Type u_6} [VAdd M α] [VAdd β α] [VAddCommClass β M α] (g : NM) :

                                          This cannot be an instance because it can cause infinite loops whenever the VAdd arguments are still metavariables.

                                          theorem mul_smul_comm {α : Type u_5} {β : Type u_6} [Mul β] [SMul α β] [SMulCommClass α β β] (s : α) (x y : β) :
                                          x * s y = s (x * y)

                                          Note that the SMulCommClass α β β typeclass argument is usually satisfied by Algebra α β.

                                          theorem add_vadd_comm {α : Type u_5} {β : Type u_6} [Add β] [VAdd α β] [VAddCommClass α β β] (s : α) (x y : β) :
                                          x + (s +ᵥ y) = s +ᵥ x + y
                                          theorem smul_mul_assoc {α : Type u_5} {β : Type u_6} [Mul β] [SMul α β] [IsScalarTower α β β] (r : α) (x y : β) :
                                          r x * y = r (x * y)

                                          Note that the IsScalarTower α β β typeclass argument is usually satisfied by Algebra α β.

                                          theorem vadd_add_assoc {α : Type u_5} {β : Type u_6} [Add β] [VAdd α β] [VAddAssocClass α β β] (r : α) (x y : β) :
                                          (r +ᵥ x) + y = r +ᵥ x + y
                                          theorem smul_div_assoc {α : Type u_5} {β : Type u_6} [DivInvMonoid β] [SMul α β] [IsScalarTower α β β] (r : α) (x y : β) :
                                          r x / y = r (x / y)

                                          Note that the IsScalarTower α β β typeclass argument is usually satisfied by Algebra α β.

                                          theorem vadd_sub_assoc {α : Type u_5} {β : Type u_6} [SubNegMonoid β] [VAdd α β] [VAddAssocClass α β β] (r : α) (x y : β) :
                                          (r +ᵥ x) - y = r +ᵥ x - y
                                          theorem smul_smul_smul_comm {α : Type u_5} {β : Type u_6} {γ : Type u_7} {δ : Type u_8} [SMul α β] [SMul α γ] [SMul β δ] [SMul α δ] [SMul γ δ] [IsScalarTower α β δ] [IsScalarTower α γ δ] [SMulCommClass β γ δ] (a : α) (b : β) (c : γ) (d : δ) :
                                          (a b) c d = (a c) b d
                                          theorem vadd_vadd_vadd_comm {α : Type u_5} {β : Type u_6} {γ : Type u_7} {δ : Type u_8} [VAdd α β] [VAdd α γ] [VAdd β δ] [VAdd α δ] [VAdd γ δ] [VAddAssocClass α β δ] [VAddAssocClass α γ δ] [VAddCommClass β γ δ] (a : α) (b : β) (c : γ) (d : δ) :
                                          (a +ᵥ b) +ᵥ c +ᵥ d = (a +ᵥ c) +ᵥ b +ᵥ d
                                          theorem smul_mul_smul_comm {α : Type u_5} {β : Type u_6} [Mul α] [Mul β] [SMul α β] [IsScalarTower α β β] [IsScalarTower α α β] [SMulCommClass α β β] (a : α) (b : β) (c : α) (d : β) :
                                          a b * c d = (a * c) (b * d)

                                          Note that the IsScalarTower α β β and SMulCommClass α β β typeclass arguments are usually satisfied by Algebra α β.

                                          theorem vadd_add_vadd_comm {α : Type u_5} {β : Type u_6} [Add α] [Add β] [VAdd α β] [VAddAssocClass α β β] [VAddAssocClass α α β] [VAddCommClass α β β] (a : α) (b : β) (c : α) (d : β) :
                                          (a +ᵥ b) + (c +ᵥ d) = (a + c) +ᵥ b + d
                                          theorem smul_mul_smul {α : Type u_5} {β : Type u_6} [Mul α] [Mul β] [SMul α β] [IsScalarTower α β β] [IsScalarTower α α β] [SMulCommClass α β β] (a : α) (b : β) (c : α) (d : β) :
                                          a b * c d = (a * c) (b * d)

                                          Alias of smul_mul_smul_comm.


                                          Note that the IsScalarTower α β β and SMulCommClass α β β typeclass arguments are usually satisfied by Algebra α β.

                                          theorem vadd_add_vadd {α : Type u_5} {β : Type u_6} [Add α] [Add β] [VAdd α β] [VAddAssocClass α β β] [VAddAssocClass α α β] [VAddCommClass α β β] (a : α) (b : β) (c : α) (d : β) :
                                          (a +ᵥ b) + (c +ᵥ d) = (a + c) +ᵥ b + d
                                          theorem mul_smul_mul_comm {α : Type u_5} {β : Type u_6} [Mul α] [Mul β] [SMul α β] [IsScalarTower α β β] [IsScalarTower α α β] [SMulCommClass α β β] (a b : α) (c d : β) :
                                          (a * b) (c * d) = a c * b d

                                          Note that the IsScalarTower α β β and SMulCommClass α β β typeclass arguments are usually satisfied by Algebra α β.

                                          theorem add_vadd_add_comm {α : Type u_5} {β : Type u_6} [Add α] [Add β] [VAdd α β] [VAddAssocClass α β β] [VAddAssocClass α α β] [VAddCommClass α β β] (a b : α) (c d : β) :
                                          (a + b) +ᵥ c + d = (a +ᵥ c) + (b +ᵥ d)
                                          theorem Commute.smul_right {M : Type u_1} {α : Type u_5} [SMul M α] [Mul α] [SMulCommClass M α α] [IsScalarTower M α α] {a b : α} (h : Commute a b) (r : M) :
                                          Commute a (r b)
                                          theorem AddCommute.vadd_right {M : Type u_1} {α : Type u_5} [VAdd M α] [Add α] [VAddCommClass M α α] [VAddAssocClass M α α] {a b : α} (h : AddCommute a b) (r : M) :
                                          theorem Commute.smul_left {M : Type u_1} {α : Type u_5} [SMul M α] [Mul α] [SMulCommClass M α α] [IsScalarTower M α α] {a b : α} (h : Commute a b) (r : M) :
                                          Commute (r a) b
                                          theorem AddCommute.vadd_left {M : Type u_1} {α : Type u_5} [VAdd M α] [Add α] [VAddCommClass M α α] [VAddAssocClass M α α] {a b : α} (h : AddCommute a b) (r : M) :
                                          theorem smul_smul {M : Type u_1} {α : Type u_5} [Monoid M] [MulAction M α] (a₁ a₂ : M) (b : α) :
                                          a₁ a₂ b = (a₁ * a₂) b
                                          theorem vadd_vadd {M : Type u_1} {α : Type u_5} [AddMonoid M] [AddAction M α] (a₁ a₂ : M) (b : α) :
                                          a₁ +ᵥ a₂ +ᵥ b = (a₁ + a₂) +ᵥ b
                                          @[simp]
                                          theorem one_smul (M : Type u_1) {α : Type u_5} [Monoid M] [MulAction M α] (b : α) :
                                          1 b = b
                                          @[simp]
                                          theorem zero_vadd (M : Type u_1) {α : Type u_5} [AddMonoid M] [AddAction M α] (b : α) :
                                          0 +ᵥ b = b
                                          theorem one_smul_eq_id (M : Type u_1) {α : Type u_5} [Monoid M] [MulAction M α] :
                                          (fun (x : α) => 1 x) = id

                                          SMul version of one_mul_eq_id

                                          theorem zero_vadd_eq_id (M : Type u_1) {α : Type u_5} [AddMonoid M] [AddAction M α] :
                                          (fun (x : α) => 0 +ᵥ x) = id

                                          VAdd version of zero_add_eq_id

                                          theorem comp_smul_left (M : Type u_1) {α : Type u_5} [Monoid M] [MulAction M α] (a₁ a₂ : M) :
                                          ((fun (x : α) => a₁ x) fun (x : α) => a₂ x) = fun (x : α) => (a₁ * a₂) x

                                          SMul version of comp_mul_left

                                          theorem comp_vadd_left (M : Type u_1) {α : Type u_5} [AddMonoid M] [AddAction M α] (a₁ a₂ : M) :
                                          ((fun (x : α) => a₁ +ᵥ x) fun (x : α) => a₂ +ᵥ x) = fun (x : α) => (a₁ + a₂) +ᵥ x

                                          VAdd version of comp_add_left

                                          @[simp]
                                          theorem smul_iterate {M : Type u_1} {α : Type u_5} [Monoid M] [MulAction M α] (a : M) (n : ) :
                                          (fun (x : α) => a x)^[n] = fun (x : α) => a ^ n x
                                          @[simp]
                                          theorem vadd_iterate {M : Type u_1} {α : Type u_5} [AddMonoid M] [AddAction M α] (a : M) (n : ) :
                                          (fun (x : α) => a +ᵥ x)^[n] = fun (x : α) => n a +ᵥ x
                                          theorem smul_iterate_apply {M : Type u_1} {α : Type u_5} [Monoid M] [MulAction M α] (a : M) (n : ) (x : α) :
                                          (fun (x : α) => a x)^[n] x = a ^ n x
                                          theorem vadd_iterate_apply {M : Type u_1} {α : Type u_5} [AddMonoid M] [AddAction M α] (a : M) (n : ) (x : α) :
                                          (fun (x : α) => a +ᵥ x)^[n] x = n a +ᵥ x
                                          @[reducible, inline]
                                          abbrev Function.Injective.mulAction {M : Type u_1} {α : Type u_5} {β : Type u_6} [Monoid M] [MulAction M α] [SMul M β] (f : βα) (hf : Injective f) (smul : ∀ (c : M) (x : β), f (c x) = c f x) :

                                          Pullback a multiplicative action along an injective map respecting . See note [reducible non-instances].

                                          Equations
                                            Instances For
                                              @[reducible, inline]
                                              abbrev Function.Injective.addAction {M : Type u_1} {α : Type u_5} {β : Type u_6} [AddMonoid M] [AddAction M α] [VAdd M β] (f : βα) (hf : Injective f) (smul : ∀ (c : M) (x : β), f (c +ᵥ x) = c +ᵥ f x) :

                                              Pullback an additive action along an injective map respecting +ᵥ.

                                              Equations
                                                Instances For
                                                  @[reducible, inline]
                                                  abbrev Function.Surjective.mulAction {M : Type u_1} {α : Type u_5} {β : Type u_6} [Monoid M] [MulAction M α] [SMul M β] (f : αβ) (hf : Surjective f) (smul : ∀ (c : M) (x : α), f (c x) = c f x) :

                                                  Pushforward a multiplicative action along a surjective map respecting . See note [reducible non-instances].

                                                  Equations
                                                    Instances For
                                                      @[reducible, inline]
                                                      abbrev Function.Surjective.addAction {M : Type u_1} {α : Type u_5} {β : Type u_6} [AddMonoid M] [AddAction M α] [VAdd M β] (f : αβ) (hf : Surjective f) (smul : ∀ (c : M) (x : α), f (c +ᵥ x) = c +ᵥ f x) :

                                                      Pushforward an additive action along a surjective map respecting +ᵥ.

                                                      Equations
                                                        Instances For
                                                          @[instance 910]
                                                          instance Monoid.toMulAction (M : Type u_1) [Monoid M] :

                                                          The regular action of a monoid on itself by left multiplication.

                                                          This is promoted to a module by Semiring.toModule.

                                                          Equations
                                                            @[instance 910]
                                                            instance AddMonoid.toAddAction (M : Type u_1) [AddMonoid M] :

                                                            The regular action of a monoid on itself by left addition.

                                                            This is promoted to an AddTorsor by addGroup_is_addTorsor.

                                                            Equations
                                                              instance IsScalarTower.left (M : Type u_1) {α : Type u_5} [Monoid M] [MulAction M α] :
                                                              instance VAddAssocClass.left (M : Type u_1) {α : Type u_5} [AddMonoid M] [AddAction M α] :
                                                              theorem smul_pow {M : Type u_1} {N : Type u_2} [Monoid M] [Monoid N] [MulAction M N] [IsScalarTower M N N] [SMulCommClass M N N] (r : M) (x : N) (n : ) :
                                                              (r x) ^ n = r ^ n x ^ n
                                                              @[simp]
                                                              theorem inv_smul_smul {G : Type u_3} {α : Type u_5} [Group G] [MulAction G α] (g : G) (a : α) :
                                                              g⁻¹ g a = a
                                                              @[simp]
                                                              theorem neg_vadd_vadd {G : Type u_3} {α : Type u_5} [AddGroup G] [AddAction G α] (g : G) (a : α) :
                                                              -g +ᵥ g +ᵥ a = a
                                                              @[simp]
                                                              theorem smul_inv_smul {G : Type u_3} {α : Type u_5} [Group G] [MulAction G α] (g : G) (a : α) :
                                                              g g⁻¹ a = a
                                                              @[simp]
                                                              theorem vadd_neg_vadd {G : Type u_3} {α : Type u_5} [AddGroup G] [AddAction G α] (g : G) (a : α) :
                                                              g +ᵥ -g +ᵥ a = a
                                                              theorem inv_smul_eq_iff {G : Type u_3} {α : Type u_5} [Group G] [MulAction G α] {g : G} {a b : α} :
                                                              g⁻¹ a = b a = g b
                                                              theorem neg_vadd_eq_iff {G : Type u_3} {α : Type u_5} [AddGroup G] [AddAction G α] {g : G} {a b : α} :
                                                              -g +ᵥ a = b a = g +ᵥ b
                                                              theorem eq_inv_smul_iff {G : Type u_3} {α : Type u_5} [Group G] [MulAction G α] {g : G} {a b : α} :
                                                              a = g⁻¹ b g a = b
                                                              theorem eq_neg_vadd_iff {G : Type u_3} {α : Type u_5} [AddGroup G] [AddAction G α] {g : G} {a b : α} :
                                                              a = -g +ᵥ b g +ᵥ a = b
                                                              @[simp]
                                                              theorem Commute.smul_right_iff {G : Type u_3} {H : Type u_4} [Group G] {g : G} [Mul H] [MulAction G H] [SMulCommClass G H H] [IsScalarTower G H H] {a b : H} :
                                                              Commute a (g b) Commute a b
                                                              @[simp]
                                                              theorem Commute.smul_left_iff {G : Type u_3} {H : Type u_4} [Group G] {g : G} [Mul H] [MulAction G H] [SMulCommClass G H H] [IsScalarTower G H H] {a b : H} :
                                                              Commute (g a) b Commute a b
                                                              theorem smul_inv {G : Type u_3} {H : Type u_4} [Group G] [Group H] [MulAction G H] [SMulCommClass G H H] [IsScalarTower G H H] (g : G) (a : H) :
                                                              theorem smul_zpow {G : Type u_3} {H : Type u_4} [Group G] [Group H] [MulAction G H] [SMulCommClass G H H] [IsScalarTower G H H] (g : G) (a : H) (n : ) :
                                                              (g a) ^ n = g ^ n a ^ n
                                                              theorem SMulCommClass.of_commMonoid (A : Type u_9) (B : Type u_10) (G : Type u_11) [CommMonoid G] [SMul A G] [SMul B G] [IsScalarTower A G G] [IsScalarTower B G G] :
                                                              theorem smul_one_smul {α : Type u_5} {M : Type u_9} (N : Type u_10) [Monoid N] [SMul M N] [MulAction N α] [SMul M α] [IsScalarTower M N α] (x : M) (y : α) :
                                                              (x 1) y = x y
                                                              theorem vadd_zero_vadd {α : Type u_5} {M : Type u_9} (N : Type u_10) [AddMonoid N] [VAdd M N] [AddAction N α] [VAdd M α] [VAddAssocClass M N α] (x : M) (y : α) :
                                                              (x +ᵥ 0) +ᵥ y = x +ᵥ y
                                                              @[simp]
                                                              theorem smul_one_mul {M : Type u_9} {N : Type u_10} [MulOneClass N] [SMul M N] [IsScalarTower M N N] (x : M) (y : N) :
                                                              x 1 * y = x y
                                                              @[simp]
                                                              theorem vadd_zero_add {M : Type u_9} {N : Type u_10} [AddZeroClass N] [VAdd M N] [VAddAssocClass M N N] (x : M) (y : N) :
                                                              (x +ᵥ 0) + y = x +ᵥ y
                                                              @[simp]
                                                              theorem mul_smul_one {M : Type u_9} {N : Type u_10} [MulOneClass N] [SMul M N] [SMulCommClass M N N] (x : M) (y : N) :
                                                              y * x 1 = x y
                                                              @[simp]
                                                              theorem add_vadd_zero {M : Type u_9} {N : Type u_10} [AddZeroClass N] [VAdd M N] [VAddCommClass M N N] (x : M) (y : N) :
                                                              y + (x +ᵥ 0) = x +ᵥ y
                                                              theorem IsScalarTower.of_smul_one_mul {M : Type u_9} {N : Type u_10} [Monoid N] [SMul M N] (h : ∀ (x : M) (y : N), x 1 * y = x y) :
                                                              theorem VAddAssocClass.of_vadd_zero_add {M : Type u_9} {N : Type u_10} [AddMonoid N] [VAdd M N] (h : ∀ (x : M) (y : N), (x +ᵥ 0) + y = x +ᵥ y) :
                                                              theorem SMulCommClass.of_mul_smul_one {M : Type u_9} {N : Type u_10} [Monoid N] [SMul M N] (H : ∀ (x : M) (y : N), y * x 1 = x y) :
                                                              theorem VAddCommClass.of_add_vadd_zero {M : Type u_9} {N : Type u_10} [AddMonoid N] [VAdd M N] (H : ∀ (x : M) (y : N), y + (x +ᵥ 0) = x +ᵥ y) :
                                                              class MulDistribMulAction (M : Type u_9) (N : Type u_10) [Monoid M] [Monoid N] extends MulAction M N :
                                                              Type (max u_10 u_9)

                                                              Typeclass for multiplicative actions on multiplicative structures.

                                                              The key axiom here is smul_mul : g • (x * y) = (g • x) * (g • y). If G is a group (with group law multiplication) and Γ is its automorphism group then there is a natural instance of MulDistribMulAction Γ G.

                                                              The axiom is also satisfied by a Galois group $Gal(L/K)$ acting on the field L, but here you can use the even stronger class MulSemiringAction, which captures how the action plays with both multiplication and addition.

                                                              • smul : MNN
                                                              • one_smul (b : N) : 1 b = b
                                                              • mul_smul (x y : M) (b : N) : (x * y) b = x y b
                                                              • smul_mul (r : M) (x y : N) : r (x * y) = r x * r y

                                                                Distributivity of across *

                                                              • smul_one (r : M) : r 1 = 1

                                                                Multiplying 1 by a scalar gives 1

                                                              Instances
                                                                theorem MulDistribMulAction.ext_iff {M : Type u_9} {N : Type u_10} {inst✝ : Monoid M} {inst✝¹ : Monoid N} {x y : MulDistribMulAction M N} :
                                                                theorem MulDistribMulAction.ext {M : Type u_9} {N : Type u_10} {inst✝ : Monoid M} {inst✝¹ : Monoid N} {x y : MulDistribMulAction M N} (smul : SMul.smul = SMul.smul) :
                                                                x = y
                                                                theorem smul_mul' {M : Type u_1} {N : Type u_2} [Monoid M] [Monoid N] [MulDistribMulAction M N] (a : M) (b₁ b₂ : N) :
                                                                a (b₁ * b₂) = a b₁ * a b₂