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

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