Concrete categories #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
A concrete category is a category C
with a fixed faithful functor
forget : C ⥤ Type*
. We define concrete categories using class concrete_category
. In particular, we impose no restrictions on the
carrier type C
, so Type
is a concrete category with the identity
forgetful functor.
Each concrete category C
comes with a canonical faithful functor
forget C : C ⥤ Type*
. We say that a concrete category C
admits a
forgetful functor to a concrete category D
, if it has a functor
forget₂ C D : C ⥤ D
such that (forget₂ C D) ⋙ (forget D) = forget C
,
see class has_forget₂
. Due to faithful.div_comp
, it suffices
to verify that forget₂.obj
and forget₂.map
agree with the equality
above; then forget₂
will satisfy the functor laws automatically, see
has_forget₂.mk'
.
Two classes helping construct concrete categories in the two most
common cases are provided in the files bundled_hom
and
unbundled_hom
, see their documentation for details.
References #
See Ahrens and Lumsdaine, Displayed Categories for related work.
- forget : C ⥤ Type ?
- forget_faithful : category_theory.faithful (category_theory.concrete_category.forget C)
A concrete category is a category C
with a fixed faithful functor forget : C ⥤ Type
.
Note that concrete_category
potentially depends on three independent universe levels,
- the universe level
w
appearing inforget : C ⥤ Type w
- the universe level
v
of the morphisms (i.e. we have acategory.{v} C
) - the universe level
u
of the objects (i.eC : Type u
) They are specified that order, to avoid unnecessary universe annotations.
Instances of this typeclass
- category_theory.concrete_category.types
- category_theory.induced_category.concrete_category
- category_theory.full_subcategory.concrete_category
- category_theory.bundled_hom.concrete_category
- Mon.concrete_category
- AddMon.concrete_category
- CommMon.concrete_category
- AddCommMon.concrete_category
- Group.concrete_category
- AddGroup.concrete_category
- CommGroup.concrete_category
- AddCommGroup.concrete_category
- Action.category_theory.concrete_category
- Module.Module_concrete_category
- Rep.concrete_category
- fgModule.concrete_category
- fdRep.concrete_category
- category_theory.graded_object.category_theory.concrete_category
- Fintype.concrete_category_Fintype
- Preord.concrete_category
- PartOrd.concrete_category
- FinPartOrd.concrete_category
- Lat.category_theory.concrete_category
- LinOrd.concrete_category
- NonemptyFinLinOrd.concrete_category
- simplex_category.category_theory.concrete_category
- Top.concrete_category
- Meas.concrete_category
- UniformSpace.concrete_category
- CpltSepUniformSpace.concrete_category
- SemiRing.concrete_category
- Ring.concrete_category
- CommSemiRing.concrete_category
- CommRing.concrete_category
- TopCommRing.category_theory.concrete_category
- Born.category_theory.concrete_category
- CompHaus.concrete_category
- Frm.concrete_category
- Profinite.concrete_category
- Compactum.category_theory.concrete_category
- category_theory.differential_object.concrete_category_of_differential_objects
- Pointed.concrete_category
- Bipointed.concrete_category
- Twop.concrete_category
- Algebra.category_theory.concrete_category
- SemiNormedGroup.concrete_category
- SemiNormedGroup₁.category_theory.concrete_category
- BddOrd.concrete_category
- SemilatSup.category_theory.concrete_category
- SemilatInf.category_theory.concrete_category
- BddLat.category_theory.concrete_category
- DistLat.concrete_category
- BddDistLat.category_theory.concrete_category
- HeytAlg.concrete_category
- BoolAlg.category_theory.concrete_category
- FinBddDistLat.concrete_category
- FinBoolAlg.concrete_category
- CompleteLat.category_theory.concrete_category
- ωCPO.concrete_category
- GroupWithZero.category_theory.concrete_category
- BoolRing.concrete_category
- Magma.concrete_category
- AddMagma.concrete_category
- Semigroup.concrete_category
- AddSemigroup.concrete_category
Instances of other typeclasses for category_theory.concrete_category
- category_theory.concrete_category.has_sizeof_inst
The forgetful functor from a concrete category to Type u
.
Equations
Provide a coercion to Type u
for a concrete category. This is not marked as an instance
as it could potentially apply to every type, and so is too expensive in typeclass search.
You can use it on particular examples as:
instance : has_coe_to_sort X := concrete_category.has_coe_to_sort X
Usually a bundled hom structure already has a coercion to function that works with different universes. So we don't use this as a global instance.
Equations
- category_theory.concrete_category.has_coe_to_fun = {coe := λ (f : X ⟶ Y), (category_theory.forget C).map f}
In any concrete category, we can test equality of morphisms by pointwise evaluations.
Analogue of congr_fun h x
,
when h : f = g
is an equality between morphisms in a concrete category.
In any concrete category, injective morphisms are monomorphisms.
In any concrete category, surjective morphisms are epimorphisms.
- forget₂ : C ⥤ D
- forget_comp : category_theory.has_forget₂.forget₂ ⋙ category_theory.forget D = category_theory.forget C . "obviously"
has_forget₂ C D
, where C
and D
are both concrete categories, provides a functor
forget₂ C D : C ⥤ D
and a proof that forget₂ ⋙ (forget D) = forget C
.
Instances of this typeclass
- category_theory.induced_category.has_forget₂
- category_theory.full_subcategory.has_forget₂
- category_theory.bundled_hom.forget₂
- CommMon.has_forget_to_Mon
- AddCommMon.has_forget_to_AddMon
- Group.has_forget_to_Mon
- AddGroup.has_forget_to_AddMon
- CommGroup.has_forget_to_Group
- AddCommGroup.has_forget_to_AddGroup
- CommGroup.has_forget_to_CommMon
- AddCommGroup.has_forget_to_AddCommMon
- Action.has_forget_to_V
- Module.has_forget_to_AddCommGroup
- fgModule.Module.category_theory.has_forget₂
- fdRep.Rep.category_theory.has_forget₂
- category_theory.graded_object.category_theory.has_forget₂
- PartOrd.has_forget_to_Preord
- FinPartOrd.has_forget_to_PartOrd
- FinPartOrd.has_forget_to_Fintype
- Lat.has_forget_to_PartOrd
- LinOrd.has_forget_to_Lat
- NonemptyFinLinOrd.has_forget_to_LinOrd
- NonemptyFinLinOrd.has_forget_to_FinPartOrd
- Top.has_forget_to_Meas
- UniformSpace.has_forget_to_Top
- CpltSepUniformSpace.has_forget_to_UniformSpace
- SemiRing.has_forget_to_Mon
- SemiRing.has_forget_to_AddCommMon
- Ring.has_forget_to_SemiRing
- Ring.has_forget_to_AddCommGroup
- CommSemiRing.has_forget_to_SemiRing
- CommSemiRing.has_forget_to_CommMon
- CommRing.has_forget_to_Ring
- CommRing.has_forget_to_CommSemiRing
- TopCommRing.has_forget_to_CommRing
- TopCommRing.has_forget_to_Top
- Frm.has_forget_to_Lat
- Profinite.has_forget₂
- category_theory.differential_object.category_theory.has_forget₂
- Twop.has_forget_to_Bipointed
- Algebra.has_forget_to_Ring
- Algebra.has_forget_to_Module
- SemiNormedGroup₁.SemiNormedGroup.category_theory.has_forget₂
- BddOrd.has_forget_to_PartOrd
- BddOrd.has_forget_to_Bipointed
- SemilatSup.has_forget_to_PartOrd
- SemilatInf.has_forget_to_PartOrd
- BddLat.has_forget_to_BddOrd
- BddLat.has_forget_to_Lat
- BddLat.has_forget_to_SemilatSup
- BddLat.has_forget_to_SemilatInf
- DistLat.has_forget_to_Lat
- BddDistLat.has_forget_to_DistLat
- BddDistLat.has_forget_to_BddLat
- HeytAlg.has_forget_to_Lat
- BoolAlg.has_forget_to_BddDistLat
- BoolAlg.has_forget_to_HeytAlg
- FinBddDistLat.has_forget_to_BddDistLat
- FinBddDistLat.has_forget_to_FinPartOrd
- FinBoolAlg.has_forget_to_BoolAlg
- FinBoolAlg.has_forget_to_FinBddDistLat
- FinBoolAlg.has_forget_to_FinPartOrd
- CompleteLat.has_forget_to_BddLat
- GroupWithZero.has_forget_to_Bipointed
- GroupWithZero.has_forget_to_Mon
- BoolRing.has_forget_to_CommRing
- BoolRing.has_forget_to_BoolAlg
- BoolAlg.has_forget_to_BoolRing
- Semigroup.has_forget_to_Magma
- AddSemigroup.has_forget_to_AddMagma
- has_forget_to_Semigroup
- has_forget_to_AddSemigroup
Instances of other typeclasses for category_theory.has_forget₂
- category_theory.has_forget₂.has_sizeof_inst
The forgetful functor C ⥤ D
between concrete categories for which we have an instance
has_forget₂ C
.
Equations
Equations
Equations
In order to construct a “partially forgetting” functor, we do not need to verify functor laws;
it suffices to ensure that compositions agree with forget₂ C D ⋙ forget D = forget C
.
Equations
- category_theory.has_forget₂.mk' obj h_obj map h_map = {forget₂ := category_theory.faithful.div (category_theory.forget C) (category_theory.forget D) (λ (X : C), obj X) h_obj (λ (X Y : C) (f : X ⟶ Y), map f) h_map, forget_comp := _}
Every forgetful functor factors through the identity functor. This is not a global instance as it is prone to creating type class resolution loops.
Equations
- category_theory.has_forget_to_Type C = {forget₂ := category_theory.forget C _inst_2, forget_comp := _}