Documentation

Mathlib.Algebra.Group.Hom.Defs

Monoid and group homomorphisms #

This file defines the bundled structures for monoid and group homomorphisms. Namely, we define MonoidHom (resp., AddMonoidHom) to be bundled homomorphisms between multiplicative (resp., additive) monoids or groups.

We also define coercion to a function, and usual operations: composition, identity homomorphism, pointwise multiplication and pointwise inversion.

This file also defines the lesser-used (and notation-less) homomorphism types which are used as building blocks for other homomorphisms:

Notations #

Implementation notes #

There's a coercion from bundled homs to fun, and the canonical notation is to use the bundled hom as a function via this coercion.

There is no GroupHom -- the idea is that MonoidHom is used. The constructor for MonoidHom needs a proof of map_one as well as map_mul; a separate constructor MonoidHom.mk' will construct group homs (i.e. monoid homs between groups) given only a proof that multiplication is preserved,

Implicit {} brackets are often used instead of type class [] brackets. This is done when the instances can be inferred because they are implicit arguments to the type MonoidHom. When they can be inferred from the type it is faster to use this method than to use type class inference.

Historically this file also included definitions of unbundled homomorphism classes; they were deprecated and moved to Deprecated/Group.

Tags #

MonoidHom, AddMonoidHom

structure ZeroHom (M : Type u_10) (N : Type u_11) [Zero M] [Zero N] :
Type (max u_10 u_11)

ZeroHom M N is the type of functions M → N that preserve zero.

When possible, instead of parametrizing results over (f : ZeroHom M N), you should parametrize over (F : Type*) [ZeroHomClass F M N] (f : F).

When you extend this structure, make sure to also extend ZeroHomClass.

  • toFun : MN

    The underlying function

  • map_zero' : self.toFun 0 = 0

    The proposition that the function preserves 0

