The strict bicategory associated to a Cat-enriched category #
If C
is a type with a EnrichedCategory Cat C
structure, then it has hom-categories, whose
objects define 1-dimensional arrows on C
and whose morphisms define 2-dimensional arrows between
these. The enriched category axioms equip this data with the structure of a strict bicategory.
We define a type alias CatEnriched C
for a type C
with a EnrichedCategory Cat C
structure. We
provide this with an instance of a strict bicategory structure constructing
Bicategory.Strict (CatEnriched C)
.
If C
is a type with a EnrichedOrdinaryCategory Cat C
structure, then it has an Enrichred Cat C
structure, so the previous construction would again produce a strict bicategory. However, in this
setting C
is also given a Category C
structure, together with an equivalence between this
category and the underlying category of the Enriched Cat C
, and in examples the given category
structure is the preferred one.
Thus, we define a type alias CatEnrichedOrdinary C
for a type C
with an
EnrichedOrdinaryCategory Cat C
structure. We provide this with an instance of a strict bicategory
structure extending the category structure provided by the given instance Category C
constructing
Bicategory.Strict (CatEnrichedOrdinary C)
.
A type synonym for C
, which should come equipped with a Cat
-enriched category structure.
This converts it to a strict bicategory where Category (X ⟶ Y)
is (X ⟶[Cat] Y)
.
Equations
Instances For
Any enriched category has an underlying category structure defined by ForgetEnrichment
.
This is equivalent but not definitionally equal the category structure constructed here, which is
more canonically associated to the data of an EnrichedCategory Cat
structure.
Equations
- One or more equations did not get rendered due to their size.
The horizonal composition on 2-morphisms is defined using the action on arrows of the composition bifunctor from the enriched category structure.
Equations
- CategoryTheory.CatEnriched.hComp η θ = (CategoryTheory.eComp CategoryTheory.Cat a b c).map (η, θ)
Instances For
The interchange law for horizontal and vertical composition of 2-cells in a bicategory.
The action on objects of the EnrichedCategory Cat
coherences proves the category axioms.
Equations
- CategoryTheory.CatEnriched.instCategory = { toCategoryStruct := CategoryTheory.CatEnriched.instCategoryStruct, id_comp := ⋯, comp_id := ⋯, assoc := ⋯ }
The category instance on CatEnriched C
promotes it to a Cat
enriched ordinary
category.
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.
As the associator and left and right unitors are defined as eqToIso of category axioms, the
bicategory structure on CatEnriched C
is strict.
A type synonym for C
, which should come equipped with a Cat
-enriched category structure.
This converts it to a strict bicategory where Category (X ⟶ Y)
is (X ⟶[Cat] Y)
.
Equations
Instances For
The forgetful map from the type alias associated to EnrichedOrdinaryCategory Cat C
and the
type alias associated to EnrichedCategory Cat C
is the identity on underlying types.
Instances For
The hom-types in a Cat
-enriched ordinary category are equivalent to the types underlying the
hom-categories.
Equations
Instances For
The 2-cells between a parallel pair of 1-cells f g
in CatEnrichedOrdinary C
are defined to
be the morphisms in the hom-categories provided by the EnrichedCategory Cat C
structure between
the corresponding objects.
- mk' :: (
- )
Instances For
Equations
- CategoryTheory.CatEnrichedOrdinary.instQuiverHom = { Hom := fun (f g : X ⟶ Y) => CategoryTheory.CatEnrichedOrdinary.Hom f g }
A 2-cell in CatEnrichedOrdinary C
has a corresponding "base" 2-cell in CatEnriched C
.
Equations
Instances For
A 2-cell in CatEnriched C
can be "made" into a 2-cell in CatEnrichedOrdinary C
.
Equations
- CategoryTheory.CatEnrichedOrdinary.Hom.mk α = { base' := α }
Instances For
Equations
- One or more equations did not get rendered due to their size.
A Cat
-enriched ordinary category comes with hom-categories X ⟶[Cat] Y
whose underlying type
of objects is equivalent to the type X ⟶ Y
defined by the category structure on C
. The following
definition transfers the category structure to the latter type of objects.
Equations
- CategoryTheory.CatEnrichedOrdinary.instCategoryHom = { toCategoryStruct := CategoryTheory.CatEnrichedOrdinary.instCategoryStructHom, id_comp := ⋯, comp_id := ⋯, assoc := ⋯ }
The horizonal composition on 2-morphisms is defined using the action on arrows of the composition bifunctor from the enriched category structure.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The interchange law for horizontal and vertical composition of 2-cells in a bicategory.
Equations
- One or more equations did not get rendered due to their size.
As the associator and left and right unitors are defined as eqToIso of category axioms, the
bicategory structure on CatEnrichedOrdinary C
is strict.