Documentation

Mathlib.GroupTheory.GroupAction.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 for faithful and transitive actions, and 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

Faithful actions #

class FaithfulVAdd (G : Type u_10) (P : Type u_11) [VAdd G P] :
  • eq_of_vadd_eq_vadd : ∀ {g₁ g₂ : G}, (∀ (p : P), g₁ +ᵥ p = g₂ +ᵥ p) → g₁ = g₂

    Two elements g₁ and g₂ are equal whenever they act in the same way on all points.

Typeclass for faithful actions.

Instances
    class FaithfulSMul (M : Type u_10) (α : Type u_11) [SMul M α] :
    • eq_of_smul_eq_smul : ∀ {m₁ m₂ : M}, (∀ (a : α), m₁ a = m₂ a) → m₁ = m₂

      Two elements m₁ and m₂ are equal whenever they act in the same way on all points.

    Typeclass for faithful actions.

    Instances
      theorem vadd_left_injective' {M : Type u_1} {α : Type u_6} [VAdd M α] [FaithfulVAdd M α] :
      Function.Injective fun x x_1 => x +ᵥ x_1
      theorem smul_left_injective' {M : Type u_1} {α : Type u_6} [SMul M α] [FaithfulSMul M α] :
      Function.Injective fun x x_1 => x x_1
      instance Add.toVAdd (α : Type u_10) [Add α] :
      VAdd α α

      See also AddMonoid.toAddAction

      Equations
      instance Mul.toSMul (α : Type u_10) [Mul α] :
      SMul α α

      See also Monoid.toMulAction and MulZeroClass.toSMulWithZero.

      Equations
      @[simp]
      theorem vadd_eq_add (α : Type u_10) [Add α] {a : α} {a' : α} :
      a +ᵥ a' = a + a'
      @[simp]
      theorem smul_eq_mul (α : Type u_10) [Mul α] {a : α} {a' : α} :
      a a' = a * a'
      class AddAction (G : Type u_10) (P : Type u_11) [AddMonoid G] extends VAdd :
      Type (max u_10 u_11)
      • 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 +ᵥ

      Type class for additive monoid actions.

      Instances
        theorem AddAction.ext {G : Type u_10} {P : Type u_11} :
        ∀ {inst : AddMonoid G} (x y : AddAction G P), VAdd.vadd = VAdd.vaddx = y
        theorem AddAction.ext_iff {G : Type u_10} {P : Type u_11} :
        ∀ {inst : AddMonoid G} (x y : AddAction G P), x = y VAdd.vadd = VAdd.vadd
        theorem MulAction.ext_iff {α : Type u_10} {β : Type u_11} :
        ∀ {inst : Monoid α} (x y : MulAction α β), x = y SMul.smul = SMul.smul
        theorem MulAction.ext {α : Type u_10} {β : Type u_11} :
        ∀ {inst : Monoid α} (x y : MulAction α β), SMul.smul = SMul.smulx = y
        class MulAction (α : Type u_10) (β : Type u_11) [Monoid α] extends SMul :
        Type (max u_10 u_11)
        • 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 *

        Typeclass for multiplicative actions by monoids. This generalizes group actions.

        Instances

          (Pre)transitive action #

          M acts pretransitively on α if for any x y there is g such that g • x = y (or g +ᵥ x = y for an additive action). A transitive action should furthermore have α nonempty.

          In this section we define typeclasses MulAction.IsPretransitive and AddAction.IsPretransitive and provide MulAction.exists_smul_eq/AddAction.exists_vadd_eq, MulAction.surjective_smul/AddAction.surjective_vadd as public interface to access this property. We do not provide typeclasses *Action.IsTransitive; users should assume [MulAction.IsPretransitive M α] [Nonempty α] instead.

          class AddAction.IsPretransitive (M : Type u_10) (α : Type u_11) [VAdd M α] :
          • exists_vadd_eq : ∀ (x y : α), g, g +ᵥ x = y

            There is g such that g +ᵥ x = y.

          M acts pretransitively on α if for any x y there is g such that g +ᵥ x = y. A transitive action should furthermore have α nonempty.

          Instances
            class MulAction.IsPretransitive (M : Type u_10) (α : Type u_11) [SMul M α] :
            • exists_smul_eq : ∀ (x y : α), g, g x = y

              There is g such that g • x = y.

            M acts pretransitively on α if for any x y there is g such that g • x = y. A transitive action should furthermore have α nonempty.

            Instances
              theorem AddAction.exists_vadd_eq (M : Type u_1) {α : Type u_6} [VAdd M α] [AddAction.IsPretransitive M α] (x : α) (y : α) :
              m, m +ᵥ x = y
              theorem MulAction.exists_smul_eq (M : Type u_1) {α : Type u_6} [SMul M α] [MulAction.IsPretransitive M α] (x : α) (y : α) :
              m, m x = y
              theorem AddAction.surjective_vadd (M : Type u_1) {α : Type u_6} [VAdd M α] [AddAction.IsPretransitive M α] (x : α) :
              theorem MulAction.surjective_smul (M : Type u_1) {α : Type u_6} [SMul M α] [MulAction.IsPretransitive M α] (x : α) :

              Scalar tower and commuting actions #

              class VAddCommClass (M : Type u_10) (N : Type u_11) (α : Type u_12) [VAdd M α] [VAdd N α] :
              • vadd_comm : ∀ (m : M) (n : N) (a : α), m +ᵥ (n +ᵥ a) = n +ᵥ (m +ᵥ a)

                +ᵥ is left commutative

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

              Instances
                class SMulCommClass (M : Type u_10) (N : Type u_11) (α : Type u_12) [SMul M α] [SMul N α] :
                • smul_comm : ∀ (m : M) (n : N) (a : α), m n a = n m a

                  is left commutative

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

                Instances
                  theorem VAddCommClass.symm (M : Type u_10) (N : Type u_11) (α : Type u_12) [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 SMulCommClass.symm (M : Type u_10) (N : Type u_11) (α : Type u_12) [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 Function.Injective.smulCommClass {M : Type u_1} {N : Type u_2} {α : Type u_6} {β : Type u_7} [SMul M α] [SMul N α] [SMul M β] [SMul N β] [SMulCommClass M N β] {f : αβ} (hf : Function.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_6} {β : Type u_7} [SMul M α] [SMul N α] [SMul M β] [SMul N β] [SMulCommClass M N α] {f : αβ} (hf : Function.Surjective f) (h₁ : ∀ (c : M) (x : α), f (c x) = c f x) (h₂ : ∀ (c : N) (x : α), f (c x) = c f x) :
                  theorem vaddCommClass_self.proof_1 (M : Type u_1) (α : Type u_2) [AddCommMonoid M] [AddAction M α] :
                  class VAddAssocClass (M : Type u_10) (N : Type u_11) (α : Type u_12) [VAdd M N] [VAdd N α] [VAdd M α] :
                  • vadd_assoc : ∀ (x : M) (y : N) (z : α), x +ᵥ y +ᵥ z = x +ᵥ (y +ᵥ z)

                    Associativity of +ᵥ

                  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 α.

                  Instances
                    class IsScalarTower (M : Type u_10) (N : Type u_11) (α : Type u_12) [SMul M N] [SMul N α] [SMul M α] :
                    • smul_assoc : ∀ (x : M) (y : N) (z : α), (x y) z = x y z

                      Associativity of

                    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 α.

                    Instances
                      @[simp]
                      theorem vadd_assoc {α : Type u_6} {M : Type u_10} {N : Type u_11} [VAdd M N] [VAdd N α] [VAdd M α] [VAddAssocClass M N α] (x : M) (y : N) (z : α) :
                      x +ᵥ y +ᵥ z = x +ᵥ (y +ᵥ z)
                      @[simp]
                      theorem smul_assoc {α : Type u_6} {M : Type u_10} {N : Type u_11} [SMul M N] [SMul N α] [SMul M α] [IsScalarTower M N α] (x : M) (y : N) (z : α) :
                      (x y) z = x y z
                      class IsCentralVAdd (M : Type u_10) (α : Type u_11) [VAdd M α] [VAdd Mᵃᵒᵖ α] :
                      • op_vadd_eq_vadd : ∀ (m : M) (a : α), AddOpposite.op m +ᵥ a = m +ᵥ a

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

                      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 +ᵥ.

                      Instances
                        class IsCentralScalar (M : Type u_10) (α : Type u_11) [SMul M α] [SMul Mᵐᵒᵖ α] :
                        • op_smul_eq_smul : ∀ (m : M) (a : α), MulOpposite.op m a = m a

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

                        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 .

                        Instances
                          theorem IsCentralVAdd.unop_vadd_eq_vadd {M : Type u_10} {α : Type u_11} [VAdd M α] [VAdd Mᵃᵒᵖ α] [IsCentralVAdd M α] (m : Mᵃᵒᵖ) (a : α) :
                          theorem IsCentralScalar.unop_smul_eq_smul {M : Type u_10} {α : Type u_11} [SMul M α] [SMul Mᵐᵒᵖ α] [IsCentralScalar M α] (m : Mᵐᵒᵖ) (a : α) :
                          theorem VAddCommClass.op_left.proof_1 {M : Type u_1} {N : Type u_2} {α : Type u_3} [VAdd M α] [VAdd Mᵃᵒᵖ α] [IsCentralVAdd M α] [VAdd N α] [VAddCommClass M N α] :
                          theorem VAddCommClass.op_right.proof_1 {M : Type u_1} {N : Type u_2} {α : Type u_3} [VAdd M α] [VAdd N α] [VAdd Nᵃᵒᵖ α] [IsCentralVAdd N α] [VAddCommClass M N α] :
                          theorem VAddAssocClass.op_left.proof_1 {M : Type u_1} {N : Type u_2} {α : Type u_3} [VAdd M α] [VAdd Mᵃᵒᵖ α] [IsCentralVAdd M α] [VAdd M N] [VAdd Mᵃᵒᵖ N] [IsCentralVAdd M N] [VAdd N α] [VAddAssocClass M N α] :
                          theorem VAddAssocClass.op_right.proof_1 {M : Type u_1} {N : Type u_2} {α : Type u_3} [VAdd M α] [VAdd M N] [VAdd N α] [VAdd Nᵃᵒᵖ α] [IsCentralVAdd N α] [VAddAssocClass M N α] :
                          def VAdd.comp.vadd {M : Type u_1} {N : Type u_2} {α : Type u_6} [VAdd M α] (g : NM) (n : N) (a : α) :
                          α

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

                          Equations
                          Instances For
                            def SMul.comp.smul {M : Type u_1} {N : Type u_2} {α : Type u_6} [SMul M α] (g : NM) (n : N) (a : α) :
                            α

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

                            Equations
                            Instances For
                              @[reducible]
                              def VAdd.comp {M : Type u_1} {N : Type u_2} (α : Type u_6) [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
                                @[reducible]
                                def SMul.comp {M : Type u_1} {N : Type u_2} (α : Type u_6) [SMul M α] (g : NM) :
                                SMul N α

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

                                See note [reducible non-instances]. Since this is reducible, we make sure to go via SMul.comp.smul to prevent typeclass inference unfolding too far.

                                Equations
                                Instances For
                                  theorem VAdd.comp.isScalarTower {M : Type u_1} {N : Type u_2} {α : Type u_6} {β : Type u_7} [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.isScalarTower {M : Type u_1} {N : Type u_2} {α : Type u_6} {β : Type u_7} [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.vaddCommClass {M : Type u_1} {N : Type u_2} {α : Type u_6} {β : Type u_7} [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_6} {β : Type u_7} [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_6} {β : Type u_7} [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_6} {β : Type u_7} [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 add_vadd_comm {α : Type u_6} {β : Type u_7} [Add β] [VAdd α β] [VAddCommClass α β β] (s : α) (x : β) (y : β) :
                                  x + (s +ᵥ y) = s +ᵥ (x + y)
                                  theorem mul_smul_comm {α : Type u_6} {β : Type u_7} [Mul β] [SMul α β] [SMulCommClass α β β] (s : α) (x : β) (y : β) :
                                  x * s y = s (x * y)

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

                                  theorem vadd_add_assoc {α : Type u_6} {β : Type u_7} [Add β] [VAdd α β] [VAddAssocClass α β β] (r : α) (x : β) (y : β) :
                                  r +ᵥ x + y = r +ᵥ (x + y)
                                  theorem smul_mul_assoc {α : Type u_6} {β : Type u_7} [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_vadd_vadd_comm {α : Type u_6} {β : Type u_7} {γ : Type u_8} {δ : Type u_9} [VAdd α β] [VAdd α γ] [VAdd β δ] [VAdd α δ] [VAdd γ δ] [VAddAssocClass α β δ] [VAddAssocClass α γ δ] [VAddCommClass β γ δ] (a : α) (b : β) (c : γ) (d : δ) :
                                  a +ᵥ b +ᵥ (c +ᵥ d) = a +ᵥ c +ᵥ (b +ᵥ d)
                                  theorem smul_smul_smul_comm {α : Type u_6} {β : Type u_7} {γ : Type u_8} {δ : Type u_9} [SMul α β] [SMul α γ] [SMul β δ] [SMul α δ] [SMul γ δ] [IsScalarTower α β δ] [IsScalarTower α γ δ] [SMulCommClass β γ δ] (a : α) (b : β) (c : γ) (d : δ) :
                                  (a b) c d = (a c) b d
                                  theorem AddCommute.vadd_right {M : Type u_1} {α : Type u_6} [VAdd M α] [Add α] [VAddCommClass M α α] [VAddAssocClass M α α] {a : α} {b : α} (h : AddCommute a b) (r : M) :
                                  theorem Commute.smul_right {M : Type u_1} {α : Type u_6} [SMul M α] [Mul α] [SMulCommClass M α α] [IsScalarTower M α α] {a : α} {b : α} (h : Commute a b) (r : M) :
                                  Commute a (r b)
                                  theorem AddCommute.vadd_left {M : Type u_1} {α : Type u_6} [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_6} [SMul M α] [Mul α] [SMulCommClass M α α] [IsScalarTower M α α] {a : α} {b : α} (h : Commute a b) (r : M) :
                                  Commute (r a) b
                                  theorem ite_vadd {M : Type u_1} {α : Type u_6} [VAdd M α] (p : Prop) [Decidable p] (a₁ : M) (a₂ : M) (b : α) :
                                  (if p then a₁ else a₂) +ᵥ b = if p then a₁ +ᵥ b else a₂ +ᵥ b
                                  theorem ite_smul {M : Type u_1} {α : Type u_6} [SMul M α] (p : Prop) [Decidable p] (a₁ : M) (a₂ : M) (b : α) :
                                  (if p then a₁ else a₂) b = if p then a₁ b else a₂ b
                                  theorem vadd_ite {M : Type u_1} {α : Type u_6} [VAdd M α] (p : Prop) [Decidable p] (a : M) (b₁ : α) (b₂ : α) :
                                  (a +ᵥ if p then b₁ else b₂) = if p then a +ᵥ b₁ else a +ᵥ b₂
                                  theorem smul_ite {M : Type u_1} {α : Type u_6} [SMul M α] (p : Prop) [Decidable p] (a : M) (b₁ : α) (b₂ : α) :
                                  (a if p then b₁ else b₂) = if p then a b₁ else a b₂
                                  theorem vadd_vadd {M : Type u_1} {α : Type u_6} [AddMonoid M] [AddAction M α] (a₁ : M) (a₂ : M) (b : α) :
                                  a₁ +ᵥ (a₂ +ᵥ b) = a₁ + a₂ +ᵥ b
                                  theorem smul_smul {M : Type u_1} {α : Type u_6} [Monoid M] [MulAction M α] (a₁ : M) (a₂ : M) (b : α) :
                                  a₁ a₂ b = (a₁ * a₂) b
                                  @[simp]
                                  theorem zero_vadd (M : Type u_1) {α : Type u_6} [AddMonoid M] [AddAction M α] (b : α) :
                                  0 +ᵥ b = b
                                  @[simp]
                                  theorem one_smul (M : Type u_1) {α : Type u_6} [Monoid M] [MulAction M α] (b : α) :
                                  1 b = b
                                  theorem zero_vadd_eq_id (M : Type u_1) {α : Type u_6} [AddMonoid M] [AddAction M α] :
                                  (fun x x_1 => x +ᵥ x_1) 0 = id

                                  VAdd version of zero_add_eq_id

                                  theorem one_smul_eq_id (M : Type u_1) {α : Type u_6} [Monoid M] [MulAction M α] :
                                  (fun x x_1 => x x_1) 1 = id

                                  SMul version of one_mul_eq_id

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

                                  VAdd version of comp_add_left

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

                                  SMul version of comp_mul_left

                                  theorem Function.Injective.addAction.proof_1 {M : Type u_2} {α : Type u_3} {β : Type u_1} [AddMonoid M] [AddAction M α] [VAdd M β] (f : βα) (hf : Function.Injective f) (smul : ∀ (c : M) (x : β), f (c +ᵥ x) = c +ᵥ f x) (x : β) :
                                  0 +ᵥ x = x
                                  @[reducible]
                                  def Function.Injective.addAction {M : Type u_1} {α : Type u_6} {β : Type u_7} [AddMonoid M] [AddAction M α] [VAdd M β] (f : βα) (hf : Function.Injective f) (smul : ∀ (c : M) (x : β), f (c +ᵥ x) = c +ᵥ f x) :

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

                                  Equations
                                  Instances For
                                    theorem Function.Injective.addAction.proof_2 {M : Type u_2} {α : Type u_3} {β : Type u_1} [AddMonoid M] [AddAction M α] [VAdd M β] (f : βα) (hf : Function.Injective f) (smul : ∀ (c : M) (x : β), f (c +ᵥ x) = c +ᵥ f x) (c₁ : M) (c₂ : M) (x : β) :
                                    c₁ + c₂ +ᵥ x = c₁ +ᵥ (c₂ +ᵥ x)
                                    @[reducible]
                                    def Function.Injective.mulAction {M : Type u_1} {α : Type u_6} {β : Type u_7} [Monoid M] [MulAction M α] [SMul M β] (f : βα) (hf : Function.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
                                      theorem Function.Surjective.addAction.proof_2 {M : Type u_2} {α : Type u_3} {β : Type u_1} [AddMonoid M] [AddAction M α] [VAdd M β] (f : αβ) (hf : Function.Surjective f) (smul : ∀ (c : M) (x : α), f (c +ᵥ x) = c +ᵥ f x) (c₁ : M) (c₂ : M) (y : β) :
                                      c₁ + c₂ +ᵥ y = c₁ +ᵥ (c₂ +ᵥ y)
                                      theorem Function.Surjective.addAction.proof_1 {M : Type u_2} {α : Type u_3} {β : Type u_1} [AddMonoid M] [AddAction M α] [VAdd M β] (f : αβ) (hf : Function.Surjective f) (smul : ∀ (c : M) (x : α), f (c +ᵥ x) = c +ᵥ f x) (y : β) :
                                      0 +ᵥ y = y
                                      @[reducible]
                                      def Function.Surjective.addAction {M : Type u_1} {α : Type u_6} {β : Type u_7} [AddMonoid M] [AddAction M α] [VAdd M β] (f : αβ) (hf : Function.Surjective f) (smul : ∀ (c : M) (x : α), f (c +ᵥ x) = c +ᵥ f x) :

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

                                      Equations
                                      Instances For
                                        @[reducible]
                                        def Function.Surjective.mulAction {M : Type u_1} {α : Type u_6} {β : Type u_7} [Monoid M] [MulAction M α] [SMul M β] (f : αβ) (hf : Function.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
                                          theorem Function.Surjective.addActionLeft.proof_2 {R : Type u_3} {S : Type u_2} {M : Type u_1} [AddMonoid R] [AddAction R M] [AddMonoid S] [VAdd S M] (f : R →+ S) (hf : Function.Surjective f) (hsmul : ∀ (c : R) (x : M), f c +ᵥ x = c +ᵥ x) (y₁ : S) (y₂ : S) (b : M) :
                                          y₁ + y₂ +ᵥ b = y₁ +ᵥ (y₂ +ᵥ b)
                                          @[reducible]
                                          def Function.Surjective.addActionLeft {R : Type u_10} {S : Type u_11} {M : Type u_12} [AddMonoid R] [AddAction R M] [AddMonoid S] [VAdd S M] (f : R →+ S) (hf : Function.Surjective f) (hsmul : ∀ (c : R) (x : M), f c +ᵥ x = c +ᵥ x) :

                                          Push forward the action of R on M along a compatible surjective map f : R →+ S.

                                          Equations
                                          Instances For
                                            theorem Function.Surjective.addActionLeft.proof_1 {R : Type u_3} {S : Type u_2} {M : Type u_1} [AddMonoid R] [AddAction R M] [AddMonoid S] [VAdd S M] (f : R →+ S) (hsmul : ∀ (c : R) (x : M), f c +ᵥ x = c +ᵥ x) (b : M) :
                                            0 +ᵥ b = b
                                            @[reducible]
                                            def Function.Surjective.mulActionLeft {R : Type u_10} {S : Type u_11} {M : Type u_12} [Monoid R] [MulAction R M] [Monoid S] [SMul S M] (f : R →* S) (hf : Function.Surjective f) (hsmul : ∀ (c : R) (x : M), f c x = c x) :

                                            Push forward the action of R on M along a compatible surjective map f : R →* S.

                                            See also Function.Surjective.distribMulActionLeft and Function.Surjective.moduleLeft.

                                            Equations
                                            Instances For
                                              theorem AddMonoid.toAddAction.proof_1 (M : Type u_1) [AddMonoid M] (a : M) :
                                              0 + a = a
                                              theorem AddMonoid.toAddAction.proof_2 (M : Type u_1) [AddMonoid M] (a : M) (b : M) (c : M) :
                                              a + b + c = a + (b + c)
                                              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 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
                                              theorem VAddAssocClass.left.proof_1 (M : Type u_1) {α : Type u_2} [AddMonoid M] [AddAction M α] :
                                              theorem vadd_add_vadd {M : Type u_1} {α : Type u_6} [AddMonoid M] [AddAction M α] [Add α] (r : M) (s : M) (x : α) (y : α) [VAddAssocClass M α α] [VAddCommClass M α α] :
                                              r +ᵥ x + (s +ᵥ y) = r + s +ᵥ (x + y)
                                              theorem smul_mul_smul {M : Type u_1} {α : Type u_6} [Monoid M] [MulAction M α] [Mul α] (r : M) (s : M) (x : α) (y : α) [IsScalarTower M α α] [SMulCommClass M α α] :
                                              r x * s y = (r * s) (x * y)

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

                                              def AddAction.toFun (M : Type u_1) (α : Type u_6) [AddMonoid M] [AddAction M α] :
                                              α Mα

                                              Embedding of α into functions M → α induced by an additive action of M on α.

                                              Equations
                                              • AddAction.toFun M α = { toFun := fun y x => x +ᵥ y, inj' := (_ : ∀ (y₁ y₂ : α), (fun y x => x +ᵥ y) y₁ = (fun y x => x +ᵥ y) y₂y₁ = y₂) }
                                              Instances For
                                                theorem AddAction.toFun.proof_1 (M : Type u_1) (α : Type u_2) [AddMonoid M] [AddAction M α] (y₁ : α) (y₂ : α) (H : (fun y x => x +ᵥ y) y₁ = (fun y x => x +ᵥ y) y₂) :
                                                y₁ = y₂
                                                def MulAction.toFun (M : Type u_1) (α : Type u_6) [Monoid M] [MulAction M α] :
                                                α Mα

                                                Embedding of α into functions M → α induced by a multiplicative action of M on α.

                                                Equations
                                                • MulAction.toFun M α = { toFun := fun y x => x y, inj' := (_ : ∀ (y₁ y₂ : α), (fun y x => x y) y₁ = (fun y x => x y) y₂y₁ = y₂) }
                                                Instances For
                                                  @[simp]
                                                  theorem AddAction.toFun_apply {M : Type u_1} {α : Type u_6} [AddMonoid M] [AddAction M α] (x : M) (y : α) :
                                                  ↑(AddAction.toFun M α) y x = x +ᵥ y
                                                  @[simp]
                                                  theorem MulAction.toFun_apply {M : Type u_1} {α : Type u_6} [Monoid M] [MulAction M α] (x : M) (y : α) :
                                                  ↑(MulAction.toFun M α) y x = x y
                                                  theorem AddAction.compHom.proof_2 {M : Type u_3} {N : Type u_2} (α : Type u_1) [AddMonoid M] [AddAction M α] [AddMonoid N] (g : N →+ M) :
                                                  ∀ (x x_1 : N) (x_2 : α), x + x_1 +ᵥ x_2 = x +ᵥ (x_1 +ᵥ x_2)
                                                  @[reducible]
                                                  def AddAction.compHom {M : Type u_1} {N : Type u_2} (α : Type u_6) [AddMonoid M] [AddAction M α] [AddMonoid N] (g : N →+ M) :

                                                  An additive action of M on α and an additive monoid homomorphism N → M induce an additive action of N on α.

                                                  See note [reducible non-instances].

                                                  Equations
                                                  Instances For
                                                    theorem AddAction.compHom.proof_1 {M : Type u_3} {N : Type u_2} (α : Type u_1) [AddMonoid M] [AddAction M α] [AddMonoid N] (g : N →+ M) :
                                                    ∀ (x : α), 0 +ᵥ x = x
                                                    @[reducible]
                                                    def MulAction.compHom {M : Type u_1} {N : Type u_2} (α : Type u_6) [Monoid M] [MulAction M α] [Monoid N] (g : N →* M) :

                                                    A multiplicative action of M on α and a monoid homomorphism N → M induce a multiplicative action of N on α.

                                                    See note [reducible non-instances].

                                                    Equations
                                                    Instances For
                                                      theorem vadd_zero_vadd {α : Type u_6} {M : Type u_10} (N : Type u_11) [AddMonoid N] [VAdd M N] [AddAction N α] [VAdd M α] [VAddAssocClass M N α] (x : M) (y : α) :
                                                      x +ᵥ 0 +ᵥ y = x +ᵥ y
                                                      theorem smul_one_smul {α : Type u_6} {M : Type u_10} (N : Type u_11) [Monoid N] [SMul M N] [MulAction N α] [SMul M α] [IsScalarTower M N α] (x : M) (y : α) :
                                                      (x 1) y = x y
                                                      @[simp]
                                                      theorem vadd_zero_add {M : Type u_10} {N : Type u_11} [AddZeroClass N] [VAdd M N] [VAddAssocClass M N N] (x : M) (y : N) :
                                                      x +ᵥ 0 + y = x +ᵥ y
                                                      @[simp]
                                                      theorem smul_one_mul {M : Type u_10} {N : Type u_11} [MulOneClass N] [SMul M N] [IsScalarTower M N N] (x : M) (y : N) :
                                                      x 1 * y = x y
                                                      @[simp]
                                                      theorem add_vadd_zero {M : Type u_10} {N : Type u_11} [AddZeroClass N] [VAdd M N] [VAddCommClass M N N] (x : M) (y : N) :
                                                      y + (x +ᵥ 0) = x +ᵥ y
                                                      @[simp]
                                                      theorem mul_smul_one {M : Type u_10} {N : Type u_11} [MulOneClass N] [SMul M N] [SMulCommClass M N N] (x : M) (y : N) :
                                                      y * x 1 = x y
                                                      theorem VAddAssocClass.of_vadd_zero_add {M : Type u_10} {N : Type u_11} [AddMonoid N] [VAdd M N] (h : ∀ (x : M) (y : N), x +ᵥ 0 + y = x +ᵥ y) :
                                                      theorem IsScalarTower.of_smul_one_mul {M : Type u_10} {N : Type u_11} [Monoid N] [SMul M N] (h : ∀ (x : M) (y : N), x 1 * y = x y) :
                                                      theorem VAddCommClass.of_add_vadd_zero {M : Type u_10} {N : Type u_11} [AddMonoid N] [VAdd M N] (H : ∀ (x : M) (y : N), y + (x +ᵥ 0) = x +ᵥ y) :
                                                      theorem SMulCommClass.of_mul_smul_one {M : Type u_10} {N : Type u_11} [Monoid N] [SMul M N] (H : ∀ (x : M) (y : N), y * x 1 = x y) :
                                                      def vaddZeroHom {M : Type u_10} {N : Type u_11} [AddMonoid M] [AddMonoid N] [AddAction M N] [VAddAssocClass M N N] :
                                                      M →+ N

                                                      If the additive action of M on N is compatible with addition on N, then fun x => x +ᵥ 0 is an additive monoid homomorphism from M to N.

                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      Instances For
                                                        theorem vaddZeroHom.proof_2 {M : Type u_2} {N : Type u_1} [AddMonoid M] [AddMonoid N] [AddAction M N] [VAddAssocClass M N N] (x : M) (y : M) :
                                                        ZeroHom.toFun { toFun := fun x => x +ᵥ 0, map_zero' := (_ : 0 +ᵥ 0 = 0) } (x + y) = ZeroHom.toFun { toFun := fun x => x +ᵥ 0, map_zero' := (_ : 0 +ᵥ 0 = 0) } x + ZeroHom.toFun { toFun := fun x => x +ᵥ 0, map_zero' := (_ : 0 +ᵥ 0 = 0) } y
                                                        theorem vaddZeroHom.proof_1 {M : Type u_2} {N : Type u_1} [AddMonoid M] [AddMonoid N] [AddAction M N] :
                                                        0 +ᵥ 0 = 0
                                                        @[simp]
                                                        theorem smulOneHom_apply {M : Type u_10} {N : Type u_11} [Monoid M] [Monoid N] [MulAction M N] [IsScalarTower M N N] (x : M) :
                                                        smulOneHom x = x 1
                                                        @[simp]
                                                        theorem vaddZeroHom_apply {M : Type u_10} {N : Type u_11} [AddMonoid M] [AddMonoid N] [AddAction M N] [VAddAssocClass M N N] (x : M) :
                                                        vaddZeroHom x = x +ᵥ 0
                                                        def smulOneHom {M : Type u_10} {N : Type u_11} [Monoid M] [Monoid N] [MulAction M N] [IsScalarTower M N N] :
                                                        M →* N

                                                        If the multiplicative action of M on N is compatible with multiplication on N, then fun x => x • 1 is a monoid homomorphism from M to N.

                                                        Equations
                                                        • One or more equations did not get rendered due to their size.
                                                        Instances For
                                                          class SMulZeroClass (M : Type u_10) (A : Type u_11) [Zero A] extends SMul :
                                                          Type (max u_10 u_11)
                                                          • smul : MAA
                                                          • smul_zero : ∀ (a : M), a 0 = 0

                                                            Multiplying 0 by a scalar gives 0

                                                          Typeclass for scalar multiplication that preserves 0 on the right.

                                                          Instances
                                                            @[simp]
                                                            theorem smul_zero {M : Type u_1} {A : Type u_4} [Zero A] [SMulZeroClass M A] (a : M) :
                                                            a 0 = 0
                                                            @[reducible]
                                                            def Function.Injective.smulZeroClass {M : Type u_1} {A : Type u_4} {B : Type u_5} [Zero A] [SMulZeroClass M A] [Zero B] [SMul M B] (f : ZeroHom B A) (hf : Function.Injective f) (smul : ∀ (c : M) (x : B), f (c x) = c f x) :

                                                            Pullback a zero-preserving scalar multiplication along an injective zero-preserving map. See note [reducible non-instances].

                                                            Equations
                                                            Instances For
                                                              @[reducible]
                                                              def ZeroHom.smulZeroClass {M : Type u_1} {A : Type u_4} {B : Type u_5} [Zero A] [SMulZeroClass M A] [Zero B] [SMul M B] (f : ZeroHom A B) (smul : ∀ (c : M) (x : A), f (c x) = c f x) :

                                                              Pushforward a zero-preserving scalar multiplication along a zero-preserving map. See note [reducible non-instances].

                                                              Equations
                                                              Instances For
                                                                @[reducible]
                                                                def Function.Surjective.smulZeroClassLeft {R : Type u_10} {S : Type u_11} {M : Type u_12} [Zero M] [SMulZeroClass R M] [SMul S M] (f : RS) (hf : Function.Surjective f) (hsmul : ∀ (c : R) (x : M), f c x = c x) :

                                                                Push forward the multiplication of R on M along a compatible surjective map f : R → S.

                                                                See also Function.Surjective.distribMulActionLeft.

                                                                Equations
                                                                Instances For
                                                                  @[reducible]
                                                                  def SMulZeroClass.compFun {M : Type u_1} {N : Type u_2} (A : Type u_4) [Zero A] [SMulZeroClass M A] (f : NM) :

                                                                  Compose a SMulZeroClass with a function, with scalar multiplication f r' • m. See note [reducible non-instances].

                                                                  Equations
                                                                  Instances For
                                                                    @[simp]
                                                                    theorem SMulZeroClass.toZeroHom_apply {M : Type u_1} (A : Type u_4) [Zero A] [SMulZeroClass M A] (x : M) :
                                                                    ∀ (x : A), ↑(SMulZeroClass.toZeroHom A x) x = (fun x x_1 => x x_1) x x
                                                                    def SMulZeroClass.toZeroHom {M : Type u_1} (A : Type u_4) [Zero A] [SMulZeroClass M A] (x : M) :

                                                                    Each element of the scalars defines a zero-preserving map.

                                                                    Equations
                                                                    Instances For
                                                                      theorem DistribSMul.ext_iff {M : Type u_10} {A : Type u_11} :
                                                                      ∀ {inst : AddZeroClass A} (x y : DistribSMul M A), x = y SMul.smul = SMul.smul
                                                                      theorem DistribSMul.ext {M : Type u_10} {A : Type u_11} :
                                                                      ∀ {inst : AddZeroClass A} (x y : DistribSMul M A), SMul.smul = SMul.smulx = y
                                                                      class DistribSMul (M : Type u_10) (A : Type u_11) [AddZeroClass A] extends SMulZeroClass :
                                                                      Type (max u_10 u_11)
                                                                      • smul : MAA
                                                                      • smul_zero : ∀ (a : M), a 0 = 0
                                                                      • smul_add : ∀ (a : M) (x y : A), a (x + y) = a x + a y

                                                                        Scalar multiplication distributes across addition

                                                                      Typeclass for scalar multiplication that preserves 0 and + on the right.

                                                                      This is exactly DistribMulAction without the MulAction part.

                                                                      Instances
                                                                        theorem smul_add {M : Type u_1} {A : Type u_4} [AddZeroClass A] [DistribSMul M A] (a : M) (b₁ : A) (b₂ : A) :
                                                                        a (b₁ + b₂) = a b₁ + a b₂
                                                                        instance AddMonoidHom.smulZeroClass {M : Type u_1} {A : Type u_4} {B : Type u_5} [AddZeroClass A] [DistribSMul M A] [AddZeroClass B] :
                                                                        Equations
                                                                        @[reducible]
                                                                        def Function.Injective.distribSMul {M : Type u_1} {A : Type u_4} {B : Type u_5} [AddZeroClass A] [DistribSMul M A] [AddZeroClass B] [SMul M B] (f : B →+ A) (hf : Function.Injective f) (smul : ∀ (c : M) (x : B), f (c x) = c f x) :

                                                                        Pullback a distributive scalar multiplication along an injective additive monoid homomorphism. See note [reducible non-instances].

                                                                        Equations
                                                                        Instances For
                                                                          @[reducible]
                                                                          def Function.Surjective.distribSMul {M : Type u_1} {A : Type u_4} {B : Type u_5} [AddZeroClass A] [DistribSMul M A] [AddZeroClass B] [SMul M B] (f : A →+ B) (hf : Function.Surjective f) (smul : ∀ (c : M) (x : A), f (c x) = c f x) :

                                                                          Pushforward a distributive scalar multiplication along a surjective additive monoid homomorphism. See note [reducible non-instances].

                                                                          Equations
                                                                          Instances For
                                                                            @[reducible]
                                                                            def Function.Surjective.distribSMulLeft {R : Type u_10} {S : Type u_11} {M : Type u_12} [AddZeroClass M] [DistribSMul R M] [SMul S M] (f : RS) (hf : Function.Surjective f) (hsmul : ∀ (c : R) (x : M), f c x = c x) :

                                                                            Push forward the multiplication of R on M along a compatible surjective map f : R → S.

                                                                            See also Function.Surjective.distribMulActionLeft.

                                                                            Equations
                                                                            Instances For
                                                                              @[reducible]
                                                                              def DistribSMul.compFun {M : Type u_1} {N : Type u_2} (A : Type u_4) [AddZeroClass A] [DistribSMul M A] (f : NM) :

                                                                              Compose a DistribSMul with a function, with scalar multiplication f r' • m. See note [reducible non-instances].

                                                                              Equations
                                                                              Instances For
                                                                                @[simp]
                                                                                theorem DistribSMul.toAddMonoidHom_apply {M : Type u_1} (A : Type u_4) [AddZeroClass A] [DistribSMul M A] (x : M) :
                                                                                ∀ (x : A), ↑(DistribSMul.toAddMonoidHom A x) x = (fun x x_1 => x x_1) x x
                                                                                def DistribSMul.toAddMonoidHom {M : Type u_1} (A : Type u_4) [AddZeroClass A] [DistribSMul M A] (x : M) :
                                                                                A →+ A

                                                                                Each element of the scalars defines an additive monoid homomorphism.

                                                                                Equations
                                                                                • One or more equations did not get rendered due to their size.
                                                                                Instances For
                                                                                  theorem DistribMulAction.ext {M : Type u_10} {A : Type u_11} :
                                                                                  ∀ {inst : Monoid M} {inst_1 : AddMonoid A} (x y : DistribMulAction M A), SMul.smul = SMul.smulx = y
                                                                                  theorem DistribMulAction.ext_iff {M : Type u_10} {A : Type u_11} :
                                                                                  ∀ {inst : Monoid M} {inst_1 : AddMonoid A} (x y : DistribMulAction M A), x = y SMul.smul = SMul.smul
                                                                                  class DistribMulAction (M : Type u_10) (A : Type u_11) [Monoid M] [AddMonoid A] extends MulAction :
                                                                                  Type (max u_10 u_11)
                                                                                  • smul : MAA
                                                                                  • one_smul : ∀ (b : A), 1 b = b
                                                                                  • mul_smul : ∀ (x y : M) (b : A), (x * y) b = x y b
                                                                                  • smul_zero : ∀ (a : M), a 0 = 0

                                                                                    Multiplying 0 by a scalar gives 0

                                                                                  • smul_add : ∀ (a : M) (x y : A), a (x + y) = a x + a y

                                                                                    Scalar multiplication distributes across addition

                                                                                  Typeclass for multiplicative actions on additive structures. This generalizes group modules.

                                                                                  Instances
                                                                                    instance DistribMulAction.toDistribSMul {M : Type u_1} {A : Type u_4} [Monoid M] [AddMonoid A] [DistribMulAction M A] :
                                                                                    Equations

                                                                                    Since Lean 3 does not have definitional eta for structures, we have to make sure that the definition of DistribMulAction.toDistribSMul was done correctly, and the two paths from DistribMulAction to SMul are indeed definitionally equal.

                                                                                    @[reducible]
                                                                                    def Function.Injective.distribMulAction {M : Type u_1} {A : Type u_4} {B : Type u_5} [Monoid M] [AddMonoid A] [DistribMulAction M A] [AddMonoid B] [SMul M B] (f : B →+ A) (hf : Function.Injective f) (smul : ∀ (c : M) (x : B), f (c x) = c f x) :

                                                                                    Pullback a distributive multiplicative action along an injective additive monoid homomorphism. See note [reducible non-instances].

                                                                                    Equations
                                                                                    • One or more equations did not get rendered due to their size.
                                                                                    Instances For
                                                                                      @[reducible]
                                                                                      def Function.Surjective.distribMulAction {M : Type u_1} {A : Type u_4} {B : Type u_5} [Monoid M] [AddMonoid A] [DistribMulAction M A] [AddMonoid B] [SMul M B] (f : A →+ B) (hf : Function.Surjective f) (smul : ∀ (c : M) (x : A), f (c x) = c f x) :

                                                                                      Pushforward a distributive multiplicative action along a surjective additive monoid homomorphism. See note [reducible non-instances].

                                                                                      Equations
                                                                                      • One or more equations did not get rendered due to their size.
                                                                                      Instances For
                                                                                        @[reducible]
                                                                                        def Function.Surjective.distribMulActionLeft {R : Type u_10} {S : Type u_11} {M : Type u_12} [Monoid R] [AddMonoid M] [DistribMulAction R M] [Monoid S] [SMul S M] (f : R →* S) (hf : Function.Surjective f) (hsmul : ∀ (c : R) (x : M), f c x = c x) :

                                                                                        Push forward the action of R on M along a compatible surjective map f : R →* S.

                                                                                        See also Function.Surjective.mulActionLeft and Function.Surjective.moduleLeft.

                                                                                        Equations
                                                                                        • One or more equations did not get rendered due to their size.
                                                                                        Instances For
                                                                                          @[reducible]
                                                                                          def DistribMulAction.compHom {M : Type u_1} {N : Type u_2} (A : Type u_4) [Monoid M] [AddMonoid A] [DistribMulAction M A] [Monoid N] (f : N →* M) :

                                                                                          Compose a DistribMulAction with a MonoidHom, with action f r' • m. See note [reducible non-instances].

                                                                                          Equations
                                                                                          • One or more equations did not get rendered due to their size.
                                                                                          Instances For
                                                                                            @[simp]
                                                                                            theorem DistribMulAction.toAddMonoidHom_apply {M : Type u_1} (A : Type u_4) [Monoid M] [AddMonoid A] [DistribMulAction M A] (x : M) :
                                                                                            ∀ (x : A), ↑(DistribMulAction.toAddMonoidHom A x) x = x x
                                                                                            def DistribMulAction.toAddMonoidHom {M : Type u_1} (A : Type u_4) [Monoid M] [AddMonoid A] [DistribMulAction M A] (x : M) :
                                                                                            A →+ A

                                                                                            Each element of the monoid defines an additive monoid homomorphism.

                                                                                            Equations
                                                                                            Instances For

                                                                                              Each element of the monoid defines an additive monoid homomorphism.

                                                                                              Equations
                                                                                              • One or more equations did not get rendered due to their size.
                                                                                              Instances For
                                                                                                @[simp]
                                                                                                theorem smul_neg {M : Type u_1} {A : Type u_4} [Monoid M] [AddGroup A] [DistribMulAction M A] (r : M) (x : A) :
                                                                                                r -x = -(r x)
                                                                                                theorem smul_sub {M : Type u_1} {A : Type u_4} [Monoid M] [AddGroup A] [DistribMulAction M A] (r : M) (x : A) (y : A) :
                                                                                                r (x - y) = r x - r y
                                                                                                theorem MulDistribMulAction.ext {M : Type u_10} {A : Type u_11} :
                                                                                                ∀ {inst : Monoid M} {inst_1 : Monoid A} (x y : MulDistribMulAction M A), SMul.smul = SMul.smulx = y
                                                                                                theorem MulDistribMulAction.ext_iff {M : Type u_10} {A : Type u_11} :
                                                                                                ∀ {inst : Monoid M} {inst_1 : Monoid A} (x y : MulDistribMulAction M A), x = y SMul.smul = SMul.smul
                                                                                                class MulDistribMulAction (M : Type u_10) (A : Type u_11) [Monoid M] [Monoid A] extends MulAction :
                                                                                                Type (max u_10 u_11)
                                                                                                • smul : MAA
                                                                                                • one_smul : ∀ (b : A), 1 b = b
                                                                                                • mul_smul : ∀ (x y : M) (b : A), (x * y) b = x y b
                                                                                                • smul_mul : ∀ (r : M) (x y : A), r (x * y) = r x * r y

                                                                                                  Distributivity of across *

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

                                                                                                  Multiplying 1 by a scalar gives 1

                                                                                                Typeclass for multiplicative actions on multiplicative structures. This generalizes conjugation actions.

                                                                                                Instances
                                                                                                  theorem smul_mul' {M : Type u_1} {A : Type u_4} [Monoid M] [Monoid A] [MulDistribMulAction M A] (a : M) (b₁ : A) (b₂ : A) :
                                                                                                  a (b₁ * b₂) = a b₁ * a b₂
                                                                                                  @[reducible]
                                                                                                  def Function.Injective.mulDistribMulAction {M : Type u_1} {A : Type u_4} {B : Type u_5} [Monoid M] [Monoid A] [MulDistribMulAction M A] [Monoid B] [SMul M B] (f : B →* A) (hf : Function.Injective f) (smul : ∀ (c : M) (x : B), f (c x) = c f x) :

                                                                                                  Pullback a multiplicative distributive multiplicative action along an injective monoid homomorphism. See note [reducible non-instances].

                                                                                                  Equations
                                                                                                  • One or more equations did not get rendered due to their size.
                                                                                                  Instances For
                                                                                                    @[reducible]
                                                                                                    def Function.Surjective.mulDistribMulAction {M : Type u_1} {A : Type u_4} {B : Type u_5} [Monoid M] [Monoid A] [MulDistribMulAction M A] [Monoid B] [SMul M B] (f : A →* B) (hf : Function.Surjective f) (smul : ∀ (c : M) (x : A), f (c x) = c f x) :

                                                                                                    Pushforward a multiplicative distributive multiplicative action along a surjective monoid homomorphism. See note [reducible non-instances].

                                                                                                    Equations
                                                                                                    • One or more equations did not get rendered due to their size.
                                                                                                    Instances For
                                                                                                      @[reducible]
                                                                                                      def MulDistribMulAction.compHom {M : Type u_1} {N : Type u_2} (A : Type u_4) [Monoid M] [Monoid A] [MulDistribMulAction M A] [Monoid N] (f : N →* M) :

                                                                                                      Compose a MulDistribMulAction with a MonoidHom, with action f r' • m. See note [reducible non-instances].

                                                                                                      Equations
                                                                                                      Instances For
                                                                                                        def MulDistribMulAction.toMonoidHom {M : Type u_1} (A : Type u_4) [Monoid M] [Monoid A] [MulDistribMulAction M A] (r : M) :
                                                                                                        A →* A

                                                                                                        Scalar multiplication by r as a MonoidHom.

                                                                                                        Equations
                                                                                                        Instances For
                                                                                                          @[simp]
                                                                                                          theorem MulDistribMulAction.toMonoidHom_apply {M : Type u_1} {A : Type u_4} [Monoid M] [Monoid A] [MulDistribMulAction M A] (r : M) (x : A) :

                                                                                                          Each element of the monoid defines a monoid homomorphism.

                                                                                                          Equations
                                                                                                          • One or more equations did not get rendered due to their size.
                                                                                                          Instances For
                                                                                                            @[simp]
                                                                                                            theorem smul_inv' {M : Type u_1} {A : Type u_4} [Monoid M] [Group A] [MulDistribMulAction M A] (r : M) (x : A) :
                                                                                                            r x⁻¹ = (r x)⁻¹
                                                                                                            theorem smul_div' {M : Type u_1} {A : Type u_4} [Monoid M] [Group A] [MulDistribMulAction M A] (r : M) (x : A) (y : A) :
                                                                                                            r (x / y) = r x / r y
                                                                                                            def Function.End (α : Type u_6) :
                                                                                                            Type u_6

                                                                                                            The monoid of endomorphisms.

                                                                                                            Note that this is generalized by CategoryTheory.End to categories other than Type u.

                                                                                                            Equations
                                                                                                            Instances For
                                                                                                              instance instMonoidEnd (α : Type u_6) :
                                                                                                              Equations
                                                                                                              instance instInhabitedEnd (α : Type u_6) :
                                                                                                              Equations

                                                                                                              The tautological action by Function.End α on α.

                                                                                                              This is generalized to bundled endomorphisms by:

                                                                                                              Equations
                                                                                                              @[simp]
                                                                                                              theorem Function.End.smul_def {α : Type u_6} (f : Function.End α) (a : α) :
                                                                                                              f a = f a
                                                                                                              theorem Function.End.mul_def {α : Type u_6} (f : Function.End α) (g : Function.End α) :
                                                                                                              f * g = f g
                                                                                                              theorem Function.End.one_def {α : Type u_6} :
                                                                                                              1 = id

                                                                                                              The tautological action by AddMonoid.End α on α.

                                                                                                              This generalizes Function.End.applyMulAction.

                                                                                                              Equations
                                                                                                              • AddMonoid.End.applyDistribMulAction = DistribMulAction.mk (_ : ∀ (f : α →+ α), f 0 = 0) (_ : ∀ (f : α →+ α) (a b : α), f (a + b) = f a + f b)
                                                                                                              @[simp]
                                                                                                              theorem AddMonoid.End.smul_def {α : Type u_6} [AddMonoid α] (f : AddMonoid.End α) (a : α) :
                                                                                                              f a = f a
                                                                                                              def MulAction.toEndHom {M : Type u_1} {α : Type u_6} [Monoid M] [MulAction M α] :

                                                                                                              The monoid hom representing a monoid action.

                                                                                                              When M is a group, see MulAction.toPermHom.

                                                                                                              Equations
                                                                                                              • One or more equations did not get rendered due to their size.
                                                                                                              Instances For
                                                                                                                @[reducible]
                                                                                                                def MulAction.ofEndHom {M : Type u_1} {α : Type u_6} [Monoid M] (f : M →* Function.End α) :

                                                                                                                The monoid action induced by a monoid hom to Function.End α

                                                                                                                See note [reducible non-instances].

                                                                                                                Equations
                                                                                                                Instances For

                                                                                                                  Additive, Multiplicative #

                                                                                                                  instance Additive.vadd {α : Type u_6} {β : Type u_7} [SMul α β] :
                                                                                                                  VAdd (Additive α) β
                                                                                                                  Equations
                                                                                                                  • Additive.vadd = { vadd := fun a x => Additive.toMul a x }
                                                                                                                  instance Multiplicative.smul {α : Type u_6} {β : Type u_7} [VAdd α β] :
                                                                                                                  Equations
                                                                                                                  • Multiplicative.smul = { smul := fun a x => Multiplicative.toAdd a +ᵥ x }
                                                                                                                  @[simp]
                                                                                                                  theorem toMul_smul {α : Type u_6} {β : Type u_7} [SMul α β] (a : Additive α) (b : β) :
                                                                                                                  Additive.toMul a b = a +ᵥ b
                                                                                                                  @[simp]
                                                                                                                  theorem ofMul_vadd {α : Type u_6} {β : Type u_7} [SMul α β] (a : α) (b : β) :
                                                                                                                  Additive.ofMul a +ᵥ b = a b
                                                                                                                  @[simp]
                                                                                                                  theorem toAdd_vadd {α : Type u_6} {β : Type u_7} [VAdd α β] (a : Multiplicative α) (b : β) :
                                                                                                                  Multiplicative.toAdd a +ᵥ b = a b
                                                                                                                  @[simp]
                                                                                                                  theorem ofAdd_smul {α : Type u_6} {β : Type u_7} [VAdd α β] (a : α) (b : β) :
                                                                                                                  Multiplicative.ofAdd a b = a +ᵥ b
                                                                                                                  instance Additive.addAction {α : Type u_6} {β : Type u_7} [Monoid α] [MulAction α β] :
                                                                                                                  Equations
                                                                                                                  instance Multiplicative.mulAction {α : Type u_6} {β : Type u_7} [AddMonoid α] [AddAction α β] :
                                                                                                                  Equations
                                                                                                                  instance Additive.vaddCommClass {α : Type u_6} {β : Type u_7} {γ : Type u_8} [SMul α γ] [SMul β γ] [SMulCommClass α β γ] :
                                                                                                                  Equations

                                                                                                                  The tautological additive action by Additive (Function.End α) on α.

                                                                                                                  Equations
                                                                                                                  • AddAction.functionEnd = inferInstance
                                                                                                                  def AddAction.toEndHom {M : Type u_1} {α : Type u_6} [AddMonoid M] [AddAction M α] :

                                                                                                                  The additive monoid hom representing an additive monoid action.

                                                                                                                  When M is a group, see AddAction.toPermHom.

                                                                                                                  Equations
                                                                                                                  • AddAction.toEndHom = MonoidHom.toAdditive'' MulAction.toEndHom
                                                                                                                  Instances For
                                                                                                                    @[reducible]
                                                                                                                    def AddAction.ofEndHom {M : Type u_1} {α : Type u_6} [AddMonoid M] (f : M →+ Additive (Function.End α)) :

                                                                                                                    The additive action induced by a hom to Additive (Function.End α)

                                                                                                                    See note [reducible non-instances].

                                                                                                                    Equations
                                                                                                                    Instances For