Documentation

Mathlib.CategoryTheory.Monoidal.DayConvolution

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).

The main way to assert that a given monoidal category structure on a category D arises as a "Day convolution monoidal structure" is given by the typeclass LawfulDayConvolutionMonoidalCategoryStruct C V D, which bundles the data and equations needed to exhibit D as a monoidal full subcategory of C ⥤ V if the latter were to have the Day convolution monoidal structure. The definition monoidalOfLawfulDayConvolutionMonoidalCategoryStruct promotes (under suitable assumptions on V) a LawfulDayConvolutionMonoidalCategoryStruct C V D to a monoidal structure.

References #

TODOs (@robin-carlier) #

class CategoryTheory.MonoidalCategory.DayConvolution {C : Type u₁} [Category.{v₁, u₁} C] {V : Type u₂} [Category.{v₂, u₂} V] [MonoidalCategory C] [MonoidalCategory V] (F G : Functor C V) :
Type (max (max (max u₁ u₂) v₁) v₂)

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.

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
        @[simp]
        theorem CategoryTheory.MonoidalCategory.DayConvolution.unit_naturality {C : Type u₁} [Category.{v₁, u₁} C] {V : Type u₂} [Category.{v₂, u₂} V] [MonoidalCategory C] [MonoidalCategory V] (F G : Functor C V) [DayConvolution F G] {x x' y y' : C} (f : x x') (g : y y') :
        CategoryStruct.comp (tensorHom (F.map f) (G.map g)) ((unit F G).app (x', y')) = CategoryStruct.comp ((unit F G).app (x, y)) ((convolution F G).map (tensorHom f g))
        @[simp]
        theorem CategoryTheory.MonoidalCategory.DayConvolution.unit_naturality_assoc {C : Type u₁} [Category.{v₁, u₁} C] {V : Type u₂} [Category.{v₂, u₂} V] [MonoidalCategory C] [MonoidalCategory V] (F G : Functor C V) [DayConvolution F G] {x x' y y' : C} (f : x x') (g : y y') {Z : V} (h : (convolution F G).obj ((tensor C).obj (x', y')) Z) :

        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
          @[simp]
          theorem CategoryTheory.MonoidalCategory.DayConvolution.unit_app_map_app {C : Type u₁} [Category.{v₁, u₁} C] {V : Type u₂} [Category.{v₂, u₂} V] [MonoidalCategory C] [MonoidalCategory V] {F G : Functor C V} [DayConvolution F G] {F' G' : Functor C V} [DayConvolution F' G'] (f : F F') (g : G G') (x y : C) :
          CategoryStruct.comp ((unit F G).app (x, y)) ((map f g).app (tensorObj x y)) = CategoryStruct.comp (tensorHom (f.app x) (g.app y)) ((unit F' G').app (x, y))
          @[simp]
          theorem CategoryTheory.MonoidalCategory.DayConvolution.unit_app_map_app_assoc {C : Type u₁} [Category.{v₁, u₁} C] {V : Type u₂} [Category.{v₂, u₂} V] [MonoidalCategory C] [MonoidalCategory V] {F G : Functor C V} [DayConvolution F G] {F' G' : Functor C V} [DayConvolution F' G'] (f : F F') (g : G G') (x y : C) {Z : V} (h : (convolution F' G').obj (tensorObj x y) Z) :

          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
                  @[simp]
                  theorem CategoryTheory.MonoidalCategory.DayConvolution.associatorCorepresentingIso_inv_app_app {C : Type u₁} [Category.{v₁, u₁} C] {V : Type u₂} [Category.{v₂, u₂} V] [MonoidalCategory C] [MonoidalCategory V] (F G H X : Functor C V) (a✝ : (((Functor.whiskeringLeft (C × C × C) C V).obj (((Functor.id C).prod (tensor C)).comp (tensor C))).comp (coyoneda.obj (Opposite.op (externalProduct F (externalProduct G H))))).obj X) (X✝ : (C × C) × C) :
                  ((associatorCorepresentingIso F G H).inv.app X a✝).app X✝ = CategoryStruct.comp (whiskerRight (whiskerRight (F.map ((prod.associativity C C C).unit.app X✝).1.1) (G.obj X✝.1.2)) (H.obj X✝.2)) (CategoryStruct.comp (MonoidalCategoryStruct.associator (F.obj X✝.1.1) (G.obj X✝.1.2) (H.obj X✝.2)).hom (CategoryStruct.comp (whiskerLeft (F.obj X✝.1.1) (whiskerRight (G.map ((prod.associativity C C C).unit.app X✝).1.2) (H.obj X✝.2))) (CategoryStruct.comp (whiskerLeft (F.obj X✝.1.1) (whiskerLeft (G.obj X✝.1.2) (H.map ((prod.associativity C C C).unit.app X✝).2))) (CategoryStruct.comp (a✝.app (X✝.1.1, X✝.1.2, X✝.2)) (CategoryStruct.comp (X.map (tensorHom ((prod.associativity C C C).unitInv.app X✝).1.1 (tensorHom ((prod.associativity C C C).unitInv.app X✝).1.2 ((prod.associativity C C C).unitInv.app X✝).2))) (X.map (MonoidalCategoryStruct.associator X✝.1.1 X✝.1.2 X✝.2).inv))))))

                  The asociator isomorphism for Day convolution

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]

                    Characterizing the forward direction of the associator isomorphism with respect to the unit transformations.

                    @[simp]

                    Characterizing the inverse direction of the associator with respect to the unit transformations

                    class CategoryTheory.MonoidalCategory.DayConvolutionUnit {C : Type u₁} [Category.{v₁, u₁} C] {V : Type u₂} [Category.{v₂, u₂} V] [MonoidalCategory C] [MonoidalCategory V] (F : Functor C V) :
                    Type (max (max (max u₁ u₂) v₁) v₂)

                    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.

                    Instances
                      @[reducible, inline]

                      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
                      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

                                    The class DayConvolutionMonoidalCategory C V D bundles the necessary data to turn a monoidal category D into a monoidal full subcategory of a category of functors C ⥤ V endowed with a Day convolution monoidal structure. The design of this class is to bundle a fully faithful functor into C ⥤ V with left extensions on its values representing the fact that it maps tensors products in D to Day convolutions, and furthermore ask that this data is "lawful", i.e that once realized in the functor category, the objects behave like the corresponding ones in the category C ⥤ V.

                                    Instances

                                      In a LawfulDayConvolutionMonoidalCategoryStruct, ι.obj (d ⊗ d') is a Day convolution of (ι C V).obj d and (ι C V).d'.

                                      Equations
                                      • One or more equations did not get rendered due to their size.
                                      Instances For

                                        In a LawfulDayConvolutionMonoidalCategoryStruct, ι.obj (d ⊗ d' ⊗ d'') is a (triple) Day convolution of (ι C V).obj d, (ι C V).obj d' and (ι C V).obj d''.

                                        Equations
                                        • One or more equations did not get rendered due to their size.
                                        Instances For

                                          In a LawfulDayConvolutionMonoidalCategoryStruct, ι.obj ((d ⊗ d') ⊗ d'') is a (triple) Day convolution of (ι C V).obj d, (ι C V).obj d' and (ι C V).obj d''.

                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Instances For

                                            In a LawfulDayConvolutionMonoidalCategoryStruct, ι.obj (𝟙_ D) is a Day convolution unit`.

                                            Equations
                                            • One or more equations did not get rendered due to their size.
                                            Instances For

                                              We can promote a LawfulDayConvolutionMonoidalCategoryStruct to a monoidal category, note that every non-prop data is already here, so this is just about showing that they satisfy the axioms of a monoidal category.

                                              Equations
                                              Instances For

                                                In what follows, we give a constructor for LawfulDayConvolutionMonoidalCategoryStruct that does not assume a pre-existing MonoidalCategoryStruct and builds one from the data of suitable convolutions, while giving definitional control over as many parameters as we can.

                                                An InducedLawfulDayConvolutionMonoidalCategoryStructCore C V D bundles the core data needed to construct a full LawfulDayConvolutionMonoidalCategoryStructCore. We’re making this a class so that it can act as a "proxy" for inferring DayConvolution instances (which is all the more important that we are modifying the instances given in the constructor to get better ones defeq-wise). As this object is purely about the internals of definitions of Day convolutions monoidal structures, it is advised to not register this class globally.

                                                Instances

                                                  With the data of chosen isomorphic objects to given day convolutions, and provably equal unit maps through these isomorphisms, we can transform a given family of Day convolutions to one with convolutions definitionally equals to the given objects, and component of units definitionally equal to the provided map family.

                                                  Equations
                                                  • One or more equations did not get rendered due to their size.
                                                  Instances For
                                                    @[reducible, inline]

                                                    Given a fully faithful functor ι : C ⥤ V ⥤ D, a family of Day convolutions, candidate functions for tensorObj and tensorHom, suitable isomorphisms ι.obj (tensorObj d d') ≅ ι.obj (tensorObj d) ⊛ ι.obj (tensorObj d') that behave in a lawful way with respect to the chosen Day convolutions, we can construct a MonoidalCategoryStruct on D.

                                                    Equations
                                                    • One or more equations did not get rendered due to their size.
                                                    Instances For

                                                      Given a fully faithful functor ι : D ⥤ C ⥤ V and mere existence of Day convolutions of ι.obj d and ι.obj d' such that the convolution remains in the essential image of ι, construct an InducedLawfulDayConvolutionMonoidalCategoryStructCore by letting all other data be the generic ones from the HasPointwiseLeftKanExtension API.

                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      Instances For
                                                        noncomputable def CategoryTheory.MonoidalCategory.monoidalOfHasDayConvolutions {C : Type u₁} [Category.{v₁, u₁} C] {V : Type u₂} [Category.{v₂, u₂} V] [MonoidalCategory C] [MonoidalCategory V] {D : Type u₃} [Category.{v₃, u₃} D] (ι : Functor D (Functor C V)) (ffι : ι.FullyFaithful) [hasDayConvolution : ∀ (d d' : D), (tensor C).HasPointwiseLeftKanExtension (externalProduct (ι.obj d) (ι.obj d'))] (essImageDayConvolution : ∀ (d d' : D), ι.essImage ((tensor C).pointwiseLeftKanExtension (externalProduct (ι.obj d) (ι.obj d')))) [hasDayConvolutionUnit : (Functor.fromPUnit (tensorUnit C)).HasPointwiseLeftKanExtension (Functor.fromPUnit (tensorUnit V))] (essImageDayConvolutionUnit : ι.essImage ((Functor.fromPUnit (tensorUnit C)).pointwiseLeftKanExtension (Functor.fromPUnit (tensorUnit V)))) [∀ (v : V) (d : C), Limits.PreservesColimitsOfShape (CostructuredArrow (tensor C) d) (tensorLeft v)] [∀ (v : V) (d : C), Limits.PreservesColimitsOfShape (CostructuredArrow (tensor C) d) (tensorRight v)] [∀ (v : V) (d : C), Limits.PreservesColimitsOfShape (CostructuredArrow (Functor.fromPUnit (tensorUnit C)) d) (tensorLeft v)] [∀ (v : V) (d : C), Limits.PreservesColimitsOfShape (CostructuredArrow (Functor.fromPUnit (tensorUnit C)) d) (tensorRight v)] [∀ (v : V) (d : C × C), Limits.PreservesColimitsOfShape (CostructuredArrow ((Functor.id C).prod (Functor.fromPUnit (tensorUnit C))) d) (tensorRight v)] [∀ (v : V) (d : C × C), Limits.PreservesColimitsOfShape (CostructuredArrow ((tensor C).prod (Functor.id C)) d) (tensorRight v)] :

                                                        Under suitable assumptions on existence of relevant Kan extensions and preservation of relevant colimits by the tensor product of V, we can define a MonoidalCategory D from the data of a fully faithful functor ι : D ⥤ C ⥤ V whose essential image contains a Day convolution unit and is stable under binary Day convolutions.

                                                        Equations
                                                        • One or more equations did not get rendered due to their size.
                                                        Instances For
                                                          noncomputable def CategoryTheory.MonoidalCategory.lawfulDayConvolutionMonoidalCategoryStructOfHasDayConvolutions {C : Type u₁} [Category.{v₁, u₁} C] {V : Type u₂} [Category.{v₂, u₂} V] [MonoidalCategory C] [MonoidalCategory V] {D : Type u₃} [Category.{v₃, u₃} D] (ι : Functor D (Functor C V)) (ffι : ι.FullyFaithful) [hasDayConvolution : ∀ (d d' : D), (tensor C).HasPointwiseLeftKanExtension (externalProduct (ι.obj d) (ι.obj d'))] (essImageDayConvolution : ∀ (d d' : D), ι.essImage ((tensor C).pointwiseLeftKanExtension (externalProduct (ι.obj d) (ι.obj d')))) [hasDayConvolutionUnit : (Functor.fromPUnit (tensorUnit C)).HasPointwiseLeftKanExtension (Functor.fromPUnit (tensorUnit V))] (essImageDayConvolutionUnit : ι.essImage ((Functor.fromPUnit (tensorUnit C)).pointwiseLeftKanExtension (Functor.fromPUnit (tensorUnit V)))) [∀ (v : V) (d : C), Limits.PreservesColimitsOfShape (CostructuredArrow (tensor C) d) (tensorLeft v)] [∀ (v : V) (d : C), Limits.PreservesColimitsOfShape (CostructuredArrow (tensor C) d) (tensorRight v)] [∀ (v : V) (d : C), Limits.PreservesColimitsOfShape (CostructuredArrow (Functor.fromPUnit (tensorUnit C)) d) (tensorLeft v)] [∀ (v : V) (d : C), Limits.PreservesColimitsOfShape (CostructuredArrow (Functor.fromPUnit (tensorUnit C)) d) (tensorRight v)] [∀ (v : V) (d : C × C), Limits.PreservesColimitsOfShape (CostructuredArrow ((Functor.id C).prod (Functor.fromPUnit (tensorUnit C))) d) (tensorRight v)] [∀ (v : V) (d : C × C), Limits.PreservesColimitsOfShape (CostructuredArrow ((tensor C).prod (Functor.id C)) d) (tensorRight v)] :

                                                          The monoidal category constructed via monoidalOfHasDayConvolutions has a canonical LawfulDayConvolutionMonoidalCategoryStruct C V D.

                                                          Equations
                                                          • One or more equations did not get rendered due to their size.
                                                          Instances For