Documentation

Mathlib.CategoryTheory.Monoidal.OfChosenFiniteProducts.Basic

The monoidal structure on a category with chosen finite products. #

This is a variant of the development in CategoryTheory.Monoidal.OfHasFiniteProducts, which uses specified choices of the terminal object and binary product, enabling the construction of a cartesian category with specific definitions of the tensor unit and tensor product.

(Because the construction in CategoryTheory.Monoidal.OfHasFiniteProducts uses HasLimit classes, the actual definitions there are opaque behind Classical.choice.)

We use this in CategoryTheory.Monoidal.TypeCat to construct the monoidal category of types so that the tensor product is the usual cartesian product of types.

For now we only do the construction from products, and not from coproducts, which seems less often useful.

Swap the two sides of a BinaryFan.

Equations
Instances For
    @[simp]
    @[simp]

    If a binary fan t over P Q is a limit cone, then t.swap is a limit cone over Q P.

    Equations
    Instances For

      Construct HasBinaryProduct Q P from HasBinaryProduct P Q. This can't be an instance, as it would cause a loop in typeclass search.

      Given a limit cone over X and Y, and another limit cone over Y and X, we can construct an isomorphism between the cone points. Relative to some fixed choice of limits cones for every pair, these isomorphisms constitute a braiding.

      Equations
      Instances For

        Given binary fans sXY over X Y, and sYZ over Y Z, and s over sXY.X Z, if sYZ is a limit cone we can construct a binary fan over X sYZ.X.

        This is an ingredient of building the associator for a cartesian category.

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

          Given binary fans sXY over X Y, and sYZ over Y Z, and s over X sYZ.X, if sYZ is a limit cone we can construct a binary fan over sXY.X Z.

          This is an ingredient of building the associator for a cartesian category.

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

            If all the binary fans involved a limit cones, BinaryFan.assoc produces another limit cone.

            Equations
            Instances For
              @[reducible, inline]

              Given two pairs of limit cones corresponding to the parenthesisations of X × Y × Z, we obtain an isomorphism between the cone points.

              Equations
              Instances For
                @[reducible, inline]
                abbrev CategoryTheory.Limits.BinaryFan.associatorOfLimitCone {C : Type u} [CategoryTheory.Category.{v, u} C] (L : (X Y : C) → CategoryTheory.Limits.LimitCone (CategoryTheory.Limits.pair X Y)) (X : C) (Y : C) (Z : C) :
                (L (L X Y).cone.pt Z).cone.pt (L X (L Y Z).cone.pt).cone.pt

                Given a fixed family of limit data for every pair X Y, we obtain an associator.

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

                  Construct a left unitor from specified limit cones.

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

                    Construct a right unitor from specified limit cones.

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

                      Implementation of the tensor product for MonoidalOfChosenFiniteProducts.

                      Equations
                      Instances For
                        @[reducible, inline]

                        Implementation of the tensor product of morphisms for MonoidalOfChosenFiniteProducts.

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

                          A category with a terminal object and binary products has a natural monoidal structure.

                          Equations
                          Instances For

                            A type synonym for C carrying a monoidal category structure corresponding to a fixed choice of limit data for the empty functor, and for pair X Y for every X Y : C.

                            This is an implementation detail for SymmetricOfChosenFiniteProducts.

                            Equations
                            Instances For