Day functors #
In this file, given a monoidal category C
and a monoidal category V
,
we define a basic type synonym DayFunctor C V
(denoted C ⊛⥤ D
)
for the category C ⥤ V
and endow it with the monoidal structure coming
from Day convolution. Such a setup is necessary as by default,
the MonoidalCategory
instance on C ⥤ V
is the "pointwise" one,
where the tensor product of F
and G
is the functor x ↦ F.obj x ⊗ G.obj x
.
TODOs #
- Given a
LawfulDayConvolutionMonoidalCategoryStruct C V D
, show that ι induce a monoidal functorD ⥤ (C ⊛⥤ V)
. - Specialize to the case
V := Type _
, and prove a universal property stating that for every monoidal categoryW
with suitable colimits, colimit-preserving monoidal functors(Cᵒᵖ ⊛⥤ Type u) ⥤ W
are equivalent to to monoidal functorsC ⥤ W
. Show that the Yoneda embedding is monoidal.
DayFunctor C V
is a type synonym for C ⥤ V
, implemented as a one-field
structure.
- functor : Functor C V
the underlying functor.
Instances For
Notation for DayFunctor
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Morphisms of Day functors are natural transformations of the underlying functors.
the underlying natural transformation
Instances For
Equations
- One or more equations did not get rendered due to their size.
The tautological equivalence of categories between C ⥤ V
and C ⊛⥤ V
.
Equations
- One or more equations did not get rendered due to their size.
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.
The unit transformation exhibiting (F ⊗ G).functor
as a left Kan extension of
F.functor ⊠ G.functor
along tensor C
.
Equations
Instances For
A natural transformation F.functor ⊠ G.functor ⟶ tensor C ⋙ H.functor
defines a morphism F ⨂ G ⟶ H
.
Equations
- CategoryTheory.MonoidalCategory.DayFunctor.tensorDesc α = { natTrans := (CategoryTheory.MonoidalCategoryStruct.tensorObj F G).functor.descOfIsLeftKanExtension (F.η G) H.functor α }
Instances For
An abstract isomorphism between (F ⊗ G).functor
and the generic pointwise
left Kan extension of F.functor ⊠ G.functor
along the
Equations
- One or more equations did not get rendered due to their size.
Instances For
The canonical map 𝟙_ V ⟶ (𝟙_ (C ⊛⥤ V)).functor.obj (𝟙_ C)
that exhibits (𝟙_ (C ⊛⥤ V)).functor
as a Day convolution unit.
Equations
Instances For
The reinterpretation of ν
as a natural transformation.
Equations
- CategoryTheory.MonoidalCategory.DayFunctor.νNatTrans C V = { app := fun (x : CategoryTheory.Discrete PUnit.{1}) => CategoryTheory.MonoidalCategory.DayFunctor.ν C V, naturality := ⋯ }
Instances For
Given F : C ⊛⥤ V
, a morphism 𝟙_ V ⟶ F.functor.obj (𝟙_ C)
induces a
(unique) morphism 𝟙_ (C ⊛⥤ V) ⟶ F
.
Equations
- One or more equations did not get rendered due to their size.