Instances For
    class ZeroHomClass (F : Type u_10) (M : outParam (Type u_11)) (N : outParam (Type u_12)) [Zero M] [Zero N] [FunLike F M N] :

    ZeroHomClass F M N states that F is a type of zero-preserving homomorphisms.

    You should extend this typeclass when you extend ZeroHom.

    • map_zero (f : F) : f 0 = 0

      The proposition that the function preserves 0

    Instances
      structure AddHom (M : Type u_10) (N : Type u_11) [Add M] [Add N] :
      Type (max u_10 u_11)

      M →ₙ+ N is the type of functions M → N that preserve addition. The in the notation stands for "non-unital" because it is intended to match the notation for NonUnitalAlgHom and NonUnitalRingHom, so a AddHom is a non-unital additive monoid hom.

      When possible, instead of parametrizing results over (f : AddHom M N), you should parametrize over (F : Type*) [AddHomClass F M N] (f : F).

      When you extend this structure, make sure to extend AddHomClass.

      • toFun : MN

        The underlying function

      • map_add' (x y : M) : self.toFun (x + y) = self.toFun x + self.toFun y

        The proposition that the function preserves addition

      Instances For

        M →ₙ+ N denotes the type of addition-preserving maps from M to N.

        Equations
          Instances For
            class AddHomClass (F : Type u_10) (M : outParam (Type u_11)) (N : outParam (Type u_12)) [Add M] [Add N] [FunLike F M N] :

            AddHomClass F M N states that F is a type of addition-preserving homomorphisms. You should declare an instance of this typeclass when you extend AddHom.

            • map_add (f : F) (x y : M) : f (x + y) = f x + f y

              The proposition that the function preserves addition

            Instances
              structure AddMonoidHom (M : Type u_10) (N : Type u_11) [AddZeroClass M] [AddZeroClass N] extends ZeroHom M N, M →ₙ+ N :
              Type (max u_10 u_11)

              M →+ N is the type of functions M → N that preserve the AddZeroClass structure.

              AddMonoidHom is also used for group homomorphisms.

              When possible, instead of parametrizing results over (f : M →+ N), you should parametrize over (F : Type*) [AddMonoidHomClass F M N] (f : F).

              When you extend this structure, make sure to extend AddMonoidHomClass.

              Instances For

                M →+ N denotes the type of additive monoid homomorphisms from M to N.

                Equations
                  Instances For
                    class AddMonoidHomClass (F : Type u_10) (M : outParam (Type u_11)) (N : outParam (Type u_12)) [AddZeroClass M] [AddZeroClass N] [FunLike F M N] extends AddHomClass F M N, ZeroHomClass F M N :

                    AddMonoidHomClass F M N states that F is a type of AddZeroClass-preserving homomorphisms.

                    You should also extend this typeclass when you extend AddMonoidHom.

                    Instances
                      structure OneHom (M : Type u_10) (N : Type u_11) [One M] [One N] :
                      Type (max u_10 u_11)

                      OneHom M N is the type of functions M → N that preserve one.

                      When possible, instead of parametrizing results over (f : OneHom M N), you should parametrize over (F : Type*) [OneHomClass F M N] (f : F).

                      When you extend this structure, make sure to also extend OneHomClass.

                      • toFun : MN

                        The underlying function

                      • map_one' : self.toFun 1 = 1

                        The proposition that the function preserves 1

                      Instances For
                        class OneHomClass (F : Type u_10) (M : outParam (Type u_11)) (N : outParam (Type u_12)) [One M] [One N] [FunLike F M N] :

                        OneHomClass F M N states that F is a type of one-preserving homomorphisms. You should extend this typeclass when you extend OneHom.

                        • map_one (f : F) : f 1 = 1

                          The proposition that the function preserves 1

                        Instances
                          instance OneHom.funLike {M : Type u_4} {N : Type u_5} [One M] [One N] :
                          FunLike (OneHom M N) M N
                          Equations
                            instance ZeroHom.funLike {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] :
                            FunLike (ZeroHom M N) M N
                            Equations
                              instance OneHom.oneHomClass {M : Type u_4} {N : Type u_5} [One M] [One N] :
                              instance ZeroHom.zeroHomClass {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] :
                              @[simp]
                              theorem map_one {M : Type u_4} {N : Type u_5} {F : Type u_9} [One M] [One N] [FunLike F M N] [OneHomClass F M N] (f : F) :
                              f 1 = 1

                              See note [hom simp lemma priority]

                              @[simp]
                              theorem map_zero {M : Type u_4} {N : Type u_5} {F : Type u_9} [Zero M] [Zero N] [FunLike F M N] [ZeroHomClass F M N] (f : F) :
                              f 0 = 0
                              theorem map_comp_one {ι : Type u_1} {M : Type u_4} {N : Type u_5} {F : Type u_9} [One M] [One N] [FunLike F M N] [OneHomClass F M N] (f : F) :
                              f 1 = 1
                              theorem map_comp_zero {ι : Type u_1} {M : Type u_4} {N : Type u_5} {F : Type u_9} [Zero M] [Zero N] [FunLike F M N] [ZeroHomClass F M N] (f : F) :
                              f 0 = 0
                              theorem Subsingleton.of_oneHomClass {M : Type u_4} {N : Type u_5} {F : Type u_9} [One M] [One N] [FunLike F M N] [Subsingleton M] [OneHomClass F M N] :

                              In principle this could be an instance, but in practice it causes performance issues.

                              theorem Subsingleton.of_zeroHomClass {M : Type u_4} {N : Type u_5} {F : Type u_9} [Zero M] [Zero N] [FunLike F M N] [Subsingleton M] [ZeroHomClass F M N] :
                              instance instSubsingletonOneHom {M : Type u_4} {N : Type u_5} [One M] [One N] [Subsingleton M] :
                              instance instSubsingletonZeroHom {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] [Subsingleton M] :
                              theorem map_eq_one_iff {M : Type u_4} {N : Type u_5} {F : Type u_9} [One M] [One N] [FunLike F M N] [OneHomClass F M N] (f : F) (hf : Function.Injective f) {x : M} :
                              f x = 1 x = 1
                              theorem map_eq_zero_iff {M : Type u_4} {N : Type u_5} {F : Type u_9} [Zero M] [Zero N] [FunLike F M N] [ZeroHomClass F M N] (f : F) (hf : Function.Injective f) {x : M} :
                              f x = 0 x = 0
                              theorem map_ne_one_iff {R : Type u_10} {S : Type u_11} {F : Type u_12} [One R] [One S] [FunLike F R S] [OneHomClass F R S] (f : F) (hf : Function.Injective f) {x : R} :
                              f x 1 x 1
                              theorem map_ne_zero_iff {R : Type u_10} {S : Type u_11} {F : Type u_12} [Zero R] [Zero S] [FunLike F R S] [ZeroHomClass F R S] (f : F) (hf : Function.Injective f) {x : R} :
                              f x 0 x 0
                              theorem ne_one_of_map {R : Type u_10} {S : Type u_11} {F : Type u_12} [One R] [One S] [FunLike F R S] [OneHomClass F R S] {f : F} {x : R} (hx : f x 1) :
                              x 1
                              theorem ne_zero_of_map {R : Type u_10} {S : Type u_11} {F : Type u_12} [Zero R] [Zero S] [FunLike F R S] [ZeroHomClass F R S] {f : F} {x : R} (hx : f x 0) :
                              x 0
                              def OneHomClass.toOneHom {M : Type u_4} {N : Type u_5} {F : Type u_9} [One M] [One N] [FunLike F M N] [OneHomClass F M N] (f : F) :
                              OneHom M N

                              Turn an element of a type F satisfying OneHomClass F M N into an actual OneHom. This is declared as the default coercion from F to OneHom M N.

                              Equations
                                Instances For
                                  def ZeroHomClass.toZeroHom {M : Type u_4} {N : Type u_5} {F : Type u_9} [Zero M] [Zero N] [FunLike F M N] [ZeroHomClass F M N] (f : F) :

                                  Turn an element of a type F satisfying ZeroHomClass F M N into an actual ZeroHom. This is declared as the default coercion from F to ZeroHom M N.

                                  Equations
                                    Instances For
                                      instance instCoeTCOneHomOfOneHomClass {M : Type u_4} {N : Type u_5} {F : Type u_9} [One M] [One N] [FunLike F M N] [OneHomClass F M N] :
                                      CoeTC F (OneHom M N)

                                      Any type satisfying OneHomClass can be cast into OneHom via OneHomClass.toOneHom.

                                      Equations
                                        instance instCoeTCZeroHomOfZeroHomClass {M : Type u_4} {N : Type u_5} {F : Type u_9} [Zero M] [Zero N] [FunLike F M N] [ZeroHomClass F M N] :
                                        CoeTC F (ZeroHom M N)

                                        Any type satisfying ZeroHomClass can be cast into ZeroHom via ZeroHomClass.toZeroHom.

                                        Equations
                                          @[simp]
                                          theorem OneHom.coe_coe {M : Type u_4} {N : Type u_5} {F : Type u_9} [One M] [One N] [FunLike F M N] [OneHomClass F M N] (f : F) :
                                          f = f
                                          @[simp]
                                          theorem ZeroHom.coe_coe {M : Type u_4} {N : Type u_5} {F : Type u_9} [Zero M] [Zero N] [FunLike F M N] [ZeroHomClass F M N] (f : F) :
                                          f = f
                                          structure MulHom (M : Type u_10) (N : Type u_11) [Mul M] [Mul N] :
                                          Type (max u_10 u_11)

                                          M →ₙ* N is the type of functions M → N that preserve multiplication. The in the notation stands for "non-unital" because it is intended to match the notation for NonUnitalAlgHom and NonUnitalRingHom, so a MulHom is a non-unital monoid hom.

                                          When possible, instead of parametrizing results over (f : M →ₙ* N), you should parametrize over (F : Type*) [MulHomClass F M N] (f : F). When you extend this structure, make sure to extend MulHomClass.

                                          • toFun : MN

                                            The underlying function

                                          • map_mul' (x y : M) : self.toFun (x * y) = self.toFun x * self.toFun y

                                            The proposition that the function preserves multiplication

                                          Instances For

                                            M →ₙ* N denotes the type of multiplication-preserving maps from M to N.

                                            Equations
                                              Instances For
                                                class MulHomClass (F : Type u_10) (M : outParam (Type u_11)) (N : outParam (Type u_12)) [Mul M] [Mul N] [FunLike F M N] :

                                                MulHomClass F M N states that F is a type of multiplication-preserving homomorphisms.

                                                You should declare an instance of this typeclass when you extend MulHom.

                                                • map_mul (f : F) (x y : M) : f (x * y) = f x * f y

                                                  The proposition that the function preserves multiplication

                                                Instances
                                                  instance MulHom.funLike {M : Type u_4} {N : Type u_5} [Mul M] [Mul N] :
                                                  FunLike (M →ₙ* N) M N
                                                  Equations
                                                    instance AddHom.funLike {M : Type u_4} {N : Type u_5} [Add M] [Add N] :
                                                    FunLike (M →ₙ+ N) M N
                                                    Equations
                                                      instance MulHom.mulHomClass {M : Type u_4} {N : Type u_5} [Mul M] [Mul N] :

                                                      MulHom is a type of multiplication-preserving homomorphisms

                                                      instance AddHom.addHomClass {M : Type u_4} {N : Type u_5} [Add M] [Add N] :

                                                      AddHom is a type of addition-preserving homomorphisms

                                                      @[simp]
                                                      theorem map_mul {M : Type u_4} {N : Type u_5} {F : Type u_9} [Mul M] [Mul N] [FunLike F M N] [MulHomClass F M N] (f : F) (x y : M) :
                                                      f (x * y) = f x * f y

                                                      See note [hom simp lemma priority]

                                                      @[simp]
                                                      theorem map_add {M : Type u_4} {N : Type u_5} {F : Type u_9} [Add M] [Add N] [FunLike F M N] [AddHomClass F M N] (f : F) (x y : M) :
                                                      f (x + y) = f x + f y
                                                      @[simp]
                                                      theorem map_comp_mul {ι : Type u_1} {M : Type u_4} {N : Type u_5} {F : Type u_9} [Mul M] [Mul N] [FunLike F M N] [MulHomClass F M N] (f : F) (g h : ιM) :
                                                      f (g * h) = f g * f h
                                                      @[simp]
                                                      theorem map_comp_add {ι : Type u_1} {M : Type u_4} {N : Type u_5} {F : Type u_9} [Add M] [Add N] [FunLike F M N] [AddHomClass F M N] (f : F) (g h : ιM) :
                                                      f (g + h) = f g + f h
                                                      def MulHomClass.toMulHom {M : Type u_4} {N : Type u_5} {F : Type u_9} [Mul M] [Mul N] [FunLike F M N] [MulHomClass F M N] (f : F) :

                                                      Turn an element of a type F satisfying MulHomClass F M N into an actual MulHom. This is declared as the default coercion from F to M →ₙ* N.

                                                      Equations
                                                        Instances For
                                                          def AddHomClass.toAddHom {M : Type u_4} {N : Type u_5} {F : Type u_9} [Add M] [Add N] [FunLike F M N] [AddHomClass F M N] (f : F) :

                                                          Turn an element of a type F satisfying AddHomClass F M N into an actual AddHom. This is declared as the default coercion from F to M →ₙ+ N.

                                                          Equations
                                                            Instances For
                                                              instance instCoeTCMulHomOfMulHomClass {M : Type u_4} {N : Type u_5} {F : Type u_9} [Mul M] [Mul N] [FunLike F M N] [MulHomClass F M N] :
                                                              CoeTC F (M →ₙ* N)

                                                              Any type satisfying MulHomClass can be cast into MulHom via MulHomClass.toMulHom.

                                                              Equations
                                                                instance instCoeTCAddHomOfAddHomClass {M : Type u_4} {N : Type u_5} {F : Type u_9} [Add M] [Add N] [FunLike F M N] [AddHomClass F M N] :
                                                                CoeTC F (M →ₙ+ N)

                                                                Any type satisfying AddHomClass can be cast into AddHom via AddHomClass.toAddHom.

                                                                Equations
                                                                  @[simp]
                                                                  theorem MulHom.coe_coe {M : Type u_4} {N : Type u_5} {F : Type u_9} [Mul M] [Mul N] [FunLike F M N] [MulHomClass F M N] (f : F) :
                                                                  f = f
                                                                  @[simp]
                                                                  theorem AddHom.coe_coe {M : Type u_4} {N : Type u_5} {F : Type u_9} [Add M] [Add N] [FunLike F M N] [AddHomClass F M N] (f : F) :
                                                                  f = f
                                                                  structure MonoidHom (M : Type u_10) (N : Type u_11) [MulOneClass M] [MulOneClass N] extends OneHom M N, M →ₙ* N :
                                                                  Type (max u_10 u_11)

                                                                  M →* N is the type of functions M → N that preserve the Monoid structure. MonoidHom is also used for group homomorphisms.

                                                                  When possible, instead of parametrizing results over (f : M →* N), you should parametrize over (F : Type*) [MonoidHomClass F M N] (f : F).

                                                                  When you extend this structure, make sure to extend MonoidHomClass.

                                                                  Instances For

                                                                    M →* N denotes the type of monoid homomorphisms from M to N.

                                                                    Equations
                                                                      Instances For
                                                                        class MonoidHomClass (F : Type u_10) (M : outParam (Type u_11)) (N : outParam (Type u_12)) [MulOneClass M] [MulOneClass N] [FunLike F M N] extends MulHomClass F M N, OneHomClass F M N :

                                                                        MonoidHomClass F M N states that F is a type of Monoid-preserving homomorphisms. You should also extend this typeclass when you extend MonoidHom.

                                                                        Instances
                                                                          instance MonoidHom.instFunLike {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] :
                                                                          FunLike (M →* N) M N
                                                                          Equations
                                                                            instance AddMonoidHom.instFunLike {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] :
                                                                            FunLike (M →+ N) M N
                                                                            Equations
                                                                              def MonoidHomClass.toMonoidHom {M : Type u_4} {N : Type u_5} {F : Type u_9} [MulOneClass M] [MulOneClass N] [FunLike F M N] [MonoidHomClass F M N] (f : F) :
                                                                              M →* N

                                                                              Turn an element of a type F satisfying MonoidHomClass F M N into an actual MonoidHom. This is declared as the default coercion from F to M →* N.

                                                                              Equations
                                                                                Instances For
                                                                                  def AddMonoidHomClass.toAddMonoidHom {M : Type u_4} {N : Type u_5} {F : Type u_9} [AddZeroClass M] [AddZeroClass N] [FunLike F M N] [AddMonoidHomClass F M N] (f : F) :
                                                                                  M →+ N

                                                                                  Turn an element of a type F satisfying AddMonoidHomClass F M N into an actual MonoidHom. This is declared as the default coercion from F to M →+ N.

                                                                                  Equations
                                                                                    Instances For
                                                                                      instance instCoeTCMonoidHomOfMonoidHomClass {M : Type u_4} {N : Type u_5} {F : Type u_9} [MulOneClass M] [MulOneClass N] [FunLike F M N] [MonoidHomClass F M N] :
                                                                                      CoeTC F (M →* N)

                                                                                      Any type satisfying MonoidHomClass can be cast into MonoidHom via MonoidHomClass.toMonoidHom.

                                                                                      Equations
                                                                                        instance instCoeTCAddMonoidHomOfAddMonoidHomClass {M : Type u_4} {N : Type u_5} {F : Type u_9} [AddZeroClass M] [AddZeroClass N] [FunLike F M N] [AddMonoidHomClass F M N] :
                                                                                        CoeTC F (M →+ N)

                                                                                        Any type satisfying AddMonoidHomClass can be cast into AddMonoidHom via AddMonoidHomClass.toAddMonoidHom.

                                                                                        Equations
                                                                                          @[simp]
                                                                                          theorem MonoidHom.coe_coe {M : Type u_4} {N : Type u_5} {F : Type u_9} [MulOneClass M] [MulOneClass N] [FunLike F M N] [MonoidHomClass F M N] (f : F) :
                                                                                          f = f
                                                                                          @[simp]
                                                                                          theorem AddMonoidHom.coe_coe {M : Type u_4} {N : Type u_5} {F : Type u_9} [AddZeroClass M] [AddZeroClass N] [FunLike F M N] [AddMonoidHomClass F M N] (f : F) :
                                                                                          f = f
                                                                                          theorem map_mul_eq_one {M : Type u_4} {N : Type u_5} {F : Type u_9} [MulOneClass M] [MulOneClass N] [FunLike F M N] [MonoidHomClass F M N] (f : F) {a b : M} (h : a * b = 1) :
                                                                                          f a * f b = 1
                                                                                          theorem map_add_eq_zero {M : Type u_4} {N : Type u_5} {F : Type u_9} [AddZeroClass M] [AddZeroClass N] [FunLike F M N] [AddMonoidHomClass F M N] (f : F) {a b : M} (h : a + b = 0) :
                                                                                          f a + f b = 0
                                                                                          theorem map_div' {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [DivInvMonoid G] [DivInvMonoid H] [MulHomClass F G H] (f : F) (hf : ∀ (a : G), f a⁻¹ = (f a)⁻¹) (a b : G) :
                                                                                          f (a / b) = f a / f b
                                                                                          theorem map_sub' {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [SubNegMonoid G] [SubNegMonoid H] [AddHomClass F G H] (f : F) (hf : ∀ (a : G), f (-a) = -f a) (a b : G) :
                                                                                          f (a - b) = f a - f b
                                                                                          theorem map_comp_div' {ι : Type u_1} {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [DivInvMonoid G] [DivInvMonoid H] [MulHomClass F G H] (f : F) (hf : ∀ (a : G), f a⁻¹ = (f a)⁻¹) (g h : ιG) :
                                                                                          f (g / h) = f g / f h
                                                                                          theorem map_comp_sub' {ι : Type u_1} {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [SubNegMonoid G] [SubNegMonoid H] [AddHomClass F G H] (f : F) (hf : ∀ (a : G), f (-a) = -f a) (g h : ιG) :
                                                                                          f (g - h) = f g - f h
                                                                                          @[simp]
                                                                                          theorem map_inv {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [Group G] [DivisionMonoid H] [MonoidHomClass F G H] (f : F) (a : G) :
                                                                                          f a⁻¹ = (f a)⁻¹

                                                                                          Group homomorphisms preserve inverse.

                                                                                          See note [hom simp lemma priority]

                                                                                          @[simp]
                                                                                          theorem map_neg {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [AddGroup G] [SubtractionMonoid H] [AddMonoidHomClass F G H] (f : F) (a : G) :
                                                                                          f (-a) = -f a

                                                                                          Additive group homomorphisms preserve negation.

                                                                                          @[simp]
                                                                                          theorem map_comp_inv {ι : Type u_1} {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [Group G] [DivisionMonoid H] [MonoidHomClass F G H] (f : F) (g : ιG) :
                                                                                          f g⁻¹ = (f g)⁻¹
                                                                                          @[simp]
                                                                                          theorem map_comp_neg {ι : Type u_1} {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [AddGroup G] [SubtractionMonoid H] [AddMonoidHomClass F G H] (f : F) (g : ιG) :
                                                                                          f (-g) = -f g
                                                                                          theorem map_mul_inv {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [Group G] [DivisionMonoid H] [MonoidHomClass F G H] (f : F) (a b : G) :
                                                                                          f (a * b⁻¹) = f a * (f b)⁻¹

                                                                                          Group homomorphisms preserve division.

                                                                                          theorem map_add_neg {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [AddGroup G] [SubtractionMonoid H] [AddMonoidHomClass F G H] (f : F) (a b : G) :
                                                                                          f (a + -b) = f a + -f b

                                                                                          Additive group homomorphisms preserve subtraction.

                                                                                          theorem map_comp_mul_inv {ι : Type u_1} {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [Group G] [DivisionMonoid H] [MonoidHomClass F G H] (f : F) (g h : ιG) :
                                                                                          f (g * h⁻¹) = f g * (f h)⁻¹
                                                                                          theorem map_comp_add_neg {ι : Type u_1} {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [AddGroup G] [SubtractionMonoid H] [AddMonoidHomClass F G H] (f : F) (g h : ιG) :
                                                                                          f (g + -h) = f g + -f h
                                                                                          @[simp]
                                                                                          theorem map_div {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [Group G] [DivisionMonoid H] [MonoidHomClass F G H] (f : F) (a b : G) :
                                                                                          f (a / b) = f a / f b

                                                                                          Group homomorphisms preserve division.

                                                                                          See note [hom simp lemma priority]

                                                                                          @[simp]
                                                                                          theorem map_sub {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [AddGroup G] [SubtractionMonoid H] [AddMonoidHomClass F G H] (f : F) (a b : G) :
                                                                                          f (a - b) = f a - f b

                                                                                          Additive group homomorphisms preserve subtraction.

                                                                                          @[simp]
                                                                                          theorem map_comp_div {ι : Type u_1} {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [Group G] [DivisionMonoid H] [MonoidHomClass F G H] (f : F) (g h : ιG) :
                                                                                          f (g / h) = f g / f h
                                                                                          @[simp]
                                                                                          theorem map_comp_sub {ι : Type u_1} {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [AddGroup G] [SubtractionMonoid H] [AddMonoidHomClass F G H] (f : F) (g h : ιG) :
                                                                                          f (g - h) = f g - f h
                                                                                          @[simp]
                                                                                          theorem map_pow {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [Monoid G] [Monoid H] [MonoidHomClass F G H] (f : F) (a : G) (n : ) :
                                                                                          f (a ^ n) = f a ^ n

                                                                                          See note [hom simp lemma priority]

                                                                                          @[simp]
                                                                                          theorem map_nsmul {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [AddMonoid G] [AddMonoid H] [AddMonoidHomClass F G H] (f : F) (n : ) (a : G) :
                                                                                          f (n a) = n f a
                                                                                          @[simp]
                                                                                          theorem map_comp_pow {ι : Type u_1} {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [Monoid G] [Monoid H] [MonoidHomClass F G H] (f : F) (g : ιG) (n : ) :
                                                                                          f (g ^ n) = f g ^ n
                                                                                          @[simp]
                                                                                          theorem map_comp_nsmul {ι : Type u_1} {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [AddMonoid G] [AddMonoid H] [AddMonoidHomClass F G H] (f : F) (g : ιG) (n : ) :
                                                                                          f (n g) = n f g
                                                                                          theorem map_zpow' {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [DivInvMonoid G] [DivInvMonoid H] [MonoidHomClass F G H] (f : F) (hf : ∀ (x : G), f x⁻¹ = (f x)⁻¹) (a : G) (n : ) :
                                                                                          f (a ^ n) = f a ^ n
                                                                                          theorem map_zsmul' {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [SubNegMonoid G] [SubNegMonoid H] [AddMonoidHomClass F G H] (f : F) (hf : ∀ (x : G), f (-x) = -f x) (a : G) (n : ) :
                                                                                          f (n a) = n f a
                                                                                          @[simp]
                                                                                          theorem map_comp_zpow' {ι : Type u_1} {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [DivInvMonoid G] [DivInvMonoid H] [MonoidHomClass F G H] (f : F) (hf : ∀ (x : G), f x⁻¹ = (f x)⁻¹) (g : ιG) (n : ) :
                                                                                          f (g ^ n) = f g ^ n
                                                                                          @[simp]
                                                                                          theorem map_comp_zsmul' {ι : Type u_1} {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [SubNegMonoid G] [SubNegMonoid H] [AddMonoidHomClass F G H] (f : F) (hf : ∀ (x : G), f (-x) = -f x) (g : ιG) (n : ) :
                                                                                          f (n g) = n f g
                                                                                          @[simp]
                                                                                          theorem map_zpow {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [Group G] [DivisionMonoid H] [MonoidHomClass F G H] (f : F) (g : G) (n : ) :
                                                                                          f (g ^ n) = f g ^ n

                                                                                          Group homomorphisms preserve integer power.

                                                                                          See note [hom simp lemma priority]

                                                                                          @[simp]
                                                                                          theorem map_zsmul {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [AddGroup G] [SubtractionMonoid H] [AddMonoidHomClass F G H] (f : F) (n : ) (g : G) :
                                                                                          f (n g) = n f g

                                                                                          Additive group homomorphisms preserve integer scaling.

                                                                                          theorem map_comp_zpow {ι : Type u_1} {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [Group G] [DivisionMonoid H] [MonoidHomClass F G H] (f : F) (g : ιG) (n : ) :
                                                                                          f (g ^ n) = f g ^ n
                                                                                          theorem map_comp_zsmul {ι : Type u_1} {G : Type u_7} {H : Type u_8} {F : Type u_9} [FunLike F G H] [AddGroup G] [SubtractionMonoid H] [AddMonoidHomClass F G H] (f : F) (g : ιG) (n : ) :
                                                                                          f (n g) = n f g

                                                                                          Bundled morphisms can be down-cast to weaker bundlings

                                                                                          instance MonoidHom.coeToOneHom {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] :
                                                                                          Coe (M →* N) (OneHom M N)

                                                                                          MonoidHom down-cast to a OneHom, forgetting the multiplicative property.

                                                                                          Equations
                                                                                            instance AddMonoidHom.coeToZeroHom {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] :
                                                                                            Coe (M →+ N) (ZeroHom M N)

                                                                                            AddMonoidHom down-cast to a ZeroHom, forgetting the additive property

                                                                                            Equations
                                                                                              instance MonoidHom.coeToMulHom {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] :
                                                                                              Coe (M →* N) (M →ₙ* N)

                                                                                              MonoidHom down-cast to a MulHom, forgetting the 1-preserving property.

                                                                                              Equations
                                                                                                instance AddMonoidHom.coeToAddHom {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] :
                                                                                                Coe (M →+ N) (M →ₙ+ N)

                                                                                                AddMonoidHom down-cast to an AddHom, forgetting the 0-preserving property.

                                                                                                Equations
                                                                                                  @[simp]
                                                                                                  theorem OneHom.coe_mk {M : Type u_4} {N : Type u_5} [One M] [One N] (f : MN) (h1 : f 1 = 1) :
                                                                                                  { toFun := f, map_one' := h1 } = f
                                                                                                  @[simp]
                                                                                                  theorem ZeroHom.coe_mk {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] (f : MN) (h1 : f 0 = 0) :
                                                                                                  { toFun := f, map_zero' := h1 } = f
                                                                                                  @[simp]
                                                                                                  theorem OneHom.toFun_eq_coe {M : Type u_4} {N : Type u_5} [One M] [One N] (f : OneHom M N) :
                                                                                                  f.toFun = f
                                                                                                  @[simp]
                                                                                                  theorem ZeroHom.toFun_eq_coe {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] (f : ZeroHom M N) :
                                                                                                  f.toFun = f
                                                                                                  @[simp]
                                                                                                  theorem MulHom.coe_mk {M : Type u_4} {N : Type u_5} [Mul M] [Mul N] (f : MN) (hmul : ∀ (x y : M), f (x * y) = f x * f y) :
                                                                                                  { toFun := f, map_mul' := hmul } = f
                                                                                                  @[simp]
                                                                                                  theorem AddHom.coe_mk {M : Type u_4} {N : Type u_5} [Add M] [Add N] (f : MN) (hmul : ∀ (x y : M), f (x + y) = f x + f y) :
                                                                                                  { toFun := f, map_add' := hmul } = f
                                                                                                  @[simp]
                                                                                                  theorem MulHom.toFun_eq_coe {M : Type u_4} {N : Type u_5} [Mul M] [Mul N] (f : M →ₙ* N) :
                                                                                                  f.toFun = f
                                                                                                  @[simp]
                                                                                                  theorem AddHom.toFun_eq_coe {M : Type u_4} {N : Type u_5} [Add M] [Add N] (f : M →ₙ+ N) :
                                                                                                  f.toFun = f
                                                                                                  @[simp]
                                                                                                  theorem MonoidHom.coe_mk {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] (f : OneHom M N) (hmul : ∀ (x y : M), f.toFun (x * y) = f.toFun x * f.toFun y) :
                                                                                                  { toOneHom := f, map_mul' := hmul } = f
                                                                                                  @[simp]
                                                                                                  theorem AddMonoidHom.coe_mk {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] (f : ZeroHom M N) (hmul : ∀ (x y : M), f.toFun (x + y) = f.toFun x + f.toFun y) :
                                                                                                  { toZeroHom := f, map_add' := hmul } = f
                                                                                                  @[simp]
                                                                                                  theorem MonoidHom.toOneHom_coe {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] (f : M →* N) :
                                                                                                  f = f
                                                                                                  @[simp]
                                                                                                  theorem AddMonoidHom.toZeroHom_coe {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) :
                                                                                                  f = f
                                                                                                  @[simp]
                                                                                                  theorem MonoidHom.toMulHom_coe {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] (f : M →* N) :
                                                                                                  (↑f).toFun = f
                                                                                                  @[simp]
                                                                                                  theorem AddMonoidHom.toAddHom_coe {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) :
                                                                                                  (↑f).toFun = f
                                                                                                  theorem MonoidHom.toFun_eq_coe {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] (f : M →* N) :
                                                                                                  (↑f).toFun = f
                                                                                                  theorem AddMonoidHom.toFun_eq_coe {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) :
                                                                                                  (↑f).toFun = f
                                                                                                  theorem OneHom.ext {M : Type u_4} {N : Type u_5} [One M] [One N] f g : OneHom M N (h : ∀ (x : M), f x = g x) :
                                                                                                  f = g
                                                                                                  theorem ZeroHom.ext {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] f g : ZeroHom M N (h : ∀ (x : M), f x = g x) :
                                                                                                  f = g
                                                                                                  theorem ZeroHom.ext_iff {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] {f g : ZeroHom M N} :
                                                                                                  f = g ∀ (x : M), f x = g x
                                                                                                  theorem OneHom.ext_iff {M : Type u_4} {N : Type u_5} [One M] [One N] {f g : OneHom M N} :
                                                                                                  f = g ∀ (x : M), f x = g x
                                                                                                  theorem MulHom.ext {M : Type u_4} {N : Type u_5} [Mul M] [Mul N] f g : M →ₙ* N (h : ∀ (x : M), f x = g x) :
                                                                                                  f = g
                                                                                                  theorem AddHom.ext {M : Type u_4} {N : Type u_5} [Add M] [Add N] f g : M →ₙ+ N (h : ∀ (x : M), f x = g x) :
                                                                                                  f = g
                                                                                                  theorem MulHom.ext_iff {M : Type u_4} {N : Type u_5} [Mul M] [Mul N] {f g : M →ₙ* N} :
                                                                                                  f = g ∀ (x : M), f x = g x
                                                                                                  theorem AddHom.ext_iff {M : Type u_4} {N : Type u_5} [Add M] [Add N] {f g : M →ₙ+ N} :
                                                                                                  f = g ∀ (x : M), f x = g x
                                                                                                  theorem MonoidHom.ext {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] f g : M →* N (h : ∀ (x : M), f x = g x) :
                                                                                                  f = g
                                                                                                  theorem AddMonoidHom.ext {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] f g : M →+ N (h : ∀ (x : M), f x = g x) :
                                                                                                  f = g
                                                                                                  theorem MonoidHom.ext_iff {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] {f g : M →* N} :
                                                                                                  f = g ∀ (x : M), f x = g x
                                                                                                  theorem AddMonoidHom.ext_iff {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] {f g : M →+ N} :
                                                                                                  f = g ∀ (x : M), f x = g x
                                                                                                  def MonoidHom.mk' {M : Type u_4} {G : Type u_7} [Group G] [MulOneClass M] (f : MG) (map_mul : ∀ (a b : M), f (a * b) = f a * f b) :
                                                                                                  M →* G

                                                                                                  Makes a group homomorphism from a proof that the map preserves multiplication.

                                                                                                  Equations
                                                                                                    Instances For
                                                                                                      def AddMonoidHom.mk' {M : Type u_4} {G : Type u_7} [AddGroup G] [AddZeroClass M] (f : MG) (map_mul : ∀ (a b : M), f (a + b) = f a + f b) :
                                                                                                      M →+ G

                                                                                                      Makes an additive group homomorphism from a proof that the map preserves addition.

                                                                                                      Equations
                                                                                                        Instances For
                                                                                                          @[simp]
                                                                                                          theorem AddMonoidHom.mk'_apply {M : Type u_4} {G : Type u_7} [AddGroup G] [AddZeroClass M] (f : MG) (map_mul : ∀ (a b : M), f (a + b) = f a + f b) :
                                                                                                          (mk' f map_mul) = f
                                                                                                          @[simp]
                                                                                                          theorem MonoidHom.mk'_apply {M : Type u_4} {G : Type u_7} [Group G] [MulOneClass M] (f : MG) (map_mul : ∀ (a b : M), f (a * b) = f a * f b) :
                                                                                                          (mk' f map_mul) = f
                                                                                                          @[simp]
                                                                                                          theorem OneHom.mk_coe {M : Type u_4} {N : Type u_5} [One M] [One N] (f : OneHom M N) (h1 : f 1 = 1) :
                                                                                                          { toFun := f, map_one' := h1 } = f
                                                                                                          @[simp]
                                                                                                          theorem ZeroHom.mk_coe {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] (f : ZeroHom M N) (h1 : f 0 = 0) :
                                                                                                          { toFun := f, map_zero' := h1 } = f
                                                                                                          @[simp]
                                                                                                          theorem MulHom.mk_coe {M : Type u_4} {N : Type u_5} [Mul M] [Mul N] (f : M →ₙ* N) (hmul : ∀ (x y : M), f (x * y) = f x * f y) :
                                                                                                          { toFun := f, map_mul' := hmul } = f
                                                                                                          @[simp]
                                                                                                          theorem AddHom.mk_coe {M : Type u_4} {N : Type u_5} [Add M] [Add N] (f : M →ₙ+ N) (hmul : ∀ (x y : M), f (x + y) = f x + f y) :
                                                                                                          { toFun := f, map_add' := hmul } = f
                                                                                                          @[simp]
                                                                                                          theorem MonoidHom.mk_coe {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] (f : M →* N) (hmul : ∀ (x y : M), (↑f).toFun (x * y) = (↑f).toFun x * (↑f).toFun y) :
                                                                                                          { toOneHom := f, map_mul' := hmul } = f
                                                                                                          @[simp]
                                                                                                          theorem AddMonoidHom.mk_coe {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) (hmul : ∀ (x y : M), (↑f).toFun (x + y) = (↑f).toFun x + (↑f).toFun y) :
                                                                                                          { toZeroHom := f, map_add' := hmul } = f
                                                                                                          def OneHom.copy {M : Type u_4} {N : Type u_5} [One M] [One N] (f : OneHom M N) (f' : MN) (h : f' = f) :
                                                                                                          OneHom M N

                                                                                                          Copy of a OneHom with a new toFun equal to the old one. Useful to fix definitional equalities.

                                                                                                          Equations
                                                                                                            Instances For
                                                                                                              def ZeroHom.copy {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] (f : ZeroHom M N) (f' : MN) (h : f' = f) :

                                                                                                              Copy of a ZeroHom with a new toFun equal to the old one. Useful to fix definitional equalities.

                                                                                                              Equations
                                                                                                                Instances For
                                                                                                                  @[simp]
                                                                                                                  theorem OneHom.coe_copy {M : Type u_4} {N : Type u_5} {x✝ : One M} {x✝¹ : One N} (f : OneHom M N) (f' : MN) (h : f' = f) :
                                                                                                                  (f.copy f' h) = f'
                                                                                                                  @[simp]
                                                                                                                  theorem ZeroHom.coe_copy {M : Type u_4} {N : Type u_5} {x✝ : Zero M} {x✝¹ : Zero N} (f : ZeroHom M N) (f' : MN) (h : f' = f) :
                                                                                                                  (f.copy f' h) = f'
                                                                                                                  theorem OneHom.coe_copy_eq {M : Type u_4} {N : Type u_5} {x✝ : One M} {x✝¹ : One N} (f : OneHom M N) (f' : MN) (h : f' = f) :
                                                                                                                  f.copy f' h = f
                                                                                                                  theorem ZeroHom.coe_copy_eq {M : Type u_4} {N : Type u_5} {x✝ : Zero M} {x✝¹ : Zero N} (f : ZeroHom M N) (f' : MN) (h : f' = f) :
                                                                                                                  f.copy f' h = f
                                                                                                                  def MulHom.copy {M : Type u_4} {N : Type u_5} [Mul M] [Mul N] (f : M →ₙ* N) (f' : MN) (h : f' = f) :

                                                                                                                  Copy of a MulHom with a new toFun equal to the old one. Useful to fix definitional equalities.

                                                                                                                  Equations
                                                                                                                    Instances For
                                                                                                                      def AddHom.copy {M : Type u_4} {N : Type u_5} [Add M] [Add N] (f : M →ₙ+ N) (f' : MN) (h : f' = f) :

                                                                                                                      Copy of an AddHom with a new toFun equal to the old one. Useful to fix definitional equalities.

                                                                                                                      Equations
                                                                                                                        Instances For
                                                                                                                          @[simp]
                                                                                                                          theorem MulHom.coe_copy {M : Type u_4} {N : Type u_5} {x✝ : Mul M} {x✝¹ : Mul N} (f : M →ₙ* N) (f' : MN) (h : f' = f) :
                                                                                                                          (f.copy f' h) = f'
                                                                                                                          @[simp]
                                                                                                                          theorem AddHom.coe_copy {M : Type u_4} {N : Type u_5} {x✝ : Add M} {x✝¹ : Add N} (f : M →ₙ+ N) (f' : MN) (h : f' = f) :
                                                                                                                          (f.copy f' h) = f'
                                                                                                                          theorem MulHom.coe_copy_eq {M : Type u_4} {N : Type u_5} {x✝ : Mul M} {x✝¹ : Mul N} (f : M →ₙ* N) (f' : MN) (h : f' = f) :
                                                                                                                          f.copy f' h = f
                                                                                                                          theorem AddHom.coe_copy_eq {M : Type u_4} {N : Type u_5} {x✝ : Add M} {x✝¹ : Add N} (f : M →ₙ+ N) (f' : MN) (h : f' = f) :
                                                                                                                          f.copy f' h = f
                                                                                                                          def MonoidHom.copy {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] (f : M →* N) (f' : MN) (h : f' = f) :
                                                                                                                          M →* N

                                                                                                                          Copy of a MonoidHom with a new toFun equal to the old one. Useful to fix definitional equalities.

                                                                                                                          Equations
                                                                                                                            Instances For
                                                                                                                              def AddMonoidHom.copy {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) (f' : MN) (h : f' = f) :
                                                                                                                              M →+ N

                                                                                                                              Copy of an AddMonoidHom with a new toFun equal to the old one. Useful to fix definitional equalities.

                                                                                                                              Equations
                                                                                                                                Instances For
                                                                                                                                  @[simp]
                                                                                                                                  theorem MonoidHom.coe_copy {M : Type u_4} {N : Type u_5} {x✝ : MulOneClass M} {x✝¹ : MulOneClass N} (f : M →* N) (f' : MN) (h : f' = f) :
                                                                                                                                  (f.copy f' h) = f'
                                                                                                                                  @[simp]
                                                                                                                                  theorem AddMonoidHom.coe_copy {M : Type u_4} {N : Type u_5} {x✝ : AddZeroClass M} {x✝¹ : AddZeroClass N} (f : M →+ N) (f' : MN) (h : f' = f) :
                                                                                                                                  (f.copy f' h) = f'
                                                                                                                                  theorem MonoidHom.copy_eq {M : Type u_4} {N : Type u_5} {x✝ : MulOneClass M} {x✝¹ : MulOneClass N} (f : M →* N) (f' : MN) (h : f' = f) :
                                                                                                                                  f.copy f' h = f
                                                                                                                                  theorem AddMonoidHom.copy_eq {M : Type u_4} {N : Type u_5} {x✝ : AddZeroClass M} {x✝¹ : AddZeroClass N} (f : M →+ N) (f' : MN) (h : f' = f) :
                                                                                                                                  f.copy f' h = f
                                                                                                                                  theorem OneHom.map_one {M : Type u_4} {N : Type u_5} [One M] [One N] (f : OneHom M N) :
                                                                                                                                  f 1 = 1
                                                                                                                                  theorem ZeroHom.map_zero {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] (f : ZeroHom M N) :
                                                                                                                                  f 0 = 0
                                                                                                                                  theorem MonoidHom.map_one {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] (f : M →* N) :
                                                                                                                                  f 1 = 1

                                                                                                                                  If f is a monoid homomorphism then f 1 = 1.

                                                                                                                                  theorem AddMonoidHom.map_zero {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) :
                                                                                                                                  f 0 = 0

                                                                                                                                  If f is an additive monoid homomorphism then f 0 = 0.

                                                                                                                                  theorem MulHom.map_mul {M : Type u_4} {N : Type u_5} [Mul M] [Mul N] (f : M →ₙ* N) (a b : M) :
                                                                                                                                  f (a * b) = f a * f b
                                                                                                                                  theorem AddHom.map_add {M : Type u_4} {N : Type u_5} [Add M] [Add N] (f : M →ₙ+ N) (a b : M) :
                                                                                                                                  f (a + b) = f a + f b
                                                                                                                                  theorem MonoidHom.map_mul {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] (f : M →* N) (a b : M) :
                                                                                                                                  f (a * b) = f a * f b

                                                                                                                                  If f is a monoid homomorphism then f (a * b) = f a * f b.

                                                                                                                                  theorem AddMonoidHom.map_add {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) (a b : M) :
                                                                                                                                  f (a + b) = f a + f b

                                                                                                                                  If f is an additive monoid homomorphism then f (a + b) = f a + f b.

                                                                                                                                  theorem MonoidHom.map_exists_right_inv {M : Type u_4} {N : Type u_5} {F : Type u_9} [MulOneClass M] [MulOneClass N] [FunLike F M N] [MonoidHomClass F M N] (f : F) {x : M} (hx : (y : M), x * y = 1) :
                                                                                                                                  (y : N), f x * y = 1

                                                                                                                                  Given a monoid homomorphism f : M →* N and an element x : M, if x has a right inverse, then f x has a right inverse too. For elements invertible on both sides see IsUnit.map.

                                                                                                                                  theorem AddMonoidHom.map_exists_right_neg {M : Type u_4} {N : Type u_5} {F : Type u_9} [AddZeroClass M] [AddZeroClass N] [FunLike F M N] [AddMonoidHomClass F M N] (f : F) {x : M} (hx : (y : M), x + y = 0) :
                                                                                                                                  (y : N), f x + y = 0

                                                                                                                                  Given an AddMonoid homomorphism f : M →+ N and an element x : M, if x has a right inverse, then f x has a right inverse too.

                                                                                                                                  theorem MonoidHom.map_exists_left_inv {M : Type u_4} {N : Type u_5} {F : Type u_9} [MulOneClass M] [MulOneClass N] [FunLike F M N] [MonoidHomClass F M N] (f : F) {x : M} (hx : (y : M), y * x = 1) :
                                                                                                                                  (y : N), y * f x = 1

                                                                                                                                  Given a monoid homomorphism f : M →* N and an element x : M, if x has a left inverse, then f x has a left inverse too. For elements invertible on both sides see IsUnit.map.

                                                                                                                                  theorem AddMonoidHom.map_exists_left_neg {M : Type u_4} {N : Type u_5} {F : Type u_9} [AddZeroClass M] [AddZeroClass N] [FunLike F M N] [AddMonoidHomClass F M N] (f : F) {x : M} (hx : (y : M), y + x = 0) :
                                                                                                                                  (y : N), y + f x = 0

                                                                                                                                  Given an AddMonoid homomorphism f : M →+ N and an element x : M, if x has a left inverse, then f x has a left inverse too. For elements invertible on both sides see IsAddUnit.map.

                                                                                                                                  def OneHom.id (M : Type u_10) [One M] :
                                                                                                                                  OneHom M M

                                                                                                                                  The identity map from a type with 1 to itself.

                                                                                                                                  Equations
                                                                                                                                    Instances For
                                                                                                                                      def ZeroHom.id (M : Type u_10) [Zero M] :

                                                                                                                                      The identity map from a type with zero to itself.

                                                                                                                                      Equations
                                                                                                                                        Instances For
                                                                                                                                          @[simp]
                                                                                                                                          theorem ZeroHom.id_apply (M : Type u_10) [Zero M] (x : M) :
                                                                                                                                          (id M) x = x
                                                                                                                                          @[simp]
                                                                                                                                          theorem OneHom.id_apply (M : Type u_10) [One M] (x : M) :
                                                                                                                                          (id M) x = x
                                                                                                                                          def MulHom.id (M : Type u_10) [Mul M] :

                                                                                                                                          The identity map from a type with multiplication to itself.

                                                                                                                                          Equations
                                                                                                                                            Instances For
                                                                                                                                              def AddHom.id (M : Type u_10) [Add M] :

                                                                                                                                              The identity map from a type with addition to itself.

                                                                                                                                              Equations
                                                                                                                                                Instances For
                                                                                                                                                  @[simp]
                                                                                                                                                  theorem MulHom.id_apply (M : Type u_10) [Mul M] (x : M) :
                                                                                                                                                  (id M) x = x
                                                                                                                                                  @[simp]
                                                                                                                                                  theorem AddHom.id_apply (M : Type u_10) [Add M] (x : M) :
                                                                                                                                                  (id M) x = x
                                                                                                                                                  def MonoidHom.id (M : Type u_10) [MulOneClass M] :
                                                                                                                                                  M →* M

                                                                                                                                                  The identity map from a monoid to itself.

                                                                                                                                                  Equations
                                                                                                                                                    Instances For
                                                                                                                                                      def AddMonoidHom.id (M : Type u_10) [AddZeroClass M] :
                                                                                                                                                      M →+ M

                                                                                                                                                      The identity map from an additive monoid to itself.

                                                                                                                                                      Equations
                                                                                                                                                        Instances For
                                                                                                                                                          @[simp]
                                                                                                                                                          theorem MonoidHom.id_apply (M : Type u_10) [MulOneClass M] (x : M) :
                                                                                                                                                          (id M) x = x
                                                                                                                                                          @[simp]
                                                                                                                                                          theorem AddMonoidHom.id_apply (M : Type u_10) [AddZeroClass M] (x : M) :
                                                                                                                                                          (id M) x = x
                                                                                                                                                          @[simp]
                                                                                                                                                          theorem OneHom.coe_id {M : Type u_10} [One M] :
                                                                                                                                                          (id M) = _root_.id
                                                                                                                                                          @[simp]
                                                                                                                                                          theorem ZeroHom.coe_id {M : Type u_10} [Zero M] :
                                                                                                                                                          (id M) = _root_.id
                                                                                                                                                          @[simp]
                                                                                                                                                          theorem MulHom.coe_id {M : Type u_10} [Mul M] :
                                                                                                                                                          (id M) = _root_.id
                                                                                                                                                          @[simp]
                                                                                                                                                          theorem AddHom.coe_id {M : Type u_10} [Add M] :
                                                                                                                                                          (id M) = _root_.id
                                                                                                                                                          @[simp]
                                                                                                                                                          theorem MonoidHom.coe_id {M : Type u_10} [MulOneClass M] :
                                                                                                                                                          (id M) = _root_.id
                                                                                                                                                          @[simp]
                                                                                                                                                          theorem AddMonoidHom.coe_id {M : Type u_10} [AddZeroClass M] :
                                                                                                                                                          (id M) = _root_.id
                                                                                                                                                          def OneHom.comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [One M] [One N] [One P] (hnp : OneHom N P) (hmn : OneHom M N) :
                                                                                                                                                          OneHom M P

                                                                                                                                                          Composition of OneHoms as a OneHom.

                                                                                                                                                          Equations
                                                                                                                                                            Instances For
                                                                                                                                                              def ZeroHom.comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [Zero M] [Zero N] [Zero P] (hnp : ZeroHom N P) (hmn : ZeroHom M N) :

                                                                                                                                                              Composition of ZeroHoms as a ZeroHom.

                                                                                                                                                              Equations
                                                                                                                                                                Instances For
                                                                                                                                                                  def MulHom.comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [Mul M] [Mul N] [Mul P] (hnp : N →ₙ* P) (hmn : M →ₙ* N) :

                                                                                                                                                                  Composition of MulHoms as a MulHom.

                                                                                                                                                                  Equations
                                                                                                                                                                    Instances For
                                                                                                                                                                      def AddHom.comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [Add M] [Add N] [Add P] (hnp : N →ₙ+ P) (hmn : M →ₙ+ N) :

                                                                                                                                                                      Composition of AddHoms as an AddHom.

                                                                                                                                                                      Equations
                                                                                                                                                                        Instances For
                                                                                                                                                                          def MonoidHom.comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [MulOneClass M] [MulOneClass N] [MulOneClass P] (hnp : N →* P) (hmn : M →* N) :
                                                                                                                                                                          M →* P

                                                                                                                                                                          Composition of monoid morphisms as a monoid morphism.

                                                                                                                                                                          Equations
                                                                                                                                                                            Instances For
                                                                                                                                                                              def AddMonoidHom.comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [AddZeroClass M] [AddZeroClass N] [AddZeroClass P] (hnp : N →+ P) (hmn : M →+ N) :
                                                                                                                                                                              M →+ P

                                                                                                                                                                              Composition of additive monoid morphisms as an additive monoid morphism.

                                                                                                                                                                              Equations
                                                                                                                                                                                Instances For
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem OneHom.coe_comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [One M] [One N] [One P] (g : OneHom N P) (f : OneHom M N) :
                                                                                                                                                                                  (g.comp f) = g f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem ZeroHom.coe_comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [Zero M] [Zero N] [Zero P] (g : ZeroHom N P) (f : ZeroHom M N) :
                                                                                                                                                                                  (g.comp f) = g f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem MulHom.coe_comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [Mul M] [Mul N] [Mul P] (g : N →ₙ* P) (f : M →ₙ* N) :
                                                                                                                                                                                  (g.comp f) = g f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem AddHom.coe_comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [Add M] [Add N] [Add P] (g : N →ₙ+ P) (f : M →ₙ+ N) :
                                                                                                                                                                                  (g.comp f) = g f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem MonoidHom.coe_comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [MulOneClass M] [MulOneClass N] [MulOneClass P] (g : N →* P) (f : M →* N) :
                                                                                                                                                                                  (g.comp f) = g f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem AddMonoidHom.coe_comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [AddZeroClass M] [AddZeroClass N] [AddZeroClass P] (g : N →+ P) (f : M →+ N) :
                                                                                                                                                                                  (g.comp f) = g f
                                                                                                                                                                                  theorem OneHom.comp_apply {M : Type u_4} {N : Type u_5} {P : Type u_6} [One M] [One N] [One P] (g : OneHom N P) (f : OneHom M N) (x : M) :
                                                                                                                                                                                  (g.comp f) x = g (f x)
                                                                                                                                                                                  theorem ZeroHom.comp_apply {M : Type u_4} {N : Type u_5} {P : Type u_6} [Zero M] [Zero N] [Zero P] (g : ZeroHom N P) (f : ZeroHom M N) (x : M) :
                                                                                                                                                                                  (g.comp f) x = g (f x)
                                                                                                                                                                                  theorem MulHom.comp_apply {M : Type u_4} {N : Type u_5} {P : Type u_6} [Mul M] [Mul N] [Mul P] (g : N →ₙ* P) (f : M →ₙ* N) (x : M) :
                                                                                                                                                                                  (g.comp f) x = g (f x)
                                                                                                                                                                                  theorem AddHom.comp_apply {M : Type u_4} {N : Type u_5} {P : Type u_6} [Add M] [Add N] [Add P] (g : N →ₙ+ P) (f : M →ₙ+ N) (x : M) :
                                                                                                                                                                                  (g.comp f) x = g (f x)
                                                                                                                                                                                  theorem MonoidHom.comp_apply {M : Type u_4} {N : Type u_5} {P : Type u_6} [MulOneClass M] [MulOneClass N] [MulOneClass P] (g : N →* P) (f : M →* N) (x : M) :
                                                                                                                                                                                  (g.comp f) x = g (f x)
                                                                                                                                                                                  theorem AddMonoidHom.comp_apply {M : Type u_4} {N : Type u_5} {P : Type u_6} [AddZeroClass M] [AddZeroClass N] [AddZeroClass P] (g : N →+ P) (f : M →+ N) (x : M) :
                                                                                                                                                                                  (g.comp f) x = g (f x)
                                                                                                                                                                                  theorem OneHom.comp_assoc {M : Type u_4} {N : Type u_5} {P : Type u_6} {Q : Type u_10} [One M] [One N] [One P] [One Q] (f : OneHom M N) (g : OneHom N P) (h : OneHom P Q) :
                                                                                                                                                                                  (h.comp g).comp f = h.comp (g.comp f)

                                                                                                                                                                                  Composition of monoid homomorphisms is associative.

                                                                                                                                                                                  theorem ZeroHom.comp_assoc {M : Type u_4} {N : Type u_5} {P : Type u_6} {Q : Type u_10} [Zero M] [Zero N] [Zero P] [Zero Q] (f : ZeroHom M N) (g : ZeroHom N P) (h : ZeroHom P Q) :
                                                                                                                                                                                  (h.comp g).comp f = h.comp (g.comp f)

                                                                                                                                                                                  Composition of additive monoid homomorphisms is associative.

                                                                                                                                                                                  theorem MulHom.comp_assoc {M : Type u_4} {N : Type u_5} {P : Type u_6} {Q : Type u_10} [Mul M] [Mul N] [Mul P] [Mul Q] (f : M →ₙ* N) (g : N →ₙ* P) (h : P →ₙ* Q) :
                                                                                                                                                                                  (h.comp g).comp f = h.comp (g.comp f)
                                                                                                                                                                                  theorem AddHom.comp_assoc {M : Type u_4} {N : Type u_5} {P : Type u_6} {Q : Type u_10} [Add M] [Add N] [Add P] [Add Q] (f : M →ₙ+ N) (g : N →ₙ+ P) (h : P →ₙ+ Q) :
                                                                                                                                                                                  (h.comp g).comp f = h.comp (g.comp f)
                                                                                                                                                                                  theorem MonoidHom.comp_assoc {M : Type u_4} {N : Type u_5} {P : Type u_6} {Q : Type u_10} [MulOneClass M] [MulOneClass N] [MulOneClass P] [MulOneClass Q] (f : M →* N) (g : N →* P) (h : P →* Q) :
                                                                                                                                                                                  (h.comp g).comp f = h.comp (g.comp f)
                                                                                                                                                                                  theorem AddMonoidHom.comp_assoc {M : Type u_4} {N : Type u_5} {P : Type u_6} {Q : Type u_10} [AddZeroClass M] [AddZeroClass N] [AddZeroClass P] [AddZeroClass Q] (f : M →+ N) (g : N →+ P) (h : P →+ Q) :
                                                                                                                                                                                  (h.comp g).comp f = h.comp (g.comp f)
                                                                                                                                                                                  theorem OneHom.cancel_right {M : Type u_4} {N : Type u_5} {P : Type u_6} [One M] [One N] [One P] {g₁ g₂ : OneHom N P} {f : OneHom M N} (hf : Function.Surjective f) :
                                                                                                                                                                                  g₁.comp f = g₂.comp f g₁ = g₂
                                                                                                                                                                                  theorem ZeroHom.cancel_right {M : Type u_4} {N : Type u_5} {P : Type u_6} [Zero M] [Zero N] [Zero P] {g₁ g₂ : ZeroHom N P} {f : ZeroHom M N} (hf : Function.Surjective f) :
                                                                                                                                                                                  g₁.comp f = g₂.comp f g₁ = g₂
                                                                                                                                                                                  theorem MulHom.cancel_right {M : Type u_4} {N : Type u_5} {P : Type u_6} [Mul M] [Mul N] [Mul P] {g₁ g₂ : N →ₙ* P} {f : M →ₙ* N} (hf : Function.Surjective f) :
                                                                                                                                                                                  g₁.comp f = g₂.comp f g₁ = g₂
                                                                                                                                                                                  theorem AddHom.cancel_right {M : Type u_4} {N : Type u_5} {P : Type u_6} [Add M] [Add N] [Add P] {g₁ g₂ : N →ₙ+ P} {f : M →ₙ+ N} (hf : Function.Surjective f) :
                                                                                                                                                                                  g₁.comp f = g₂.comp f g₁ = g₂
                                                                                                                                                                                  theorem MonoidHom.cancel_right {M : Type u_4} {N : Type u_5} {P : Type u_6} [MulOneClass M] [MulOneClass N] [MulOneClass P] {g₁ g₂ : N →* P} {f : M →* N} (hf : Function.Surjective f) :
                                                                                                                                                                                  g₁.comp f = g₂.comp f g₁ = g₂
                                                                                                                                                                                  theorem AddMonoidHom.cancel_right {M : Type u_4} {N : Type u_5} {P : Type u_6} [AddZeroClass M] [AddZeroClass N] [AddZeroClass P] {g₁ g₂ : N →+ P} {f : M →+ N} (hf : Function.Surjective f) :
                                                                                                                                                                                  g₁.comp f = g₂.comp f g₁ = g₂
                                                                                                                                                                                  theorem OneHom.cancel_left {M : Type u_4} {N : Type u_5} {P : Type u_6} [One M] [One N] [One P] {g : OneHom N P} {f₁ f₂ : OneHom M N} (hg : Function.Injective g) :
                                                                                                                                                                                  g.comp f₁ = g.comp f₂ f₁ = f₂
                                                                                                                                                                                  theorem ZeroHom.cancel_left {M : Type u_4} {N : Type u_5} {P : Type u_6} [Zero M] [Zero N] [Zero P] {g : ZeroHom N P} {f₁ f₂ : ZeroHom M N} (hg : Function.Injective g) :
                                                                                                                                                                                  g.comp f₁ = g.comp f₂ f₁ = f₂
                                                                                                                                                                                  theorem MulHom.cancel_left {M : Type u_4} {N : Type u_5} {P : Type u_6} [Mul M] [Mul N] [Mul P] {g : N →ₙ* P} {f₁ f₂ : M →ₙ* N} (hg : Function.Injective g) :
                                                                                                                                                                                  g.comp f₁ = g.comp f₂ f₁ = f₂
                                                                                                                                                                                  theorem AddHom.cancel_left {M : Type u_4} {N : Type u_5} {P : Type u_6} [Add M] [Add N] [Add P] {g : N →ₙ+ P} {f₁ f₂ : M →ₙ+ N} (hg : Function.Injective g) :
                                                                                                                                                                                  g.comp f₁ = g.comp f₂ f₁ = f₂
                                                                                                                                                                                  theorem MonoidHom.cancel_left {M : Type u_4} {N : Type u_5} {P : Type u_6} [MulOneClass M] [MulOneClass N] [MulOneClass P] {g : N →* P} {f₁ f₂ : M →* N} (hg : Function.Injective g) :
                                                                                                                                                                                  g.comp f₁ = g.comp f₂ f₁ = f₂
                                                                                                                                                                                  theorem AddMonoidHom.cancel_left {M : Type u_4} {N : Type u_5} {P : Type u_6} [AddZeroClass M] [AddZeroClass N] [AddZeroClass P] {g : N →+ P} {f₁ f₂ : M →+ N} (hg : Function.Injective g) :
                                                                                                                                                                                  g.comp f₁ = g.comp f₂ f₁ = f₂
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem OneHom.comp_id {M : Type u_4} {N : Type u_5} [One M] [One N] (f : OneHom M N) :
                                                                                                                                                                                  f.comp (id M) = f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem ZeroHom.comp_id {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] (f : ZeroHom M N) :
                                                                                                                                                                                  f.comp (id M) = f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem MulHom.comp_id {M : Type u_4} {N : Type u_5} [Mul M] [Mul N] (f : M →ₙ* N) :
                                                                                                                                                                                  f.comp (id M) = f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem AddHom.comp_id {M : Type u_4} {N : Type u_5} [Add M] [Add N] (f : M →ₙ+ N) :
                                                                                                                                                                                  f.comp (id M) = f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem MonoidHom.comp_id {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] (f : M →* N) :
                                                                                                                                                                                  f.comp (id M) = f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem AddMonoidHom.comp_id {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) :
                                                                                                                                                                                  f.comp (id M) = f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem OneHom.id_comp {M : Type u_4} {N : Type u_5} [One M] [One N] (f : OneHom M N) :
                                                                                                                                                                                  (id N).comp f = f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem ZeroHom.id_comp {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] (f : ZeroHom M N) :
                                                                                                                                                                                  (id N).comp f = f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem MulHom.id_comp {M : Type u_4} {N : Type u_5} [Mul M] [Mul N] (f : M →ₙ* N) :
                                                                                                                                                                                  (id N).comp f = f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem AddHom.id_comp {M : Type u_4} {N : Type u_5} [Add M] [Add N] (f : M →ₙ+ N) :
                                                                                                                                                                                  (id N).comp f = f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem MonoidHom.id_comp {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] (f : M →* N) :
                                                                                                                                                                                  (id N).comp f = f
                                                                                                                                                                                  @[simp]
                                                                                                                                                                                  theorem AddMonoidHom.id_comp {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) :
                                                                                                                                                                                  (id N).comp f = f
                                                                                                                                                                                  theorem MonoidHom.map_pow {M : Type u_4} {N : Type u_5} [Monoid M] [Monoid N] (f : M →* N) (a : M) (n : ) :
                                                                                                                                                                                  f (a ^ n) = f a ^ n
                                                                                                                                                                                  theorem AddMonoidHom.map_nsmul {M : Type u_4} {N : Type u_5} [AddMonoid M] [AddMonoid N] (f : M →+ N) (a : M) (n : ) :
                                                                                                                                                                                  f (n a) = n f a
                                                                                                                                                                                  theorem MonoidHom.map_zpow' {M : Type u_4} {N : Type u_5} [DivInvMonoid M] [DivInvMonoid N] (f : M →* N) (hf : ∀ (x : M), f x⁻¹ = (f x)⁻¹) (a : M) (n : ) :
                                                                                                                                                                                  f (a ^ n) = f a ^ n
                                                                                                                                                                                  theorem AddMonoidHom.map_zsmul' {M : Type u_4} {N : Type u_5} [SubNegMonoid M] [SubNegMonoid N] (f : M →+ N) (hf : ∀ (x : M), f (-x) = -f x) (a : M) (n : ) :
                                                                                                                                                                                  f (n a) = n f a
                                                                                                                                                                                  def OneHom.inverse {M : Type u_4} {N : Type u_5} [One M] [One N] (f : OneHom M N) (g : NM) (h₁ : Function.LeftInverse g f) :
                                                                                                                                                                                  OneHom N M

                                                                                                                                                                                  Makes a OneHom inverse from the bijective inverse of a OneHom

                                                                                                                                                                                  Equations
                                                                                                                                                                                    Instances For
                                                                                                                                                                                      def ZeroHom.inverse {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] (f : ZeroHom M N) (g : NM) (h₁ : Function.LeftInverse g f) :

                                                                                                                                                                                      Make a ZeroHom inverse from the bijective inverse of a ZeroHom

                                                                                                                                                                                      Equations
                                                                                                                                                                                        Instances For
                                                                                                                                                                                          @[simp]
                                                                                                                                                                                          theorem OneHom.inverse_apply {M : Type u_4} {N : Type u_5} [One M] [One N] (f : OneHom M N) (g : NM) (h₁ : Function.LeftInverse g f) (a✝ : N) :
                                                                                                                                                                                          (f.inverse g h₁) a✝ = g a✝
                                                                                                                                                                                          @[simp]
                                                                                                                                                                                          theorem ZeroHom.inverse_apply {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] (f : ZeroHom M N) (g : NM) (h₁ : Function.LeftInverse g f) (a✝ : N) :
                                                                                                                                                                                          (f.inverse g h₁) a✝ = g a✝
                                                                                                                                                                                          def MulHom.inverse {M : Type u_4} {N : Type u_5} [Mul M] [Mul N] (f : M →ₙ* N) (g : NM) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) :

                                                                                                                                                                                          Makes a multiplicative inverse from a bijection which preserves multiplication.

                                                                                                                                                                                          Equations
                                                                                                                                                                                            Instances For
                                                                                                                                                                                              def AddHom.inverse {M : Type u_4} {N : Type u_5} [Add M] [Add N] (f : M →ₙ+ N) (g : NM) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) :

                                                                                                                                                                                              Makes an additive inverse from a bijection which preserves addition.

                                                                                                                                                                                              Equations
                                                                                                                                                                                                Instances For
                                                                                                                                                                                                  @[simp]
                                                                                                                                                                                                  theorem MulHom.inverse_apply {M : Type u_4} {N : Type u_5} [Mul M] [Mul N] (f : M →ₙ* N) (g : NM) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) (a✝ : N) :
                                                                                                                                                                                                  (f.inverse g h₁ h₂) a✝ = g a✝
                                                                                                                                                                                                  @[simp]
                                                                                                                                                                                                  theorem AddHom.inverse_apply {M : Type u_4} {N : Type u_5} [Add M] [Add N] (f : M →ₙ+ N) (g : NM) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) (a✝ : N) :
                                                                                                                                                                                                  (f.inverse g h₁ h₂) a✝ = g a✝
                                                                                                                                                                                                  theorem Function.Surjective.mul_comm {M : Type u_4} {N : Type u_5} [Mul M] [Mul N] {f : M →ₙ* N} (is_surj : Surjective f) (is_comm : Std.Commutative fun (x1 x2 : M) => x1 * x2) :
                                                                                                                                                                                                  Std.Commutative fun (x1 x2 : N) => x1 * x2

                                                                                                                                                                                                  If M and N have multiplications, f : M →ₙ* N is a surjective multiplicative map, and M is commutative, then N is commutative.

                                                                                                                                                                                                  theorem Function.Surjective.add_comm {M : Type u_4} {N : Type u_5} [Add M] [Add N] {f : M →ₙ+ N} (is_surj : Surjective f) (is_comm : Std.Commutative fun (x1 x2 : M) => x1 + x2) :
                                                                                                                                                                                                  Std.Commutative fun (x1 x2 : N) => x1 + x2

                                                                                                                                                                                                  If M and N have additions, f : M →ₙ+ N is a surjective additive map, and M is commutative, then N is commutative.

                                                                                                                                                                                                  def MonoidHom.inverse {A : Type u_10} {B : Type u_11} [Monoid A] [Monoid B] (f : A →* B) (g : BA) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) :
                                                                                                                                                                                                  B →* A

                                                                                                                                                                                                  The inverse of a bijective MonoidHom is a MonoidHom.

                                                                                                                                                                                                  Equations
                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                      def AddMonoidHom.inverse {A : Type u_10} {B : Type u_11} [AddMonoid A] [AddMonoid B] (f : A →+ B) (g : BA) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) :
                                                                                                                                                                                                      B →+ A

                                                                                                                                                                                                      The inverse of a bijective AddMonoidHom is an AddMonoidHom.

                                                                                                                                                                                                      Equations
                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                          @[simp]
                                                                                                                                                                                                          theorem AddMonoidHom.inverse_apply {A : Type u_10} {B : Type u_11} [AddMonoid A] [AddMonoid B] (f : A →+ B) (g : BA) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) (a✝ : B) :
                                                                                                                                                                                                          (f.inverse g h₁ h₂) a✝ = g a✝
                                                                                                                                                                                                          @[simp]
                                                                                                                                                                                                          theorem MonoidHom.inverse_apply {A : Type u_10} {B : Type u_11} [Monoid A] [Monoid B] (f : A →* B) (g : BA) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) (a✝ : B) :
                                                                                                                                                                                                          (f.inverse g h₁ h₂) a✝ = g a✝
                                                                                                                                                                                                          def Monoid.End (M : Type u_4) [MulOneClass M] :
                                                                                                                                                                                                          Type u_4

                                                                                                                                                                                                          The monoid of endomorphisms.

                                                                                                                                                                                                          Equations
                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                              def AddMonoid.End (M : Type u_4) [AddZeroClass M] :
                                                                                                                                                                                                              Type u_4

                                                                                                                                                                                                              The monoid of endomorphisms.

                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                  instance Monoid.End.instFunLike (M : Type u_4) [MulOneClass M] :
                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                      instance Monoid.End.instOne (M : Type u_4) [MulOneClass M] :
                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                          instance Monoid.End.instMul (M : Type u_4) [MulOneClass M] :
                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                @[simp]
                                                                                                                                                                                                                                theorem Monoid.End.coe_pow (M : Type u_4) [MulOneClass M] (f : Monoid.End M) (n : ) :
                                                                                                                                                                                                                                ⇑(f ^ n) = (⇑f)^[n]
                                                                                                                                                                                                                                @[simp]
                                                                                                                                                                                                                                theorem AddMonoid.End.coe_pow (M : Type u_4) [AddZeroClass M] (f : AddMonoid.End M) (n : ) :
                                                                                                                                                                                                                                ⇑(f ^ n) = (⇑f)^[n]
                                                                                                                                                                                                                                @[simp]
                                                                                                                                                                                                                                theorem Monoid.End.coe_one (M : Type u_4) [MulOneClass M] :
                                                                                                                                                                                                                                1 = id
                                                                                                                                                                                                                                @[simp]
                                                                                                                                                                                                                                theorem AddMonoid.End.coe_one (M : Type u_4) [AddZeroClass M] :
                                                                                                                                                                                                                                1 = id
                                                                                                                                                                                                                                @[simp]
                                                                                                                                                                                                                                theorem Monoid.End.coe_mul (M : Type u_4) [MulOneClass M] (f g : Monoid.End M) :
                                                                                                                                                                                                                                ⇑(f * g) = f g
                                                                                                                                                                                                                                @[simp]
                                                                                                                                                                                                                                theorem AddMonoid.End.coe_mul (M : Type u_4) [AddZeroClass M] (f g : AddMonoid.End M) :
                                                                                                                                                                                                                                ⇑(f * g) = f g
                                                                                                                                                                                                                                @[deprecated Monoid.End.coe_one (since := "2024-11-20")]
                                                                                                                                                                                                                                theorem Monoid.coe_one (M : Type u_4) [MulOneClass M] :
                                                                                                                                                                                                                                1 = id

                                                                                                                                                                                                                                Alias of Monoid.End.coe_one.

                                                                                                                                                                                                                                @[deprecated Monoid.End.coe_mul (since := "2024-11-20")]
                                                                                                                                                                                                                                theorem Monoid.coe_mul (M : Type u_4) [MulOneClass M] (f g : Monoid.End M) :
                                                                                                                                                                                                                                ⇑(f * g) = f g

                                                                                                                                                                                                                                Alias of Monoid.End.coe_mul.

                                                                                                                                                                                                                                instance instOneOneHom {M : Type u_4} {N : Type u_5} [One M] [One N] :
                                                                                                                                                                                                                                One (OneHom M N)

                                                                                                                                                                                                                                1 is the homomorphism sending all elements to 1.

                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                  instance instZeroZeroHom {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] :

                                                                                                                                                                                                                                  0 is the homomorphism sending all elements to 0.

                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                    instance instOneMulHom {M : Type u_4} {N : Type u_5} [Mul M] [MulOneClass N] :
                                                                                                                                                                                                                                    One (M →ₙ* N)

                                                                                                                                                                                                                                    1 is the multiplicative homomorphism sending all elements to 1.

                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                      instance instZeroAddHom {M : Type u_4} {N : Type u_5} [Add M] [AddZeroClass N] :

                                                                                                                                                                                                                                      0 is the additive homomorphism sending all elements to 0

                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                        instance instOneMonoidHom {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] :
                                                                                                                                                                                                                                        One (M →* N)

                                                                                                                                                                                                                                        1 is the monoid homomorphism sending all elements to 1.

                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                          instance instZeroAddMonoidHom {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] :
                                                                                                                                                                                                                                          Zero (M →+ N)

                                                                                                                                                                                                                                          0 is the additive monoid homomorphism sending all elements to 0.

                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                            @[simp]
                                                                                                                                                                                                                                            theorem OneHom.one_apply {M : Type u_4} {N : Type u_5} [One M] [One N] (x : M) :
                                                                                                                                                                                                                                            1 x = 1
                                                                                                                                                                                                                                            @[simp]
                                                                                                                                                                                                                                            theorem ZeroHom.zero_apply {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] (x : M) :
                                                                                                                                                                                                                                            0 x = 0
                                                                                                                                                                                                                                            @[simp]
                                                                                                                                                                                                                                            theorem MonoidHom.one_apply {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] (x : M) :
                                                                                                                                                                                                                                            1 x = 1
                                                                                                                                                                                                                                            @[simp]
                                                                                                                                                                                                                                            theorem AddMonoidHom.zero_apply {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] (x : M) :
                                                                                                                                                                                                                                            0 x = 0
                                                                                                                                                                                                                                            @[simp]
                                                                                                                                                                                                                                            theorem OneHom.one_comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [One M] [One N] [One P] (f : OneHom M N) :
                                                                                                                                                                                                                                            comp 1 f = 1
                                                                                                                                                                                                                                            @[simp]
                                                                                                                                                                                                                                            theorem ZeroHom.zero_comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [Zero M] [Zero N] [Zero P] (f : ZeroHom M N) :
                                                                                                                                                                                                                                            comp 0 f = 0
                                                                                                                                                                                                                                            @[simp]
                                                                                                                                                                                                                                            theorem OneHom.comp_one {M : Type u_4} {N : Type u_5} {P : Type u_6} [One M] [One N] [One P] (f : OneHom N P) :
                                                                                                                                                                                                                                            f.comp 1 = 1
                                                                                                                                                                                                                                            @[simp]
                                                                                                                                                                                                                                            theorem ZeroHom.comp_zero {M : Type u_4} {N : Type u_5} {P : Type u_6} [Zero M] [Zero N] [Zero P] (f : ZeroHom N P) :
                                                                                                                                                                                                                                            f.comp 0 = 0
                                                                                                                                                                                                                                            instance instInhabitedOneHom {M : Type u_4} {N : Type u_5} [One M] [One N] :
                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                              instance instInhabitedZeroHom {M : Type u_4} {N : Type u_5} [Zero M] [Zero N] :
                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                instance instInhabitedMulHom {M : Type u_4} {N : Type u_5} [Mul M] [MulOneClass N] :
                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                  instance instInhabitedAddHom {M : Type u_4} {N : Type u_5} [Add M] [AddZeroClass N] :
                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                    instance instInhabitedMonoidHom {M : Type u_4} {N : Type u_5} [MulOneClass M] [MulOneClass N] :
                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                      instance instInhabitedAddMonoidHom {M : Type u_4} {N : Type u_5} [AddZeroClass M] [AddZeroClass N] :
                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem MonoidHom.one_comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [MulOneClass M] [MulOneClass N] [MulOneClass P] (f : M →* N) :
                                                                                                                                                                                                                                                        comp 1 f = 1
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem AddMonoidHom.zero_comp {M : Type u_4} {N : Type u_5} {P : Type u_6} [AddZeroClass M] [AddZeroClass N] [AddZeroClass P] (f : M →+ N) :
                                                                                                                                                                                                                                                        comp 0 f = 0
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem MonoidHom.comp_one {M : Type u_4} {N : Type u_5} {P : Type u_6} [MulOneClass M] [MulOneClass N] [MulOneClass P] (f : N →* P) :
                                                                                                                                                                                                                                                        f.comp 1 = 1
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem AddMonoidHom.comp_zero {M : Type u_4} {N : Type u_5} {P : Type u_6} [AddZeroClass M] [AddZeroClass N] [AddZeroClass P] (f : N →+ P) :
                                                                                                                                                                                                                                                        f.comp 0 = 0
                                                                                                                                                                                                                                                        theorem MonoidHom.map_inv {α : Type u_2} {β : Type u_3} [Group α] [DivisionMonoid β] (f : α →* β) (a : α) :
                                                                                                                                                                                                                                                        f a⁻¹ = (f a)⁻¹

                                                                                                                                                                                                                                                        Group homomorphisms preserve inverse.

                                                                                                                                                                                                                                                        theorem AddMonoidHom.map_neg {α : Type u_2} {β : Type u_3} [AddGroup α] [SubtractionMonoid β] (f : α →+ β) (a : α) :
                                                                                                                                                                                                                                                        f (-a) = -f a

                                                                                                                                                                                                                                                        Additive group homomorphisms preserve negation.

                                                                                                                                                                                                                                                        theorem MonoidHom.map_zpow {α : Type u_2} {β : Type u_3} [Group α] [DivisionMonoid β] (f : α →* β) (g : α) (n : ) :
                                                                                                                                                                                                                                                        f (g ^ n) = f g ^ n

                                                                                                                                                                                                                                                        Group homomorphisms preserve integer power.

                                                                                                                                                                                                                                                        theorem AddMonoidHom.map_zsmul {α : Type u_2} {β : Type u_3} [AddGroup α] [SubtractionMonoid β] (f : α →+ β) (g : α) (n : ) :
                                                                                                                                                                                                                                                        f (n g) = n f g

                                                                                                                                                                                                                                                        Additive group homomorphisms preserve integer scaling.

                                                                                                                                                                                                                                                        theorem MonoidHom.map_div {α : Type u_2} {β : Type u_3} [Group α] [DivisionMonoid β] (f : α →* β) (g h : α) :
                                                                                                                                                                                                                                                        f (g / h) = f g / f h

                                                                                                                                                                                                                                                        Group homomorphisms preserve division.

                                                                                                                                                                                                                                                        theorem AddMonoidHom.map_sub {α : Type u_2} {β : Type u_3} [AddGroup α] [SubtractionMonoid β] (f : α →+ β) (g h : α) :
                                                                                                                                                                                                                                                        f (g - h) = f g - f h

                                                                                                                                                                                                                                                        Additive group homomorphisms preserve subtraction.

                                                                                                                                                                                                                                                        theorem MonoidHom.map_mul_inv {α : Type u_2} {β : Type u_3} [Group α] [DivisionMonoid β] (f : α →* β) (g h : α) :
                                                                                                                                                                                                                                                        f (g * h⁻¹) = f g * (f h)⁻¹

                                                                                                                                                                                                                                                        Group homomorphisms preserve division.

                                                                                                                                                                                                                                                        theorem AddMonoidHom.map_add_neg {α : Type u_2} {β : Type u_3} [AddGroup α] [SubtractionMonoid β] (f : α →+ β) (g h : α) :
                                                                                                                                                                                                                                                        f (g + -h) = f g + -f h

                                                                                                                                                                                                                                                        Additive group homomorphisms preserve subtraction.

                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem iterate_map_mul {M : Type u_10} {F : Type u_11} [Mul M] [FunLike F M M] [MulHomClass F M M] (f : F) (n : ) (x y : M) :
                                                                                                                                                                                                                                                        (⇑f)^[n] (x * y) = (⇑f)^[n] x * (⇑f)^[n] y
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem iterate_map_add {M : Type u_10} {F : Type u_11} [Add M] [FunLike F M M] [AddHomClass F M M] (f : F) (n : ) (x y : M) :
                                                                                                                                                                                                                                                        (⇑f)^[n] (x + y) = (⇑f)^[n] x + (⇑f)^[n] y
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem iterate_map_one {M : Type u_10} {F : Type u_11} [One M] [FunLike F M M] [OneHomClass F M M] (f : F) (n : ) :
                                                                                                                                                                                                                                                        (⇑f)^[n] 1 = 1
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem iterate_map_zero {M : Type u_10} {F : Type u_11} [Zero M] [FunLike F M M] [ZeroHomClass F M M] (f : F) (n : ) :
                                                                                                                                                                                                                                                        (⇑f)^[n] 0 = 0
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem iterate_map_inv {M : Type u_10} {F : Type u_11} [Group M] [FunLike F M M] [MonoidHomClass F M M] (f : F) (n : ) (x : M) :
                                                                                                                                                                                                                                                        (⇑f)^[n] x⁻¹ = ((⇑f)^[n] x)⁻¹
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem iterate_map_neg {M : Type u_10} {F : Type u_11} [AddGroup M] [FunLike F M M] [AddMonoidHomClass F M M] (f : F) (n : ) (x : M) :
                                                                                                                                                                                                                                                        (⇑f)^[n] (-x) = -(⇑f)^[n] x
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem iterate_map_div {M : Type u_10} {F : Type u_11} [Group M] [FunLike F M M] [MonoidHomClass F M M] (f : F) (n : ) (x y : M) :
                                                                                                                                                                                                                                                        (⇑f)^[n] (x / y) = (⇑f)^[n] x / (⇑f)^[n] y
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem iterate_map_sub {M : Type u_10} {F : Type u_11} [AddGroup M] [FunLike F M M] [AddMonoidHomClass F M M] (f : F) (n : ) (x y : M) :
                                                                                                                                                                                                                                                        (⇑f)^[n] (x - y) = (⇑f)^[n] x - (⇑f)^[n] y
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem iterate_map_pow {M : Type u_10} {F : Type u_11} [Monoid M] [FunLike F M M] [MonoidHomClass F M M] (f : F) (n : ) (x : M) (k : ) :
                                                                                                                                                                                                                                                        (⇑f)^[n] (x ^ k) = (⇑f)^[n] x ^ k
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem iterate_map_nsmul {M : Type u_10} {F : Type u_11} [AddMonoid M] [FunLike F M M] [AddMonoidHomClass F M M] (f : F) (n : ) (x : M) (k : ) :
                                                                                                                                                                                                                                                        (⇑f)^[n] (k x) = k (⇑f)^[n] x
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem iterate_map_zpow {M : Type u_10} {F : Type u_11} [Group M] [FunLike F M M] [MonoidHomClass F M M] (f : F) (n : ) (x : M) (k : ) :
                                                                                                                                                                                                                                                        (⇑f)^[n] (x ^ k) = (⇑f)^[n] x ^ k
                                                                                                                                                                                                                                                        @[simp]
                                                                                                                                                                                                                                                        theorem iterate_map_zsmul {M : Type u_10} {F : Type u_11} [AddGroup M] [FunLike F M M] [AddMonoidHomClass F M M] (f : F) (n : ) (x : M) (k : ) :
                                                                                                                                                                                                                                                        (⇑f)^[n] (k x) = k (⇑f)^[n] x