Documentation

Mathlib.Algebra.Category.Grp.Basic

Category instances for Group, AddGroup, CommGroup, and AddCommGroup. #

We introduce the bundled categories:

structure AddGrp :
Type (u + 1)

The category of additive groups and group morphisms.

Instances For
    structure Grp :
    Type (u + 1)

    The category of groups and group morphisms.

    • carrier : Type u

      The underlying type.

    • str : Group self
    Instances For
      @[reducible, inline]
      abbrev Grp.of (M : Type u) [Group M] :

      Construct a bundled Grp from the underlying type and typeclass.

      Equations
      Instances For
        @[reducible, inline]
        abbrev AddGrp.of (M : Type u) [AddGroup M] :

        Construct a bundled AddGrp from the underlying type and typeclass.

        Equations
        Instances For
          structure AddGrp.Hom (A B : AddGrp) :

          The type of morphisms in AddGrp R.

          • hom' : A →+ B

            The underlying monoid homomorphism.

          Instances For
            theorem AddGrp.Hom.ext {A B : AddGrp} {x y : A.Hom B} (hom' : x.hom' = y.hom') :
            x = y
            structure Grp.Hom (A B : Grp) :

            The type of morphisms in Grp R.

            • hom' : A →* B

              The underlying monoid homomorphism.

            Instances For
              theorem Grp.Hom.ext {A B : Grp} {x y : A.Hom B} (hom' : x.hom' = y.hom') :
              x = y
              Equations
              • One or more equations did not get rendered due to their size.
              Equations
              • One or more equations did not get rendered due to their size.
              @[reducible, inline]
              abbrev Grp.Hom.hom {X Y : Grp} (f : X.Hom Y) :
              X →* Y

              Turn a morphism in Grp back into a MonoidHom.

              Equations
              Instances For
                @[reducible, inline]
                abbrev AddGrp.Hom.hom {X Y : AddGrp} (f : X.Hom Y) :
                X →+ Y

                Turn a morphism in AddGrp back into an AddMonoidHom.

                Equations
                Instances For
                  @[reducible, inline]
                  abbrev Grp.ofHom {X Y : Type u} [Group X] [Group Y] (f : X →* Y) :
                  of X of Y

                  Typecheck a MonoidHom as a morphism in Grp.

                  Equations
                  Instances For
                    @[reducible, inline]
                    abbrev AddGrp.ofHom {X Y : Type u} [AddGroup X] [AddGroup Y] (f : X →+ Y) :
                    of X of Y

                    Typecheck an AddMonoidHom as a morphism in AddGrp.

                    Equations
                    Instances For
                      def Grp.Hom.Simps.hom (X Y : Grp) (f : X.Hom Y) :
                      X →* Y

                      Use the ConcreteCategory.hom projection for @[simps] lemmas.

                      Equations
                      Instances For

                        The results below duplicate the ConcreteCategory simp lemmas, but we can keep them for dsimp.

                        @[deprecated "Use hom_comp instead" (since := "2025-01-28")]
                        theorem Grp.comp_def {X Y Z : Grp} {f : X Y} {g : Y Z} :
                        @[deprecated "Use hom_comp instead" (since := "2025-01-28")]
                        theorem AddGrp.comp_def {X Y Z : AddGrp} {f : X Y} {g : Y Z} :
                        theorem AddGrp.ext {X Y : AddGrp} {f g : X Y} (w : ∀ (x : X), (CategoryTheory.ConcreteCategory.hom f) x = (CategoryTheory.ConcreteCategory.hom g) x) :
                        f = g
                        theorem Grp.ext {X Y : Grp} {f g : X Y} (w : ∀ (x : X), (CategoryTheory.ConcreteCategory.hom f) x = (CategoryTheory.ConcreteCategory.hom g) x) :
                        f = g
                        theorem Grp.coe_of (R : Type u) [Group R] :
                        (of R) = R
                        theorem AddGrp.coe_of (R : Type u) [AddGroup R] :
                        (of R) = R
                        @[simp]
                        theorem Grp.hom_comp {X Y T : Grp} (f : X Y) (g : Y T) :
                        @[simp]
                        theorem AddGrp.hom_comp {X Y T : AddGrp} (f : X Y) (g : Y T) :
                        theorem AddGrp.hom_ext {X Y : AddGrp} {f g : X Y} (hf : Hom.hom f = Hom.hom g) :
                        f = g
                        theorem Grp.hom_ext {X Y : Grp} {f g : X Y} (hf : Hom.hom f = Hom.hom g) :
                        f = g
                        @[simp]
                        theorem Grp.hom_ofHom {R S : Type u} [Group R] [Group S] (f : R →* S) :
                        @[simp]
                        theorem AddGrp.hom_ofHom {R S : Type u} [AddGroup R] [AddGroup S] (f : R →+ S) :
                        @[simp]
                        theorem Grp.ofHom_hom {X Y : Grp} (f : X Y) :
                        @[simp]
                        theorem AddGrp.ofHom_hom {X Y : AddGrp} (f : X Y) :
                        @[simp]
                        theorem Grp.ofHom_comp {X Y Z : Type u} [Group X] [Group Y] [Group Z] (f : X →* Y) (g : Y →* Z) :
                        @[simp]
                        theorem AddGrp.ofHom_comp {X Y Z : Type u} [AddGroup X] [AddGroup Y] [AddGroup Z] (f : X →+ Y) (g : Y →+ Z) :
                        theorem Grp.ofHom_apply {X Y : Type u} [Group X] [Group Y] (f : X →* Y) (x : X) :
                        theorem AddGrp.ofHom_apply {X Y : Type u} [AddGroup X] [AddGroup Y] (f : X →+ Y) (x : X) :
                        @[deprecated "use `coe_comp` instead" (since := "2025-01-28")]

                        Alias of Grp.coe_comp.

                        @[deprecated "use `coe_comp` instead" (since := "2025-01-28")]
                        @[deprecated "use `coe_id` instead" (since := "2025-01-28")]

                        Alias of Grp.coe_id.

                        @[deprecated "use `coe_id` instead" (since := "2025-01-28")]
                        Equations
                        • One or more equations did not get rendered due to their size.
                        Equations
                        • One or more equations did not get rendered due to their size.
                        instance Grp.instOneHom (G H : Grp) :
                        One (G H)
                        Equations
                        instance AddGrp.instZeroHom (G H : AddGrp) :
                        Zero (G H)
                        Equations
                        @[simp]
                        theorem Grp.one_apply (G H : Grp) (g : G) :
                        @[simp]
                        theorem Grp.ofHom_injective {X Y : Type u} [Group X] [Group Y] :
                        Function.Injective fun (f : X →* Y) => ofHom f
                        theorem AddGrp.ofHom_injective {X Y : Type u} [AddGroup X] [AddGroup Y] :
                        Function.Injective fun (f : X →+ Y) => ofHom f
                        instance Grp.ofUnique (G : Type u_1) [Group G] [i : Unique G] :
                        Unique (of G)
                        Equations
                        instance AddGrp.ofUnique (G : Type u_1) [AddGroup G] [i : Unique G] :
                        Unique (of G)
                        Equations

                        Universe lift functor for groups.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For

                          Universe lift functor for additive groups.

                          Equations
                          • One or more equations did not get rendered due to their size.
                          Instances For
                            structure AddCommGrp :
                            Type (u + 1)

                            The category of additive groups and group morphisms.

                            Instances For
                              structure CommGrp :
                              Type (u + 1)

                              The category of groups and group morphisms.

                              Instances For
                                @[reducible, inline]
                                abbrev Ab :
                                Type (u_1 + 1)

                                Ab is an abbreviation for AddCommGroup, for the sake of mathematicians' sanity.

                                Equations
                                Instances For
                                  @[reducible, inline]
                                  abbrev CommGrp.of (M : Type u) [CommGroup M] :

                                  Construct a bundled CommGrp from the underlying type and typeclass.

                                  Equations
                                  Instances For
                                    @[reducible, inline]

                                    Construct a bundled AddCommGrp from the underlying type and typeclass.

                                    Equations
                                    Instances For
                                      structure AddCommGrp.Hom (A B : AddCommGrp) :

                                      The type of morphisms in AddCommGrp R.

                                      • hom' : A →+ B

                                        The underlying monoid homomorphism.

                                      Instances For
                                        theorem AddCommGrp.Hom.ext {A B : AddCommGrp} {x y : A.Hom B} (hom' : x.hom' = y.hom') :
                                        x = y
                                        structure CommGrp.Hom (A B : CommGrp) :

                                        The type of morphisms in CommGrp R.

                                        • hom' : A →* B

                                          The underlying monoid homomorphism.

                                        Instances For
                                          theorem CommGrp.Hom.ext {A B : CommGrp} {x y : A.Hom B} (hom' : x.hom' = y.hom') :
                                          x = y
                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          @[reducible, inline]
                                          abbrev CommGrp.Hom.hom {X Y : CommGrp} (f : X.Hom Y) :
                                          X →* Y

                                          Turn a morphism in CommGrp back into a MonoidHom.

                                          Equations
                                          Instances For
                                            @[reducible, inline]
                                            abbrev AddCommGrp.Hom.hom {X Y : AddCommGrp} (f : X.Hom Y) :
                                            X →+ Y

                                            Turn a morphism in AddCommGrp back into an AddMonoidHom.

                                            Equations
                                            Instances For
                                              @[reducible, inline]
                                              abbrev CommGrp.ofHom {X Y : Type u} [CommGroup X] [CommGroup Y] (f : X →* Y) :
                                              of X of Y

                                              Typecheck a MonoidHom as a morphism in CommGrp.

                                              Equations
                                              Instances For
                                                @[reducible, inline]
                                                abbrev AddCommGrp.ofHom {X Y : Type u} [AddCommGroup X] [AddCommGroup Y] (f : X →+ Y) :
                                                of X of Y

                                                Typecheck an AddMonoidHom as a morphism in AddCommGrp.

                                                Equations
                                                Instances For
                                                  def CommGrp.Hom.Simps.hom (X Y : CommGrp) (f : X.Hom Y) :
                                                  X →* Y

                                                  Use the ConcreteCategory.hom projection for @[simps] lemmas.

                                                  Equations
                                                  Instances For
                                                    def AddCommGrp.Hom.Simps.hom (X Y : AddCommGrp) (f : X.Hom Y) :
                                                    X →+ Y

                                                    Use the ConcreteCategory.hom projection for @[simps] lemmas.

                                                    Equations
                                                    Instances For

                                                      The results below duplicate the ConcreteCategory simp lemmas, but we can keep them for dsimp.

                                                      @[deprecated "Use hom_comp instead" (since := "2025-01-28")]
                                                      theorem CommGrp.comp_def {X Y Z : CommGrp} {f : X Y} {g : Y Z} :
                                                      @[deprecated "Use hom_comp instead" (since := "2025-01-28")]
                                                      theorem AddCommGrp.ext {X Y : AddCommGrp} {f g : X Y} (w : ∀ (x : X), (CategoryTheory.ConcreteCategory.hom f) x = (CategoryTheory.ConcreteCategory.hom g) x) :
                                                      f = g
                                                      theorem CommGrp.ext {X Y : CommGrp} {f g : X Y} (w : ∀ (x : X), (CategoryTheory.ConcreteCategory.hom f) x = (CategoryTheory.ConcreteCategory.hom g) x) :
                                                      f = g
                                                      theorem CommGrp.coe_of (R : Type u) [CommGroup R] :
                                                      (of R) = R
                                                      theorem AddCommGrp.coe_of (R : Type u) [AddCommGroup R] :
                                                      (of R) = R
                                                      @[simp]
                                                      theorem CommGrp.hom_comp {X Y T : CommGrp} (f : X Y) (g : Y T) :
                                                      @[simp]
                                                      theorem AddCommGrp.hom_ext {X Y : AddCommGrp} {f g : X Y} (hf : Hom.hom f = Hom.hom g) :
                                                      f = g
                                                      theorem CommGrp.hom_ext {X Y : CommGrp} {f g : X Y} (hf : Hom.hom f = Hom.hom g) :
                                                      f = g
                                                      @[simp]
                                                      theorem CommGrp.hom_ofHom {X Y : Type u} [CommGroup X] [CommGroup Y] (f : X →* Y) :
                                                      @[simp]
                                                      theorem AddCommGrp.hom_ofHom {X Y : Type u} [AddCommGroup X] [AddCommGroup Y] (f : X →+ Y) :
                                                      @[simp]
                                                      theorem CommGrp.ofHom_hom {X Y : CommGrp} (f : X Y) :
                                                      @[simp]
                                                      theorem AddCommGrp.ofHom_hom {X Y : AddCommGrp} (f : X Y) :
                                                      @[simp]
                                                      theorem CommGrp.ofHom_comp {X Y Z : Type u} [CommGroup X] [CommGroup Y] [CommGroup Z] (f : X →* Y) (g : Y →* Z) :
                                                      @[simp]
                                                      theorem CommGrp.ofHom_apply {X Y : Type u} [CommGroup X] [CommGroup Y] (f : X →* Y) (x : X) :
                                                      @[deprecated "use `coe_comp` instead" (since := "2025-01-28")]

                                                      Alias of CommGrp.coe_comp.

                                                      @[deprecated "use `coe_comp` instead" (since := "2025-01-28")]
                                                      @[deprecated "use `coe_id` instead" (since := "2025-01-28")]

                                                      Alias of CommGrp.coe_id.

                                                      @[deprecated "use `coe_id` instead" (since := "2025-01-28")]
                                                      instance CommGrp.ofUnique (G : Type u_1) [CommGroup G] [i : Unique G] :
                                                      Unique (of G)
                                                      Equations
                                                      instance AddCommGrp.ofUnique (G : Type u_1) [AddCommGroup G] [i : Unique G] :
                                                      Unique (of G)
                                                      Equations
                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      instance CommGrp.instOneHom (G H : CommGrp) :
                                                      One (G H)
                                                      Equations
                                                      instance AddCommGrp.instZeroHom (G H : AddCommGrp) :
                                                      Zero (G H)
                                                      Equations
                                                      @[simp]
                                                      theorem CommGrp.ofHom_injective {X Y : Type u} [CommGroup X] [CommGroup Y] :
                                                      Function.Injective fun (f : X →* Y) => ofHom f

                                                      Universe lift functor for commutative groups.

                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      Instances For

                                                        Universe lift functor for additive commutative groups.

                                                        Equations
                                                        • One or more equations did not get rendered due to their size.
                                                        Instances For
                                                          def AddCommGrp.asHom {G : AddCommGrp} (g : G) :

                                                          Any element of an abelian group gives a unique morphism from sending 1 to that element.

                                                          Equations
                                                          Instances For
                                                            @[simp]
                                                            theorem AddCommGrp.asHom_hom_apply {G : AddCommGrp} (g : G) (n : ) :
                                                            (Hom.hom (asHom g)) n = n g
                                                            def MulEquiv.toGrpIso {X Y : Grp} (e : X ≃* Y) :
                                                            X Y

                                                            Build an isomorphism in the category Grp from a MulEquiv between Groups.

                                                            Equations
                                                            Instances For
                                                              def AddEquiv.toAddGrpIso {X Y : AddGrp} (e : X ≃+ Y) :
                                                              X Y

                                                              Build an isomorphism in the category AddGroup from an AddEquiv between AddGroups.

                                                              Equations
                                                              Instances For
                                                                @[simp]
                                                                theorem MulEquiv.toGrpIso_inv {X Y : Grp} (e : X ≃* Y) :
                                                                @[simp]
                                                                theorem MulEquiv.toGrpIso_hom {X Y : Grp} (e : X ≃* Y) :
                                                                def MulEquiv.toCommGrpIso {X Y : CommGrp} (e : X ≃* Y) :
                                                                X Y

                                                                Build an isomorphism in the category CommGrp from a MulEquiv between CommGroups.

                                                                Equations
                                                                Instances For
                                                                  def AddEquiv.toAddCommGrpIso {X Y : AddCommGrp} (e : X ≃+ Y) :
                                                                  X Y

                                                                  Build an isomorphism in the category AddCommGrp from an AddEquiv between AddCommGroups.

                                                                  Equations
                                                                  Instances For
                                                                    def CategoryTheory.Iso.groupIsoToMulEquiv {X Y : Grp} (i : X Y) :
                                                                    X ≃* Y

                                                                    Build a MulEquiv from an isomorphism in the category Grp.

                                                                    Equations
                                                                    Instances For
                                                                      def CategoryTheory.Iso.addGroupIsoToAddEquiv {X Y : AddGrp} (i : X Y) :
                                                                      X ≃+ Y

                                                                      Build an addEquiv from an isomorphism in the category AddGroup

                                                                      Equations
                                                                      Instances For

                                                                        Build a MulEquiv from an isomorphism in the category CommGroup.

                                                                        Equations
                                                                        Instances For

                                                                          Build an AddEquiv from an isomorphism in the category AddCommGroup.

                                                                          Equations
                                                                          Instances For
                                                                            def mulEquivIsoGroupIso {X Y : Grp} :
                                                                            X ≃* Y X Y

                                                                            multiplicative equivalences between Groups are the same as (isomorphic to) isomorphisms in Grp

                                                                            Equations
                                                                            Instances For
                                                                              def addEquivIsoAddGroupIso {X Y : AddGrp} :
                                                                              X ≃+ Y X Y

                                                                              Additive equivalences between AddGroups are the same as (isomorphic to) isomorphisms in AddGrp.

                                                                              Equations
                                                                              Instances For
                                                                                def mulEquivIsoCommGroupIso {X Y : CommGrp} :
                                                                                X ≃* Y X Y

                                                                                Multiplicative equivalences between CommGroups are the same as (isomorphic to) isomorphisms in CommGrp.

                                                                                Equations
                                                                                Instances For
                                                                                  def addEquivIsoAddCommGroupIso {X Y : AddCommGrp} :
                                                                                  X ≃+ Y X Y

                                                                                  Additive equivalences between AddCommGroups are the same as (isomorphic to) isomorphisms in AddCommGrp.

                                                                                  Equations
                                                                                  Instances For

                                                                                    The (bundled) group of automorphisms of a type is isomorphic to the (bundled) group of permutations.

                                                                                    Equations
                                                                                    • One or more equations did not get rendered due to their size.
                                                                                    Instances For

                                                                                      The (unbundled) group of automorphisms of a type is MulEquiv to the (unbundled) group of permutations.

                                                                                      Equations
                                                                                      Instances For
                                                                                        @[reducible, inline]
                                                                                        abbrev GrpMax :
                                                                                        Type ((max u1 u2) + 1)

                                                                                        An alias for Grp.{max u v}, to deal around unification issues.

                                                                                        Equations
                                                                                        Instances For
                                                                                          @[reducible, inline]
                                                                                          abbrev GrpMaxAux :
                                                                                          Type ((max u1 u2) + 1)

                                                                                          An alias for AddGrp.{max u v}, to deal around unification issues.

                                                                                          Equations
                                                                                          Instances For
                                                                                            @[reducible, inline]
                                                                                            abbrev AddGrpMax :
                                                                                            Type ((max u1 u2) + 1)

                                                                                            An alias for AddGrp.{max u v}, to deal around unification issues.

                                                                                            Equations
                                                                                            Instances For
                                                                                              @[reducible, inline]
                                                                                              abbrev CommGrpMax :
                                                                                              Type ((max u1 u2) + 1)

                                                                                              An alias for CommGrp.{max u v}, to deal around unification issues.

                                                                                              Equations
                                                                                              Instances For
                                                                                                @[reducible, inline]
                                                                                                abbrev AddCommGrpMaxAux :
                                                                                                Type ((max u1 u2) + 1)

                                                                                                An alias for AddCommGrp.{max u v}, to deal around unification issues.

                                                                                                Equations
                                                                                                Instances For
                                                                                                  @[reducible, inline]
                                                                                                  abbrev AddCommGrpMax :
                                                                                                  Type ((max u1 u2) + 1)

                                                                                                  An alias for AddCommGrp.{max u v}, to deal around unification issues.

                                                                                                  Equations
                                                                                                  Instances For

                                                                                                    Deprecated lemmas for MonoidHom.comp and categorical identities.

                                                                                                    @[deprecated "Proven by `simp only [Grp.hom_id, comp_id]`" (since := "2025-01-28")]
                                                                                                    theorem MonoidHom.comp_id_grp {G : Grp} {H : Type u} [Monoid H] (f : G →* H) :
                                                                                                    @[deprecated "Proven by `simp only [Grp.hom_id, comp_id]`" (since := "2025-01-28")]
                                                                                                    @[deprecated "Proven by `simp only [Grp.hom_id, id_comp]`" (since := "2025-01-28")]
                                                                                                    theorem MonoidHom.id_grp_comp {G : Type u} [Monoid G] {H : Grp} (f : G →* H) :
                                                                                                    @[deprecated "Proven by `simp only [Grp.hom_id, id_comp]`" (since := "2025-01-28")]
                                                                                                    @[deprecated "Proven by `simp only [CommGrp.hom_id, comp_id]`" (since := "2025-01-28")]
                                                                                                    @[deprecated "Proven by `simp only [CommGrp.hom_id, comp_id]`" (since := "2025-01-28")]
                                                                                                    @[deprecated "Proven by `simp only [CommGrp.hom_id, id_comp]`" (since := "2025-01-28")]
                                                                                                    @[deprecated "Proven by `simp only [CommGrp.hom_id, id_comp]`" (since := "2025-01-28")]