Category instances for monoid, add_monoid, comm_monoid, and add_comm_monoid.
We introduce the bundled categories:
Mon
AddMon
CommMon
AddCommMon
along with the relevant forgetful functors between them.
Equations
- Mon.bundled_hom = {to_fun := monoid_hom.to_fun, id := monoid_hom.id, comp := monoid_hom.comp, hom_ext := monoid_hom.coe_inj, id_to_fun := Mon.bundled_hom._proof_1, comp_to_fun := Mon.bundled_hom._proof_2}
Construct a bundled Mon
from the underlying type and typeclass.
Equations
The category of additive commutative monoids and monoid morphisms.
The category of commutative monoids and monoid morphisms.
Equations
Construct a bundled CommMon
from the underlying type and typeclass.
Equations
Construct a bundled AddCommMon
from the underlying type and typeclass.
Equations
Equations
- M.comm_monoid = M.str
Build an isomorphism in the category AddMon
from
an add_equiv
between add_monoid
s.
Build an isomorphism in the category Mon
from a mul_equiv
between monoid
s.
Equations
- e.to_Mon_iso = {hom := e.to_monoid_hom, inv := e.symm.to_monoid_hom, hom_inv_id' := _, inv_hom_id' := _}
Build an isomorphism in the category CommMon
from a mul_equiv
between comm_monoid
s.
Equations
- e.to_CommMon_iso = {hom := e.to_monoid_hom, inv := e.symm.to_monoid_hom, hom_inv_id' := _, inv_hom_id' := _}
Build an isomorphism in the category AddCommMon
from
an add_equiv
between add_comm_monoid
s.
Build an add_equiv
from an isomorphism in the category
AddCommMon
.
multiplicative equivalences between monoid
s are the same as (isomorphic to) isomorphisms
in Mon
Equations
- mul_equiv_iso_Mon_iso = {hom := λ (e : X ≃* Y), e.to_Mon_iso, inv := λ (i : Mon.of X ≅ Mon.of Y), i.Mon_iso_to_mul_equiv, hom_inv_id' := _, inv_hom_id' := _}
additive equivalences between add_monoid
s are the same
as (isomorphic to) isomorphisms in AddMon
additive equivalences between add_comm_monoid
s are
the same as (isomorphic to) isomorphisms in AddCommMon
multiplicative equivalences between comm_monoid
s are the same as (isomorphic to) isomorphisms
in CommMon
Equations
- mul_equiv_iso_CommMon_iso = {hom := λ (e : X ≃* Y), e.to_CommMon_iso, inv := λ (i : CommMon.of X ≅ CommMon.of Y), i.CommMon_iso_to_mul_equiv, hom_inv_id' := _, inv_hom_id' := _}
Equations
- Mon.forget_reflects_isos = {reflects := λ (X Y : Mon) (f : X ⟶ Y) (_x : category_theory.is_iso ((category_theory.forget Mon).map f)), let i : (category_theory.forget Mon).obj X ≅ (category_theory.forget Mon).obj Y := category_theory.as_iso ((category_theory.forget Mon).map f), e : ↥X ≃* ↥Y := {to_fun := f.to_fun, inv_fun := i.to_equiv.inv_fun, left_inv := _, right_inv := _, map_mul' := _} in {inv := e.to_Mon_iso.inv, hom_inv_id' := _, inv_hom_id' := _}}
Equations
- CommMon.forget_reflects_isos = {reflects := λ (X Y : CommMon) (f : X ⟶ Y) (_x : category_theory.is_iso ((category_theory.forget CommMon).map f)), let i : (category_theory.forget CommMon).obj X ≅ (category_theory.forget CommMon).obj Y := category_theory.as_iso ((category_theory.forget CommMon).map f), e : ↥X ≃* ↥Y := {to_fun := f.to_fun, inv_fun := i.to_equiv.inv_fun, left_inv := _, right_inv := _, map_mul' := _} in {inv := e.to_CommMon_iso.inv, hom_inv_id' := _, inv_hom_id' := _}}
Once we've shown that the forgetful functors to type reflect isomorphisms,
we automatically obtain that the forget₂
functors between our concrete categories
reflect isomorphisms.