Documentation

Mathlib.Algebra.Group.End

Monoids of endomorphisms, groups of automorphisms #

This file defines

Implementation notes #

The definition of multiplication in the endomorphism monoids and automorphism groups agrees with function composition, and multiplication in CategoryTheory.End, but not with CategoryTheory.comp.

Tags #

end monoid, aut group

Type endomorphisms #

def Function.End (α : Type u_4) :
Type u_4

The monoid of endomorphisms.

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

Equations
Instances For
    instance instMonoidEnd {α : Type u_4} :
    Equations
    instance instInhabitedEnd {α : Type u_4} :
    Equations

    Monoid endomorphisms #

    instance Equiv.Perm.instOne {α : Type u_4} :
    One (Perm α)
    Equations
    instance Equiv.Perm.instMul {α : Type u_4} :
    Mul (Perm α)
    Equations
    instance Equiv.Perm.instInv {α : Type u_4} :
    Inv (Perm α)
    Equations
    instance Equiv.Perm.instPowNat {α : Type u_4} :
    Pow (Perm α)
    Equations
    @[simp]
    theorem Equiv.Perm.default_eq {α : Type u_4} :

    The permutation of a type is equivalent to the units group of the endomorphisms monoid of this type.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      @[simp]
      theorem Equiv.Perm.val_equivUnitsEnd_apply {α : Type u_4} (e : Perm α) (a✝ : α) :
      (equivUnitsEnd e) a✝ = e.toFun a✝
      @[simp]
      theorem Equiv.Perm.val_inv_equivUnitsEnd_apply {α : Type u_4} (e : Perm α) (a✝ : α) :
      (equivUnitsEnd e)⁻¹ a✝ = (Equiv.symm e).toFun a✝
      def MonoidHom.toHomPerm {α : Type u_4} {G : Type u_6} [Group G] (f : G →* Function.End α) :

      Lift a monoid homomorphism f : G →* Function.End α to a monoid homomorphism f : G →* Equiv.Perm α.

      Equations
      Instances For
        @[simp]
        theorem MonoidHom.toHomPerm_apply_apply {α : Type u_4} {G : Type u_6} [Group G] (f : G →* Function.End α) (a✝ : G) :
        (f.toHomPerm a✝) = f a✝
        @[simp]
        theorem MonoidHom.toHomPerm_apply_symm_apply {α : Type u_4} {G : Type u_6} [Group G] (f : G →* Function.End α) (a✝ : G) :
        (Equiv.symm (f.toHomPerm a✝)) = (f.toHomUnits a✝)⁻¹
        theorem Equiv.Perm.mul_apply {α : Type u_4} (f g : Perm α) (x : α) :
        (f * g) x = f (g x)
        theorem Equiv.Perm.one_apply {α : Type u_4} (x : α) :
        1 x = x
        @[simp]
        theorem Equiv.Perm.inv_apply_self {α : Type u_4} (f : Perm α) (x : α) :
        f⁻¹ (f x) = x
        @[simp]
        theorem Equiv.Perm.apply_inv_self {α : Type u_4} (f : Perm α) (x : α) :
        f (f⁻¹ x) = x
        theorem Equiv.Perm.one_def {α : Type u_4} :
        theorem Equiv.Perm.mul_def {α : Type u_4} (f g : Perm α) :
        f * g = Equiv.trans g f
        theorem Equiv.Perm.inv_def {α : Type u_4} (f : Perm α) :
        @[simp]
        theorem Equiv.Perm.coe_one {α : Type u_4} :
        1 = id
        @[simp]
        theorem Equiv.Perm.coe_mul {α : Type u_4} (f g : Perm α) :
        ⇑(f * g) = f g
        theorem Equiv.Perm.coe_pow {α : Type u_4} (f : Perm α) (n : ) :
        ⇑(f ^ n) = (⇑f)^[n]
        @[simp]
        theorem Equiv.Perm.iterate_eq_pow {α : Type u_4} (f : Perm α) (n : ) :
        (⇑f)^[n] = ⇑(f ^ n)
        theorem Equiv.Perm.eq_inv_iff_eq {α : Type u_4} {f : Perm α} {x y : α} :
        x = f⁻¹ y f x = y
        theorem Equiv.Perm.inv_eq_iff_eq {α : Type u_4} {f : Perm α} {x y : α} :
        f⁻¹ x = y x = f y
        theorem Equiv.Perm.zpow_apply_comm {α : Type u_6} (σ : Perm α) (m n : ) {x : α} :
        (σ ^ m) ((σ ^ n) x) = (σ ^ n) ((σ ^ m) x)

        Lemmas about mixing Perm with Equiv. Because we have multiple ways to express Equiv.refl, Equiv.symm, and Equiv.trans, we want simp lemmas for every combination. The assumption made here is that if you're using the group structure, you want to preserve it after simp.

        @[simp]
        theorem Equiv.Perm.trans_one {α : Sort u_6} {β : Type u_7} (e : α β) :
        e.trans 1 = e
        @[simp]
        theorem Equiv.Perm.mul_refl {α : Type u_4} (e : Perm α) :
        e * Equiv.refl α = e
        @[simp]
        theorem Equiv.Perm.one_symm {α : Type u_4} :
        @[simp]
        theorem Equiv.Perm.refl_inv {α : Type u_4} :
        @[simp]
        theorem Equiv.Perm.one_trans {α : Type u_6} {β : Sort u_7} (e : α β) :
        @[simp]
        theorem Equiv.Perm.refl_mul {α : Type u_4} (e : Perm α) :
        Equiv.refl α * e = e
        @[simp]
        theorem Equiv.Perm.inv_trans_self {α : Type u_4} (e : Perm α) :
        @[simp]
        theorem Equiv.Perm.mul_symm {α : Type u_4} (e : Perm α) :
        e * Equiv.symm e = 1
        @[simp]
        theorem Equiv.Perm.self_trans_inv {α : Type u_4} (e : Perm α) :
        @[simp]
        theorem Equiv.Perm.symm_mul {α : Type u_4} (e : Perm α) :
        Equiv.symm e * e = 1

        Lemmas about Equiv.Perm.sumCongr re-expressed via the group structure.

        @[simp]
        theorem Equiv.Perm.sumCongr_mul {α : Type u_6} {β : Type u_7} (e : Perm α) (f : Perm β) (g : Perm α) (h : Perm β) :
        e.sumCongr f * g.sumCongr h = (e * g).sumCongr (f * h)
        @[simp]
        theorem Equiv.Perm.sumCongr_inv {α : Type u_6} {β : Type u_7} (e : Perm α) (f : Perm β) :
        @[simp]
        theorem Equiv.Perm.sumCongr_one {α : Type u_6} {β : Type u_7} :
        sumCongr 1 1 = 1
        def Equiv.Perm.sumCongrHom (α : Type u_6) (β : Type u_7) :
        Perm α × Perm β →* Perm (α β)

        Equiv.Perm.sumCongr as a MonoidHom, with its two arguments bundled into a single Prod.

        This is particularly useful for its MonoidHom.range projection, which is the subgroup of permutations which do not exchange elements between α and β.

        Equations
        Instances For
          @[simp]
          theorem Equiv.Perm.sumCongrHom_apply (α : Type u_6) (β : Type u_7) (a : Perm α × Perm β) :
          (sumCongrHom α β) a = a.1.sumCongr a.2
          @[simp]
          theorem Equiv.Perm.sumCongr_swap_one {α : Type u_6} {β : Type u_7} [DecidableEq α] [DecidableEq β] (i j : α) :
          (swap i j).sumCongr 1 = swap (Sum.inl i) (Sum.inl j)
          @[simp]
          theorem Equiv.Perm.sumCongr_one_swap {α : Type u_6} {β : Type u_7} [DecidableEq α] [DecidableEq β] (i j : β) :
          sumCongr 1 (swap i j) = swap (Sum.inr i) (Sum.inr j)

          Lemmas about Equiv.Perm.sigmaCongrRight re-expressed via the group structure.

          @[simp]
          theorem Equiv.Perm.sigmaCongrRight_mul {α : Type u_6} {β : αType u_7} (F G : (a : α) → Perm (β a)) :
          @[simp]
          theorem Equiv.Perm.sigmaCongrRight_inv {α : Type u_6} {β : αType u_7} (F : (a : α) → Perm (β a)) :
          (sigmaCongrRight F)⁻¹ = sigmaCongrRight fun (a : α) => (F a)⁻¹
          @[simp]
          theorem Equiv.Perm.sigmaCongrRight_one {α : Type u_6} {β : αType u_7} :
          def Equiv.Perm.sigmaCongrRightHom {α : Type u_6} (β : αType u_7) :
          ((a : α) → Perm (β a)) →* Perm ((a : α) × β a)

          Equiv.Perm.sigmaCongrRight as a MonoidHom.

          This is particularly useful for its MonoidHom.range projection, which is the subgroup of permutations which do not exchange elements between fibers.

          Equations
          Instances For
            @[simp]
            theorem Equiv.Perm.sigmaCongrRightHom_apply {α : Type u_6} (β : αType u_7) (F : (a : α) → Perm (β a)) :
            def Equiv.Perm.subtypeCongrHom {α : Type u_4} (p : αProp) [DecidablePred p] :
            Perm { a : α // p a } × Perm { a : α // ¬p a } →* Perm α

            Equiv.Perm.subtypeCongr as a MonoidHom.

            Equations
            Instances For
              @[simp]
              theorem Equiv.Perm.subtypeCongrHom_apply {α : Type u_4} (p : αProp) [DecidablePred p] (pair : Perm { a : α // p a } × Perm { a : α // ¬p a }) :
              (subtypeCongrHom p) pair = pair.1.subtypeCongr pair.2
              @[simp]
              theorem Equiv.Perm.permCongr_eq_mul {α : Type u_4} (e p : Perm α) :
              (permCongr e) p = e * p * e⁻¹

              If e is also a permutation, we can write permCongr completely in terms of the group structure.

              Lemmas about Equiv.Perm.extendDomain re-expressed via the group structure.

              @[simp]
              theorem Equiv.Perm.extendDomain_one {α : Type u_4} {β : Type u_5} {p : βProp} [DecidablePred p] (f : α Subtype p) :
              @[simp]
              theorem Equiv.Perm.extendDomain_inv {α : Type u_4} {β : Type u_5} (e : Perm α) {p : βProp} [DecidablePred p] (f : α Subtype p) :
              @[simp]
              theorem Equiv.Perm.extendDomain_mul {α : Type u_4} {β : Type u_5} {p : βProp} [DecidablePred p] (f : α Subtype p) (e e' : Perm α) :
              def Equiv.Perm.extendDomainHom {α : Type u_4} {β : Type u_5} {p : βProp} [DecidablePred p] (f : α Subtype p) :
              Perm α →* Perm β

              extendDomain as a group homomorphism

              Equations
              Instances For
                @[simp]
                theorem Equiv.Perm.extendDomainHom_apply {α : Type u_4} {β : Type u_5} {p : βProp} [DecidablePred p] (f : α Subtype p) (e : Perm α) :
                theorem Equiv.Perm.extendDomainHom_injective {α : Type u_4} {β : Type u_5} {p : βProp} [DecidablePred p] (f : α Subtype p) :
                @[simp]
                theorem Equiv.Perm.extendDomain_eq_one_iff {α : Type u_4} {β : Type u_5} {p : βProp} [DecidablePred p] {e : Perm α} {f : α Subtype p} :
                e.extendDomain f = 1 e = 1
                @[simp]
                theorem Equiv.Perm.extendDomain_pow {α : Type u_4} {β : Type u_5} (e : Perm α) {p : βProp} [DecidablePred p] (f : α Subtype p) (n : ) :
                (e ^ n).extendDomain f = e.extendDomain f ^ n
                @[simp]
                theorem Equiv.Perm.extendDomain_zpow {α : Type u_4} {β : Type u_5} (e : Perm α) {p : βProp} [DecidablePred p] (f : α Subtype p) (n : ) :
                (e ^ n).extendDomain f = e.extendDomain f ^ n
                def Equiv.Perm.subtypePerm {α : Type u_4} {p : αProp} (f : Perm α) (h : ∀ (x : α), p x p (f x)) :
                Perm { x : α // p x }

                If the permutation f fixes the subtype {x // p x}, then this returns the permutation on {x // p x} induced by f.

                Equations
                • f.subtypePerm h = { toFun := fun (x : { x : α // p x }) => f x, , invFun := fun (x : { x : α // p x }) => f⁻¹ x, , left_inv := , right_inv := }
                Instances For
                  @[simp]
                  theorem Equiv.Perm.subtypePerm_apply {α : Type u_4} {p : αProp} (f : Perm α) (h : ∀ (x : α), p x p (f x)) (x : { x : α // p x }) :
                  (f.subtypePerm h) x = f x,
                  @[simp]
                  theorem Equiv.Perm.subtypePerm_one {α : Type u_4} (p : αProp) (h : ∀ (x : α), p x p x := ) :
                  @[simp]
                  theorem Equiv.Perm.subtypePerm_mul {α : Type u_4} {p : αProp} (f g : Perm α) (hf : ∀ (x : α), p x p (f x)) (hg : ∀ (x : α), p x p (g x)) :
                  f.subtypePerm hf * g.subtypePerm hg = (f * g).subtypePerm
                  theorem Equiv.Perm.subtypePerm_inv {α : Type u_4} {p : αProp} (f : Perm α) (hf : ∀ (x : α), p x p (f⁻¹ x)) :

                  See Equiv.Perm.inv_subtypePerm.

                  @[simp]
                  theorem Equiv.Perm.inv_subtypePerm {α : Type u_4} {p : αProp} (f : Perm α) (hf : ∀ (x : α), p x p (f x)) :

                  See Equiv.Perm.subtypePerm_inv.

                  @[simp]
                  theorem Equiv.Perm.subtypePerm_pow {α : Type u_4} {p : αProp} (f : Perm α) (n : ) (hf : ∀ (x : α), p x p (f x)) :
                  f.subtypePerm hf ^ n = (f ^ n).subtypePerm
                  @[simp]
                  theorem Equiv.Perm.subtypePerm_zpow {α : Type u_4} {p : αProp} (f : Perm α) (n : ) (hf : ∀ (x : α), p x p (f x)) :
                  f.subtypePerm hf ^ n = (f ^ n).subtypePerm
                  def Equiv.Perm.ofSubtype {α : Type u_4} {p : αProp} [DecidablePred p] :

                  The inclusion map of permutations on a subtype of α into permutations of α, fixing the other points.

                  Equations
                  Instances For
                    theorem Equiv.Perm.ofSubtype_subtypePerm {α : Type u_4} {p : αProp} [DecidablePred p] {f : Perm α} (h₁ : ∀ (x : α), p x p (f x)) (h₂ : ∀ (x : α), f x xp x) :
                    theorem Equiv.Perm.ofSubtype_apply_of_mem {α : Type u_4} {p : αProp} [DecidablePred p] {a : α} (f : Perm (Subtype p)) (ha : p a) :
                    (ofSubtype f) a = (f a, ha)
                    @[simp]
                    theorem Equiv.Perm.ofSubtype_apply_coe {α : Type u_4} {p : αProp} [DecidablePred p] (f : Perm (Subtype p)) (x : Subtype p) :
                    (ofSubtype f) x = (f x)
                    theorem Equiv.Perm.ofSubtype_apply_of_not_mem {α : Type u_4} {p : αProp} [DecidablePred p] {a : α} (f : Perm (Subtype p)) (ha : ¬p a) :
                    (ofSubtype f) a = a
                    theorem Equiv.Perm.mem_iff_ofSubtype_apply_mem {α : Type u_4} {p : αProp} [DecidablePred p] (f : Perm (Subtype p)) (x : α) :
                    p x p ((ofSubtype f) x)
                    @[simp]
                    theorem Equiv.Perm.subtypePerm_ofSubtype {α : Type u_4} {p : αProp} [DecidablePred p] (f : Perm (Subtype p)) :
                    theorem Equiv.Perm.ofSubtype_subtypePerm_of_mem {α : Type u_4} {p : αProp} [DecidablePred p] {g : Perm α} (hg : ∀ (x : α), p x p (g x)) {a : α} (ha : p a) :
                    (ofSubtype (g.subtypePerm hg)) a = g a
                    theorem Equiv.Perm.ofSubtype_subtypePerm_of_not_mem {α : Type u_4} {p : αProp} [DecidablePred p] {g : Perm α} (hg : ∀ (x : α), p x p (g x)) {a : α} (ha : ¬p a) :
                    (ofSubtype (g.subtypePerm hg)) a = a
                    def Equiv.Perm.subtypeEquivSubtypePerm {α : Type u_4} (p : αProp) [DecidablePred p] :
                    Perm (Subtype p) { f : Perm α // ∀ (a : α), ¬p af a = a }

                    Permutations on a subtype are equivalent to permutations on the original type that fix pointwise the rest.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      @[simp]
                      theorem Equiv.Perm.subtypeEquivSubtypePerm_symm_apply {α : Type u_4} (p : αProp) [DecidablePred p] (f : { f : Perm α // ∀ (a : α), ¬p af a = a }) :
                      theorem Equiv.Perm.subtypeEquivSubtypePerm_apply_of_mem {α : Type u_4} {p : αProp} [DecidablePred p] {a : α} (f : Perm (Subtype p)) (h : p a) :
                      ((Perm.subtypeEquivSubtypePerm p) f) a = (f a, h)
                      theorem Equiv.Perm.subtypeEquivSubtypePerm_apply_of_not_mem {α : Type u_4} {p : αProp} [DecidablePred p] {a : α} (f : Perm (Subtype p)) (h : ¬p a) :
                      @[simp]
                      theorem Equiv.swap_inv {α : Type u_4} [DecidableEq α] (x y : α) :
                      (swap x y)⁻¹ = swap x y
                      @[simp]
                      theorem Equiv.swap_mul_self {α : Type u_4} [DecidableEq α] (i j : α) :
                      swap i j * swap i j = 1
                      theorem Equiv.swap_mul_eq_mul_swap {α : Type u_4} [DecidableEq α] (f : Perm α) (x y : α) :
                      swap x y * f = f * swap (f⁻¹ x) (f⁻¹ y)
                      theorem Equiv.mul_swap_eq_swap_mul {α : Type u_4} [DecidableEq α] (f : Perm α) (x y : α) :
                      f * swap x y = swap (f x) (f y) * f
                      theorem Equiv.swap_apply_apply {α : Type u_4} [DecidableEq α] (f : Perm α) (x y : α) :
                      swap (f x) (f y) = f * swap x y * f⁻¹
                      @[simp]
                      theorem Equiv.swap_mul_self_mul {α : Type u_4} [DecidableEq α] (i j : α) (σ : Perm α) :
                      swap i j * (swap i j * σ) = σ

                      Left-multiplying a permutation with swap i j twice gives the original permutation.

                      This specialization of swap_mul_self is useful when using cosets of permutations.

                      @[simp]
                      theorem Equiv.mul_swap_mul_self {α : Type u_4} [DecidableEq α] (i j : α) (σ : Perm α) :
                      σ * swap i j * swap i j = σ

                      Right-multiplying a permutation with swap i j twice gives the original permutation.

                      This specialization of swap_mul_self is useful when using cosets of permutations.

                      @[simp]
                      theorem Equiv.swap_mul_involutive {α : Type u_4} [DecidableEq α] (i j : α) :
                      Function.Involutive fun (x : Perm α) => swap i j * x

                      A stronger version of mul_right_injective

                      @[simp]
                      theorem Equiv.mul_swap_involutive {α : Type u_4} [DecidableEq α] (i j : α) :
                      Function.Involutive fun (x : Perm α) => x * swap i j

                      A stronger version of mul_left_injective

                      @[simp]
                      theorem Equiv.swap_eq_one_iff {α : Type u_4} [DecidableEq α] {i j : α} :
                      swap i j = 1 i = j
                      theorem Equiv.swap_mul_eq_iff {α : Type u_4} [DecidableEq α] {i j : α} {σ : Perm α} :
                      swap i j * σ = σ i = j
                      theorem Equiv.mul_swap_eq_iff {α : Type u_4} [DecidableEq α] {i j : α} {σ : Perm α} :
                      σ * swap i j = σ i = j
                      theorem Equiv.swap_mul_swap_mul_swap {α : Type u_4} [DecidableEq α] {x y z : α} (hxy : x y) (hxz : x z) :
                      swap y z * swap x y * swap y z = swap z x
                      @[simp]
                      theorem Equiv.addLeft_zero {α : Type u_4} [AddGroup α] :
                      @[simp]
                      theorem Equiv.addRight_zero {α : Type u_4} [AddGroup α] :
                      @[simp]
                      theorem Equiv.addLeft_add {α : Type u_4} [AddGroup α] (a b : α) :
                      @[simp]
                      theorem Equiv.addRight_add {α : Type u_4} [AddGroup α] (a b : α) :
                      @[simp]
                      theorem Equiv.inv_addLeft {α : Type u_4} [AddGroup α] (a : α) :
                      @[simp]
                      theorem Equiv.inv_addRight {α : Type u_4} [AddGroup α] (a : α) :
                      @[simp]
                      theorem Equiv.pow_addLeft {α : Type u_4} [AddGroup α] (a : α) (n : ) :
                      @[simp]
                      theorem Equiv.pow_addRight {α : Type u_4} [AddGroup α] (a : α) (n : ) :
                      @[simp]
                      theorem Equiv.zpow_addLeft {α : Type u_4} [AddGroup α] (a : α) (n : ) :
                      @[simp]
                      theorem Equiv.zpow_addRight {α : Type u_4} [AddGroup α] (a : α) (n : ) :
                      @[simp]
                      theorem Equiv.mulLeft_one {α : Type u_4} [Group α] :
                      @[simp]
                      theorem Equiv.mulRight_one {α : Type u_4} [Group α] :
                      @[simp]
                      theorem Equiv.mulLeft_mul {α : Type u_4} [Group α] (a b : α) :
                      @[simp]
                      theorem Equiv.mulRight_mul {α : Type u_4} [Group α] (a b : α) :
                      @[simp]
                      theorem Equiv.inv_mulLeft {α : Type u_4} [Group α] (a : α) :
                      @[simp]
                      theorem Equiv.inv_mulRight {α : Type u_4} [Group α] (a : α) :
                      @[simp]
                      theorem Equiv.pow_mulLeft {α : Type u_4} [Group α] (a : α) (n : ) :
                      @[simp]
                      theorem Equiv.pow_mulRight {α : Type u_4} [Group α] (a : α) (n : ) :
                      @[simp]
                      theorem Equiv.zpow_mulLeft {α : Type u_4} [Group α] (a : α) (n : ) :
                      @[simp]
                      theorem Equiv.zpow_mulRight {α : Type u_4} [Group α] (a : α) (n : ) :
                      @[reducible]
                      def MulAut (M : Type u_6) [Mul M] :
                      Type u_6

                      The group of multiplicative automorphisms.

                      Equations
                      Instances For
                        @[reducible]
                        def AddAut (M : Type u_6) [Add M] :
                        Type u_6

                        The group of additive automorphisms.

                        Equations
                        Instances For
                          instance MulAut.instGroup (M : Type u_2) [Mul M] :

                          The group operation on multiplicative automorphisms is defined by g h => MulEquiv.trans h g. This means that multiplication agrees with composition, (g*h)(x) = g (h x).

                          Equations
                          instance MulAut.instInhabited (M : Type u_2) [Mul M] :
                          Equations
                          @[simp]
                          theorem MulAut.coe_mul (M : Type u_2) [Mul M] (e₁ e₂ : MulAut M) :
                          ⇑(e₁ * e₂) = e₁ e₂
                          @[simp]
                          theorem MulAut.coe_one (M : Type u_2) [Mul M] :
                          1 = id
                          theorem MulAut.mul_def (M : Type u_2) [Mul M] (e₁ e₂ : MulAut M) :
                          e₁ * e₂ = MulEquiv.trans e₂ e₁
                          theorem MulAut.one_def (M : Type u_2) [Mul M] :
                          theorem MulAut.inv_def (M : Type u_2) [Mul M] (e₁ : MulAut M) :
                          @[simp]
                          theorem MulAut.mul_apply (M : Type u_2) [Mul M] (e₁ e₂ : MulAut M) (m : M) :
                          (e₁ * e₂) m = e₁ (e₂ m)
                          @[simp]
                          theorem MulAut.one_apply (M : Type u_2) [Mul M] (m : M) :
                          1 m = m
                          @[simp]
                          theorem MulAut.apply_inv_self (M : Type u_2) [Mul M] (e : MulAut M) (m : M) :
                          e (e⁻¹ m) = m
                          @[simp]
                          theorem MulAut.inv_apply_self (M : Type u_2) [Mul M] (e : MulAut M) (m : M) :
                          e⁻¹ (e m) = m
                          def MulAut.toPerm (M : Type u_2) [Mul M] :

                          Monoid hom from the group of multiplicative automorphisms to the group of permutations.

                          Equations
                          Instances For
                            def MulAut.conj {G : Type u_3} [Group G] :

                            Group conjugation, MulAut.conj g h = g * h * g⁻¹, as a monoid homomorphism mapping multiplication in G into multiplication in the automorphism group MulAut G. See also the type ConjAct G for any group G, which has a MulAction (ConjAct G) G instance where conj G acts on G by conjugation.

                            Equations
                            • One or more equations did not get rendered due to their size.
                            Instances For
                              @[simp]
                              theorem MulAut.conj_apply {G : Type u_3} [Group G] (g h : G) :
                              (conj g) h = g * h * g⁻¹
                              @[simp]
                              theorem MulAut.conj_symm_apply {G : Type u_3} [Group G] (g h : G) :
                              (MulEquiv.symm (conj g)) h = g⁻¹ * h * g
                              @[simp]
                              theorem MulAut.conj_inv_apply {G : Type u_3} [Group G] (g h : G) :
                              (conj g)⁻¹ h = g⁻¹ * h * g
                              def MulAut.congr {G : Type u_3} [Group G] {H : Type u_6} [Group H] (ϕ : G ≃* H) :

                              Isomorphic groups have isomorphic automorphism groups.

                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                @[simp]
                                theorem MulAut.congr_apply {G : Type u_3} [Group G] {H : Type u_6} [Group H] (ϕ : G ≃* H) (f : MulAut G) :
                                (congr ϕ) f = ϕ.symm.trans (MulEquiv.trans f ϕ)
                                @[simp]
                                theorem MulAut.congr_symm_apply {G : Type u_3} [Group G] {H : Type u_6} [Group H] (ϕ : G ≃* H) (f : MulAut H) :
                                (congr ϕ).symm f = ϕ.trans (MulEquiv.trans f ϕ.symm)
                                instance AddAut.group (A : Type u_1) [Add A] :

                                The group operation on additive automorphisms is defined by g h => AddEquiv.trans h g. This means that multiplication agrees with composition, (g*h)(x) = g (h x).

                                Equations
                                instance AddAut.instInhabited (A : Type u_1) [Add A] :
                                Equations
                                @[simp]
                                theorem AddAut.coe_mul (A : Type u_1) [Add A] (e₁ e₂ : AddAut A) :
                                ⇑(e₁ * e₂) = e₁ e₂
                                @[simp]
                                theorem AddAut.coe_one (A : Type u_1) [Add A] :
                                1 = id
                                theorem AddAut.mul_def (A : Type u_1) [Add A] (e₁ e₂ : AddAut A) :
                                e₁ * e₂ = AddEquiv.trans e₂ e₁
                                theorem AddAut.one_def (A : Type u_1) [Add A] :
                                theorem AddAut.inv_def (A : Type u_1) [Add A] (e₁ : AddAut A) :
                                @[simp]
                                theorem AddAut.mul_apply (A : Type u_1) [Add A] (e₁ e₂ : AddAut A) (a : A) :
                                (e₁ * e₂) a = e₁ (e₂ a)
                                @[simp]
                                theorem AddAut.one_apply (A : Type u_1) [Add A] (a : A) :
                                1 a = a
                                @[simp]
                                theorem AddAut.apply_inv_self (A : Type u_1) [Add A] (e : AddAut A) (a : A) :
                                e⁻¹ (e a) = a
                                @[simp]
                                theorem AddAut.inv_apply_self (A : Type u_1) [Add A] (e : AddAut A) (a : A) :
                                e (e⁻¹ a) = a
                                def AddAut.toPerm (A : Type u_1) [Add A] :

                                Monoid hom from the group of multiplicative automorphisms to the group of permutations.

                                Equations
                                Instances For
                                  def AddAut.conj {G : Type u_3} [AddGroup G] :

                                  Additive group conjugation, AddAut.conj g h = g + h - g, as an additive monoid homomorphism mapping addition in G into multiplication in the automorphism group AddAut G (written additively in order to define the map).

                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For
                                    @[simp]
                                    theorem AddAut.conj_apply {G : Type u_3} [AddGroup G] (g h : G) :
                                    (Additive.toMul (conj g)) h = g + h + -g
                                    @[simp]
                                    theorem AddAut.conj_symm_apply {G : Type u_3} [AddGroup G] (g h : G) :
                                    (AddEquiv.symm (conj g)) h = -g + h + g
                                    @[simp]
                                    theorem AddAut.conj_inv_apply {G : Type u_3} [AddGroup G] (g h : G) :
                                    (Additive.toMul (conj g))⁻¹ h = -g + h + g
                                    theorem AddAut.neg_conj_apply {G : Type u_3} [AddGroup G] (g h : G) :
                                    (Additive.toMul (-conj g)) h = -g + h + g
                                    def AddAut.congr {G : Type u_3} [AddGroup G] {H : Type u_6} [AddGroup H] (ϕ : G ≃+ H) :

                                    Isomorphic additive groups have isomorphic automorphism groups.

                                    Equations
                                    • One or more equations did not get rendered due to their size.
                                    Instances For
                                      @[simp]
                                      theorem AddAut.congr_apply {G : Type u_3} [AddGroup G] {H : Type u_6} [AddGroup H] (ϕ : G ≃+ H) (f : AddAut G) :
                                      (congr ϕ) f = ϕ.symm.trans (AddEquiv.trans f ϕ)
                                      @[simp]
                                      theorem AddAut.congr_symm_apply {G : Type u_3} [AddGroup G] {H : Type u_6} [AddGroup H] (ϕ : G ≃+ H) (f : AddAut H) :
                                      (congr ϕ).symm f = ϕ.trans (AddEquiv.trans f ϕ.symm)

                                      Multiplicative G and G have isomorphic automorphism groups.

                                      Equations
                                      Instances For

                                        Additive G and G have isomorphic automorphism groups.

                                        Equations
                                        Instances For
                                          @[simp]
                                          theorem AddAutAdditive_symm_apply_apply (G : Type u_3) [Group G] (a✝ : G ≃* G) (a : Additive G) :
                                          @[simp]
                                          theorem AddAutAdditive_apply_apply (G : Type u_3) [Group G] (a✝ : Additive G ≃+ Additive G) (a : G) :
                                          @[simp]
                                          theorem AddAutAdditive_apply_symm_apply (G : Type u_3) [Group G] (a✝ : Additive G ≃+ Additive G) (a : G) :
                                          @[simp]
                                          theorem AddAutAdditive_symm_apply_symm_apply (G : Type u_3) [Group G] (a✝ : G ≃* G) (a : Additive G) :