Enriched ordinary categories #
If V
is a monoidal category, a V
-enriched category C
does not need
to be a category. However, when we have both Category C
and EnrichedCategory V C
,
we may require that the type of morphisms X ⟶ Y
in C
identify to
𝟙_ V ⟶ EnrichedCategory.Hom X Y
. This data shall be packaged in the
typeclass EnrichedOrdinaryCategory V C
.
In particular, if C
is a V
-enriched category, it is shown that
the "underlying" category ForgetEnrichment V C
is equipped with a
EnrichedOrdinaryCategory V C
instance.
Simplicial categories are implemented in AlgebraicTopology.SimplicialCategory.Basic
using an abbreviation for EnrichedOrdinaryCategory SSet C
.
An enriched ordinary category is a category C
that is also enriched
over a category V
in such a way that morphisms X ⟶ Y
in C
identify
to morphisms 𝟙_ V ⟶ (X ⟶[V] Y)
in V
.
- Hom : C → C → V
- id_comp (X Y : C) : CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.leftUnitor (CategoryTheory.EnrichedCategory.Hom X Y)).inv (CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.whiskerRight (CategoryTheory.EnrichedCategory.id X) (CategoryTheory.EnrichedCategory.Hom X Y)) (CategoryTheory.EnrichedCategory.comp X X Y)) = CategoryTheory.CategoryStruct.id (CategoryTheory.EnrichedCategory.Hom X Y)
- comp_id (X Y : C) : CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.rightUnitor (CategoryTheory.EnrichedCategory.Hom X Y)).inv (CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.whiskerLeft (CategoryTheory.EnrichedCategory.Hom X Y) (CategoryTheory.EnrichedCategory.id Y)) (CategoryTheory.EnrichedCategory.comp X Y Y)) = CategoryTheory.CategoryStruct.id (CategoryTheory.EnrichedCategory.Hom X Y)
- assoc (W X Y Z : C) : CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.associator (CategoryTheory.EnrichedCategory.Hom W X) (CategoryTheory.EnrichedCategory.Hom X Y) (CategoryTheory.EnrichedCategory.Hom Y Z)).inv (CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.whiskerRight (CategoryTheory.EnrichedCategory.comp W X Y) (CategoryTheory.EnrichedCategory.Hom Y Z)) (CategoryTheory.EnrichedCategory.comp W Y Z)) = CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.whiskerLeft (CategoryTheory.EnrichedCategory.Hom W X) (CategoryTheory.EnrichedCategory.comp X Y Z)) (CategoryTheory.EnrichedCategory.comp W X Z)
- homEquiv {X Y : C} : (X ⟶ Y) ≃ (𝟙_ V ⟶ CategoryTheory.EnrichedCategory.Hom X Y)
morphisms
X ⟶ Y
in the category identify morphisms𝟙_ V ⟶ (X ⟶[V] Y)
inV
- homEquiv_id (X : C) : CategoryTheory.EnrichedOrdinaryCategory.homEquiv (CategoryTheory.CategoryStruct.id X) = CategoryTheory.eId V X
- homEquiv_comp {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) : CategoryTheory.EnrichedOrdinaryCategory.homEquiv (CategoryTheory.CategoryStruct.comp f g) = CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.leftUnitor (𝟙_ V)).inv (CategoryTheory.CategoryStruct.comp (CategoryTheory.MonoidalCategory.tensorHom (CategoryTheory.EnrichedOrdinaryCategory.homEquiv f) (CategoryTheory.EnrichedOrdinaryCategory.homEquiv g)) (CategoryTheory.eComp V X Y Z))
Instances
The bijection (X ⟶ Y) ≃ (𝟙_ V ⟶ (X ⟶[V] Y))
given by a
EnrichedOrdinaryCategory
instance.
Equations
- CategoryTheory.eHomEquiv V = CategoryTheory.EnrichedOrdinaryCategory.homEquiv
Instances For
The morphism (X' ⟶[V] Y) ⟶ (X ⟶[V] Y)
induced by a morphism X ⟶ X'
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Whiskering commutes with the enriched composition.
The morphism (X ⟶[V] Y) ⟶ (X ⟶[V] Y')
induced by a morphism Y ⟶ Y'
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Whiskering commutes with the enriched composition.
Given an isomorphism α : Y ≅ Y₁
in C, the enriched composition map
eComp V X Y Z : (X ⟶[V] Y) ⊗ (Y ⟶[V] Z) ⟶ (X ⟶[V] Z)
factors through the V
object (X ⟶[V] Y₁) ⊗ (Y₁ ⟶[V] Z)
via the map defined by whiskering in the
middle with α.hom
and α.inv
.
The bifunctor Cᵒᵖ ⥤ C ⥤ V
which sends X : Cᵒᵖ
and Y : C
to X ⟶[V] Y
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- CategoryTheory.ForgetEnrichment.EnrichedOrdinaryCategory V = CategoryTheory.EnrichedOrdinaryCategory.mk (fun {X Y : CategoryTheory.ForgetEnrichment V D} => Equiv.refl (X ⟶ Y)) ⋯ ⋯