Documentation

Mathlib.Algebra.Group.Aut

Multiplicative and additive group automorphisms #

This file defines the automorphism group structure on AddAut R := AddEquiv R R and MulAut R := MulEquiv R R.

Implementation notes #

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

This file is kept separate from Algebra/Group/Equiv/* so that GroupTheory.Perm is free to use equivalences (and other files that use them) before the group structure is defined.

Tags #

MulAut, AddAut

@[reducible]
def MulAut (M : Type u_4) [Mul M] :
Type u_4

The group of multiplicative automorphisms.

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

    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
      • MulAut.toPerm M = { toFun := MulEquiv.toEquiv, map_one' := , map_mul' := }
      Instances For
        instance MulAut.applyMulAction {M : Type u_4} [Monoid M] :

        The tautological action by MulAut M on M.

        This generalizes Function.End.applyMulAction.

        Equations
        @[simp]
        theorem MulAut.smul_def {M : Type u_4} [Monoid M] (f : MulAut M) (a : M) :
        f a = f a

        MulAut.applyDistribMulAction is faithful.

        Equations
        • =
        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) :
          (MulAut.conj g) h = g * h * g⁻¹
          @[simp]
          theorem MulAut.conj_symm_apply {G : Type u_3} [Group G] (g h : G) :
          (MulEquiv.symm (MulAut.conj g)) h = g⁻¹ * h * g
          @[simp]
          theorem MulAut.conj_inv_apply {G : Type u_3} [Group G] (g h : G) :
          (MulAut.conj g)⁻¹ h = g⁻¹ * h * g
          def MulAut.congr {G : Type u_3} [Group G] {H : Type u_4} [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_4} [Group H] (ϕ : G ≃* H) (f : MulAut G) :
            (MulAut.congr ϕ) f = ϕ.symm.trans (MulEquiv.trans f ϕ)
            @[simp]
            theorem MulAut.congr_symm_apply {G : Type u_3} [Group G] {H : Type u_4} [Group H] (ϕ : G ≃* H) (f : MulAut H) :
            (MulAut.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
            • AddAut.toPerm A = { toFun := AddEquiv.toEquiv, map_one' := , map_mul' := }
            Instances For
              instance AddAut.applyMulAction {A : Type u_4} [AddMonoid A] :

              The tautological action by AddAut A on A.

              This generalizes Function.End.applyMulAction.

              Equations
              @[simp]
              theorem AddAut.smul_def {A : Type u_4} [AddMonoid A] (f : AddAut A) (a : A) :
              f a = f a

              AddAut.applyDistribMulAction is faithful.

              Equations
              • =
              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 (AddAut.conj g)) h = g + h + -g
                @[simp]
                theorem AddAut.conj_symm_apply {G : Type u_3} [AddGroup G] (g h : G) :
                (AddEquiv.symm (AddAut.conj g)) h = -g + h + g
                @[simp]
                theorem AddAut.conj_inv_apply {G : Type u_3} [AddGroup G] (g h : G) :
                (Additive.toMul (AddAut.conj g))⁻¹ h = -g + h + g
                def AddAut.congr {G : Type u_3} [AddGroup G] {H : Type u_4} [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_4} [AddGroup H] (ϕ : G ≃+ H) (f : AddAut G) :
                  (AddAut.congr ϕ) f = ϕ.symm.trans (AddEquiv.trans f ϕ)
                  @[simp]
                  theorem AddAut.congr_symm_apply {G : Type u_3} [AddGroup G] {H : Type u_4} [AddGroup H] (ϕ : G ≃+ H) (f : AddAut H) :
                  (AddAut.congr ϕ).symm f = ϕ.trans (AddEquiv.trans f ϕ.symm)