Documentation

Mathlib.CategoryTheory.Monoidal.Free.Basic

The free monoidal category over a type #

Given a type C, the free monoidal category over C has as objects formal expressions built from (formal) tensor products of terms of C and a formal unit. Its morphisms are compositions and tensor products of identities, unitors and associators.

In this file, we construct the free monoidal category and prove that it is a monoidal category. If D is a monoidal category, we construct the functor FreeMonoidalCategory C ⥤ D associated to a function C → D.

The free monoidal category has two important properties: it is a groupoid and it is thin. The former is obvious from the construction, and the latter is what is commonly known as the monoidal coherence theorem. Both of these properties are proved in the file Coherence.lean.

Given a type C, the free monoidal category over C has as objects formal expressions built from (formal) tensor products of terms of C and a formal unit. Its morphisms are compositions and tensor products of identities, unitors and associators.

Instances For
    Equations
    • CategoryTheory.instInhabitedFreeMonoidalCategory = { default := CategoryTheory.FreeMonoidalCategory.unit }

    Formal compositions and tensor products of identities, unitors and associators. The morphisms of the free monoidal category are obtained as a quotient of these formal morphisms by the relations defining a monoidal category.

    Instances For

      The morphisms of the free monoidal category satisfy 21 relations ensuring that the resulting category is in fact a category and that it is monoidal.

      Instances For

        We say that two formal morphisms in the free monoidal category are equivalent if they become equal if we apply the relations that are true in a monoidal category. Note that we will prove that there is only one equivalence class -- this is the monoidal coherence theorem.

        Equations
        Equations
        Equations
        • CategoryTheory.FreeMonoidalCategory.instMonoidalCategoryFreeMonoidalCategoryCategoryFreeMonoidalCategory = CategoryTheory.MonoidalCategory.mk
        @[simp]
        theorem CategoryTheory.FreeMonoidalCategory.unit_eq_unit {C : Type u} :
        CategoryTheory.FreeMonoidalCategory.unit = 𝟙_ (CategoryTheory.FreeMonoidalCategory C)
        theorem CategoryTheory.FreeMonoidalCategory.Hom.inductionOn {C : Type u} {motive : {X Y : CategoryTheory.FreeMonoidalCategory C} → (X Y)Prop} {X : CategoryTheory.FreeMonoidalCategory C} {Y : CategoryTheory.FreeMonoidalCategory C} (t : X Y) (id : ∀ (X : CategoryTheory.FreeMonoidalCategory C), motive (CategoryTheory.CategoryStruct.id X)) (α_hom : ∀ (X Y Z : CategoryTheory.FreeMonoidalCategory C), motive (CategoryTheory.MonoidalCategory.associator X Y Z).hom) (α_inv : ∀ (X Y Z : CategoryTheory.FreeMonoidalCategory C), motive (CategoryTheory.MonoidalCategory.associator X Y Z).inv) (l_hom : ∀ (X : CategoryTheory.FreeMonoidalCategory C), motive (CategoryTheory.MonoidalCategory.leftUnitor X).hom) (l_inv : ∀ (X : CategoryTheory.FreeMonoidalCategory C), motive (CategoryTheory.MonoidalCategory.leftUnitor X).inv) (ρ_hom : ∀ (X : CategoryTheory.FreeMonoidalCategory C), motive (CategoryTheory.MonoidalCategory.rightUnitor X).hom) (ρ_inv : ∀ (X : CategoryTheory.FreeMonoidalCategory C), motive (CategoryTheory.MonoidalCategory.rightUnitor X).inv) (comp : ∀ {X Y Z : CategoryTheory.FreeMonoidalCategory C} (f : X Y) (g : Y Z), motive fmotive gmotive (CategoryTheory.CategoryStruct.comp f g)) (whiskerLeft : ∀ (X : CategoryTheory.FreeMonoidalCategory C) {Y Z : CategoryTheory.FreeMonoidalCategory C} (f : Y Z), motive fmotive (CategoryTheory.MonoidalCategory.whiskerLeft X f)) (whiskerRight : ∀ {X Y : CategoryTheory.FreeMonoidalCategory C} (f : X Y) (Z : CategoryTheory.FreeMonoidalCategory C), motive fmotive (CategoryTheory.MonoidalCategory.whiskerRight f Z)) :
        motive t

        Auxiliary definition for free_monoidal_category.project.

        Equations
        Instances For

          If D is a monoidal category and we have a function C → D, then we have a functor from the free monoidal category over C to the category D.

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