Day convolution monoidal structure #
Given functors F G : C ⥤ V
between two monoidal categories,
this file defines a typeclass DayConvolution
on functors F
G
that contains
a functor F ⊛ G
, as well as the required data to exhibit F ⊛ G
as a pointwise
left Kan extension of F ⊠ G
(see Mathlib/CategoryTheory/Monoidal/ExternalProduct/Basic.lean
for the definition) along the tensor product of C
.
Such a functor is called a Day convolution of F
and G
, and
although we do not show it yet, this operation defines a monoidal structure on C ⥤ V
.
We also define a typeclass DayConvolutionUnit
on a functor U : C ⥤ V
that bundle the data
required to make it a unit for the Day convolution monoidal structure: said data is that of
a map 𝟙_ V ⟶ U.obj (𝟙_ C)
that exhibits U
as a pointwise left Kan extension of
fromPUnit (𝟙_ V)
along fromPUnit (𝟙_ C)
.
References #
TODOs (@robin-carlier) #
- Braided/symmetric case.
- Case where
V
is closed. - Define a typeclass
DayConvolutionMonoidalCategory
extendingMonoidalCategory
- Characterization of lax monoidal functors out of a day convolution monoidal category.
- Case
V = Type u
and its universal property.
A DayConvolution
structure on functors F G : C ⥤ V
is the data of
a functor F ⊛ G : C ⥤ V
, along with a unit F ⊠ G ⟶ tensor C ⋙ F ⊛ G
that exhibits this functor as a pointwise left Kan extension of F ⊠ G
along
tensor C
. This is a class
used to prove various property of such extensions,
but registering global instances of this class is probably a bad idea.
- convolution : Functor C V
The chosen convolution between the functors. Denoted
F ⊛ G
. The chosen convolution between the functors.
- isPointwiseLeftKanExtensionUnit : (Functor.LeftExtension.mk (convolution F G) (unit F G)).IsPointwiseLeftKanExtension
The transformation
unit
exhibitsF ⊛ G
as a pointwise left Kan extension ofF ⊠ G
alongtensor C
.
Instances
A notation for the Day convolution of two functors.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Two day convolution structures on the same functors gives an isomorphic functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The morphism between day convolutions (provided they exist) induced by a pair of morphisms.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The universal property of left Kan extensions characterizes the functor
corepresented by F ⊛ G
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Use the fact that (F ⊛ G).obj c
is a colimit to characterize morphisms out of it at a
point.
The CorepresentableBy
structure asserting that the Type-valued functor
Y ↦ (F ⊠ G ⊠ H ⟶ (𝟭 C).prod (tensor C) ⋙ tensor C ⋙ Y)
is corepresented by
F ⊛ G ⊛ H
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The CorepresentableBy
structure asserting that the Type-valued functor
Y ↦ ((F ⊠ G) ⊠ H ⟶ (tensor C).prod (𝟭 C) ⋙ tensor C ⋙ Y)
is corepresented by
(F ⊛ G) ⊛ H
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The isomorphism of functors between
((F ⊠ G) ⊠ H ⟶ (tensor C).prod (𝟭 C) ⋙ tensor C ⋙ -)
and
(F ⊠ G ⊠ H ⟶ (𝟭 C).prod (tensor C) ⋙ tensor C ⋙ -)
that coresponsds to the associator
isomorphism for Day convolution through corepresentableBy₂
and corepresentableBy₂
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The asociator isomorphism for Day convolution
Equations
- One or more equations did not get rendered due to their size.
Instances For
Characterizing the forward direction of the associator isomorphism with respect to the unit transformations.
Characterizing the inverse direction of the associator with respect to the unit transformations
A DayConvolutionUnit
structure on a functor C ⥤ V
is the data of a pointwise
left Kan extension of fromPUnit (𝟙_ V)
along fromPUnit (𝟙_ C)
. Again, this is
made a class to ease proofs when constructing DayConvolutionMonoidalCategory
structures, but one
should avoid registering it globally.
A "canonical" structure map
𝟙_ V ⟶ F.obj (𝟙_ C)
that defines a natural transformationfromPUnit (𝟙_ V) ⟶ fromPUnit (𝟙_ C) ⋙ F
.- isPointwiseLeftKanExtensionCan : (Functor.LeftExtension.mk F { app := fun (x : Discrete PUnit.{1}) => can, naturality := ⋯ }).IsPointwiseLeftKanExtension
The canonical map
𝟙_ V ⟶ F.obj (𝟙_ C)
exhibitsF
as a pointwise left kan extension offromPUnit.{0} 𝟙_ V
alongfromPUnit.{0} 𝟙_ C
.
Instances
A shorthand for the natural transformation of functors out of PUnit defined by
the canonical morphism 𝟙_ V ⟶ U.obj (𝟙_ C)
when U
is a unit for Day convolution.
Equations
- CategoryTheory.MonoidalCategory.DayConvolutionUnit.φ U = { app := fun (x : CategoryTheory.Discrete PUnit.{1}) => CategoryTheory.MonoidalCategory.DayConvolutionUnit.can, naturality := ⋯ }
Instances For
Since a convolution unit is a pointwise left Kan extension, maps out of it at any object are uniquely characterized.
A CorepresentableBy
structure that characterizes maps out of U ⊛ F
by leveraging the fact that U ⊠ F
is a left Kan extension of (fromPUnit 𝟙_ V) ⊠ F
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A CorepresentableBy
structure that characterizes maps out of F ⊛ U
by
leveraging the fact that F ⊠ U
is a left Kan extension of F ⊠ (fromPUnit 𝟙_ V)
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The isomorphism of corepresentable functors that defines the left unitor for Day convolution.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The isomorphism of corepresentable functors that defines the right unitor for Day convolution.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The left unitor isomorphism for Day convolution.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The right unitor isomorphism for Day convolution.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Characterizing the forward direction of leftUnitor
via the universal maps.
Characterizing the forward direction of rightUnitor
via the universal maps.