Category instances for Group, AddGroup, CommGroup, and AddCommGroup. #
We introduce the bundled categories:
Grp
AddGrp
CommGrp
AddCommGrp
along with the relevant forgetful functors between them, and to the bundled monoid categories.
Equations
- Grp.instCoeSortType = { coe := Grp.carrier }
Equations
- AddGrp.instCoeSortType = { coe := AddGrp.carrier }
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.
Use the ConcreteCategory.hom
projection for @[simps]
lemmas.
Equations
- Grp.Hom.Simps.hom X Y f = f.hom
Instances For
The results below duplicate the ConcreteCategory
simp lemmas, but we can keep them for dsimp
.
Alias of Grp.coe_comp
.
Alias of Grp.coe_id
.
Equations
- Grp.instInhabited = { default := Grp.of PUnit.{?u.1 + 1} }
Equations
- AddGrp.instInhabited = { default := AddGrp.of PUnit.{?u.1 + 1} }
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
- Grp.instCoeMonCat = { coe := (CategoryTheory.forget₂ Grp MonCat).obj }
Equations
- AddGrp.instCoeMonCat = { coe := (CategoryTheory.forget₂ AddGrp AddMonCat).obj }
Equations
- G.instOneHom H = { one := Grp.ofHom 1 }
Equations
- G.instZeroHom H = { zero := AddGrp.ofHom 0 }
Equations
- Grp.ofUnique G = i
Equations
- AddGrp.ofUnique G = i
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
The category of additive groups and group morphisms.
- carrier : Type u
The underlying type.
- str : AddCommGroup ↑self
Instances For
Ab
is an abbreviation for AddCommGroup
, for the sake of mathematicians' sanity.
Equations
Instances For
Equations
- CommGrp.instCoeSortType = { coe := CommGrp.carrier }
Equations
- AddCommGrp.instCoeSortType = { coe := AddCommGrp.carrier }
Construct a bundled CommGrp
from the underlying type and typeclass.
Equations
- CommGrp.of M = CommGrp.mk M
Instances For
Construct a bundled AddCommGrp
from the underlying type and typeclass.
Equations
Instances For
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.
Use the ConcreteCategory.hom
projection for @[simps]
lemmas.
Equations
- CommGrp.Hom.Simps.hom X Y f = f.hom
Instances For
Use the ConcreteCategory.hom
projection for @[simps]
lemmas.
Equations
- AddCommGrp.Hom.Simps.hom X Y f = f.hom
Instances For
The results below duplicate the ConcreteCategory
simp lemmas, but we can keep them for dsimp
.
Equations
- CommGrp.instInhabited = { default := CommGrp.of PUnit.{?u.1 + 1} }
Equations
- AddCommGrp.instInhabited = { default := AddCommGrp.of PUnit.{?u.1 + 1} }
Alias of CommGrp.coe_comp
.
Alias of CommGrp.coe_id
.
Equations
- CommGrp.ofUnique G = i
Equations
- AddCommGrp.ofUnique G = i
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
- CommGrp.instCoeGrp = { coe := (CategoryTheory.forget₂ CommGrp Grp).obj }
Equations
- AddCommGrp.instCoeAddGrp = { coe := (CategoryTheory.forget₂ AddCommGrp AddGrp).obj }
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
- CommGrp.instCoeCommMonCat = { coe := (CategoryTheory.forget₂ CommGrp CommMonCat).obj }
Equations
Equations
- G.instOneHom H = { one := CommGrp.ofHom 1 }
Equations
- G.instZeroHom H = { zero := AddCommGrp.ofHom 0 }
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
Any element of an abelian group gives a unique morphism from ℤ
sending
1
to that element.
Equations
- AddCommGrp.asHom g = AddCommGrp.ofHom ((zmultiplesHom ↑G) g)
Instances For
Build an isomorphism in the category Grp
from a MulEquiv
between Group
s.
Equations
- e.toGrpIso = { hom := Grp.ofHom e.toMonoidHom, inv := Grp.ofHom e.symm.toMonoidHom, hom_inv_id := ⋯, inv_hom_id := ⋯ }
Instances For
Build an isomorphism in the category AddGroup
from an AddEquiv
between AddGroup
s.
Equations
- e.toAddGrpIso = { hom := AddGrp.ofHom e.toAddMonoidHom, inv := AddGrp.ofHom e.symm.toAddMonoidHom, hom_inv_id := ⋯, inv_hom_id := ⋯ }
Instances For
Build an isomorphism in the category CommGrp
from a MulEquiv
between CommGroup
s.
Equations
- e.toCommGrpIso = { hom := CommGrp.ofHom e.toMonoidHom, inv := CommGrp.ofHom e.symm.toMonoidHom, hom_inv_id := ⋯, inv_hom_id := ⋯ }
Instances For
Build an isomorphism in the category AddCommGrp
from an AddEquiv
between AddCommGroup
s.
Equations
- e.toAddCommGrpIso = { hom := AddCommGrp.ofHom e.toAddMonoidHom, inv := AddCommGrp.ofHom e.symm.toAddMonoidHom, hom_inv_id := ⋯, inv_hom_id := ⋯ }
Instances For
Build a MulEquiv
from an isomorphism in the category Grp
.
Equations
- i.groupIsoToMulEquiv = (Grp.Hom.hom i.hom).toMulEquiv (Grp.Hom.hom i.inv) ⋯ ⋯
Instances For
Build an addEquiv
from an isomorphism in the category AddGroup
Equations
- i.addGroupIsoToAddEquiv = (AddGrp.Hom.hom i.hom).toAddEquiv (AddGrp.Hom.hom i.inv) ⋯ ⋯
Instances For
Build a MulEquiv
from an isomorphism in the category CommGroup
.
Equations
- i.commGroupIsoToMulEquiv = (CommGrp.Hom.hom i.hom).toMulEquiv (CommGrp.Hom.hom i.inv) ⋯ ⋯
Instances For
Build an AddEquiv
from an isomorphism in the category AddCommGroup
.
Equations
- i.addCommGroupIsoToAddEquiv = (AddCommGrp.Hom.hom i.hom).toAddEquiv (AddCommGrp.Hom.hom i.inv) ⋯ ⋯
Instances For
multiplicative equivalences between Group
s are the same as (isomorphic to) isomorphisms
in Grp
Equations
- mulEquivIsoGroupIso = { hom := fun (e : ↑X ≃* ↑Y) => e.toGrpIso, inv := fun (i : X ≅ Y) => i.groupIsoToMulEquiv, hom_inv_id := ⋯, inv_hom_id := ⋯ }
Instances For
Additive equivalences between AddGroup
s are the same
as (isomorphic to) isomorphisms in AddGrp
.
Equations
- addEquivIsoAddGroupIso = { hom := fun (e : ↑X ≃+ ↑Y) => e.toAddGrpIso, inv := fun (i : X ≅ Y) => i.addGroupIsoToAddEquiv, hom_inv_id := ⋯, inv_hom_id := ⋯ }
Instances For
Multiplicative equivalences between CommGroup
s are the same as (isomorphic to) isomorphisms
in CommGrp
.
Equations
- mulEquivIsoCommGroupIso = { hom := fun (e : ↑X ≃* ↑Y) => e.toCommGrpIso, inv := fun (i : X ≅ Y) => i.commGroupIsoToMulEquiv, hom_inv_id := ⋯, inv_hom_id := ⋯ }
Instances For
Additive equivalences between AddCommGroup
s are
the same as (isomorphic to) isomorphisms in AddCommGrp
.
Equations
- addEquivIsoAddCommGroupIso = { hom := fun (e : ↑X ≃+ ↑Y) => e.toAddCommGrpIso, inv := fun (i : X ≅ Y) => i.addCommGroupIsoToAddEquiv, hom_inv_id := ⋯, inv_hom_id := ⋯ }
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.
Instances For
An alias for CommGrp.{max u v}
, to deal around unification issues.
Equations
Instances For
An alias for AddCommGrp.{max u v}
, to deal around unification issues.
Equations
Instances For
An alias for AddCommGrp.{max u v}
, to deal around unification issues.
Equations
Instances For
Deprecated lemmas for MonoidHom.comp
and categorical identities.