Documentation

Mathlib.CategoryTheory.Monoidal.DayConvolution.DayFunctor

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 #

structure CategoryTheory.MonoidalCategory.DayFunctor (C : Type u₁) [Category.{v₁, u₁} C] (V : Type u₂) [Category.{v₂, u₂} V] [MonoidalCategory C] [MonoidalCategory V] :
Type (max (max (max u₁ u₂) v₁) v₂)

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.

      Instances For

        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

          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
            Instances For
              @[simp]

              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
                @[simp]
                theorem CategoryTheory.MonoidalCategory.DayFunctor.η_comp_isoPointwiseLeftKanExtension_hom {C : Type u₁} [Category.{v₁, u₁} C] {V : Type u₂} [Category.{v₂, u₂} V] [MonoidalCategory C] [MonoidalCategory V] [hasDayConvolution : ∀ (F G : Functor C V), (tensor C).HasPointwiseLeftKanExtension (externalProduct F G)] [hasDayConvolutionUnit : (Functor.fromPUnit (tensorUnit C)).HasPointwiseLeftKanExtension (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)] (F G : DayFunctor C V) (x y : C) :
                @[simp]
                theorem CategoryTheory.MonoidalCategory.DayFunctor.ι_comp_isoPointwiseLeftKanExtension_inv {C : Type u₁} [Category.{v₁, u₁} C] {V : Type u₂} [Category.{v₂, u₂} V] [MonoidalCategory C] [MonoidalCategory V] [hasDayConvolution : ∀ (F G : Functor C V), (tensor C).HasPointwiseLeftKanExtension (externalProduct F G)] [hasDayConvolutionUnit : (Functor.fromPUnit (tensorUnit C)).HasPointwiseLeftKanExtension (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)] (F G : DayFunctor C V) (x y : C) :

                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
                  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.
                    Instances For