Documentation

Mathlib.CategoryTheory.Monoidal.FunctorCategory

Monoidal structure on C ⥤ D when D is monoidal. #

When C is any category, and D is a monoidal category, there is a natural "pointwise" monoidal structure on C ⥤ D.

The initial intended application is tensor product of presheaves.

(An auxiliary definition for functorCategoryMonoidal.) Tensor product of functors C ⥤ D, when D is monoidal.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    @[simp]
    theorem CategoryTheory.Monoidal.FunctorCategory.tensorObj_map {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] [MonoidalCategory D] (F G : Functor C D) {X✝ Y✝ : C} (f : X✝ Y✝) :
    def CategoryTheory.Monoidal.FunctorCategory.tensorHom {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] [MonoidalCategory D] {F G F' G' : Functor C D} (α : F G) (β : F' G') :

    (An auxiliary definition for functorCategoryMonoidal.) Tensor product of natural transformations into D, when D is monoidal.

    Equations
    Instances For
      @[simp]
      theorem CategoryTheory.Monoidal.FunctorCategory.tensorHom_app {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] [MonoidalCategory D] {F G F' G' : Functor C D} (α : F G) (β : F' G') (X : C) :

      (An auxiliary definition for functorCategoryMonoidal.)

      Equations
      Instances For

        (An auxiliary definition for functorCategoryMonoidal.)

        Equations
        Instances For

          When C is any category, and D is a monoidal category, the functor category C ⥤ D has a natural pointwise monoidal structure, where (F ⊗ G).obj X = F.obj X ⊗ G.obj X.

          Equations
          • One or more equations did not get rendered due to their size.
          @[simp]
          theorem CategoryTheory.Monoidal.tensorUnit_obj {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] [MonoidalCategory D] {X : C} :
          (𝟙_ (Functor C D)).obj X = 𝟙_ D
          @[simp]
          theorem CategoryTheory.Monoidal.tensorUnit_map {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] [MonoidalCategory D] {X Y : C} {f : X Y} :
          (𝟙_ (Functor C D)).map f = CategoryStruct.id (𝟙_ D)
          @[simp]
          theorem CategoryTheory.Monoidal.tensorHom_app {C : Type u₁} [Category.{v₁, u₁} C] {D : Type u₂} [Category.{v₂, u₂} D] [MonoidalCategory D] {F G F' G' : Functor C D} {α : F G} {β : F' G'} {X : C} :

          When C is any category, and D is a monoidal category, the functor category C ⥤ D has a natural pointwise monoidal structure, where (F ⊗ G).obj X = F.obj X ⊗ G.obj X.

          Equations

          When C is any category, and D is a braided monoidal category, the natural pointwise monoidal structure on the functor category C ⥤ D is also braided.

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

          When C is any category, and D is a symmetric monoidal category, the natural pointwise monoidal structure on the functor category C ⥤ D is also symmetric.

          Equations