module Cat.Diagram.Product wheremodule _ {o β} (C : Precategory o β) where
  open Cat.Reasoning CProductsπ
The product of two objects and if it exists, is the smallest object equipped with βprojectionβ maps and This situation can be visualised by putting the data of a product into a commutative diagram, as the one below: To express that is the smallest object with projections to and we ask that any other object with projections through and factors uniquely through
In the sense that (univalent) categories generalise posets, the product of and β if it exists β generalises the binary meet Since products are unique when they exist, we may safely denote any product of and by
For a diagram to be a product diagram, it must be able to cough up an arrow given the data of another span which must not only fit into the diagram above but be unique among the arrows that do so.
This factoring is called the pairing of the arrows and since in the special case where is the terminal object (hence the two arrows are global elements of resp. the pairing is a global element of the product
  record is-product {A B P} (Οβ : Hom P A) (Οβ : Hom P B) : Type (o β β) where
    field
      β¨_,_β© : β {Q} (p1 : Hom Q A) (p2 : Hom Q B) β Hom Q P
      Οβββ¨β© : β {Q} {p1 : Hom Q _} {p2} β Οβ β β¨ p1 , p2 β© β‘ p1
      Οβββ¨β© : β {Q} {p1 : Hom Q _} {p2} β Οβ β β¨ p1 , p2 β© β‘ p2
      unique : β {Q} {p1 : Hom Q A} {p2}
             β {other : Hom Q P}
             β Οβ β other β‘ p1
             β Οβ β other β‘ p2
             β other β‘ β¨ p1 , p2 β©
    uniqueβ : β {Q} {pr1 : Hom Q A} {pr2}
            β β {o1} (p1 : Οβ β o1 β‘ pr1) (q1 : Οβ β o1 β‘ pr2)
            β β {o2} (p2 : Οβ β o2 β‘ pr1) (q2 : Οβ β o2 β‘ pr2)
            β o1 β‘ o2
    uniqueβ p1 q1 p2 q2 = unique p1 q1 β sym (unique p2 q2)
    β¨β©β : β {Q R} {p1 : Hom Q A} {p2 : Hom Q B} (f : Hom R Q)
        β β¨ p1 , p2 β© β f β‘ β¨ p1 β f , p2 β f β©
    β¨β©β f = unique (pulll Οβββ¨β©) (pulll Οβββ¨β©)
    β¨β©-Ξ· : β¨ Οβ , Οβ β© β‘ id
    β¨β©-Ξ· = sym $ unique (idr _) (idr _)A product of and is an explicit choice of product diagram:
  record Product (A B : Ob) : Type (o β β) where
    no-eta-equality
    field
      apex : Ob
      Οβ : Hom apex A
      Οβ : Hom apex B
      has-is-product : is-product Οβ Οβ
    open is-product has-is-product publicmodule _ {o β} {C : Precategory o β} where
  open Cat.Reasoning C
  open Product hiding (β¨_,_β© ; Οβ ; Οβ ; β¨β©β)
  private variable
    A B a b c d : Ob
  is-product-is-prop : β {X Y P} {pβ : Hom P X} {pβ : Hom P Y} β is-prop (is-product C pβ pβ)
  is-product-is-prop {X = X} {Y = Y} {pβ = pβ} {pβ} x y = q where
    open is-product
    p : Path (β {P'} β Hom P' X β Hom P' Y β _) (x .β¨_,_β©) (y .β¨_,_β©)
    p i p1 p2 = y .unique {p1 = p1} {p2} (x .Οβββ¨β©) (x .Οβββ¨β©) i
    q : x β‘ y
    q i .β¨_,_β© = p i
    q i .Οβββ¨β© {p1 = p1} {p2} = is-propβpathp (Ξ» i β Hom-set _ _ (pβ β p i p1 p2) p1) (x .Οβββ¨β©) (y .Οβββ¨β©) i
    q i .Οβββ¨β© {p1 = p1} {p2} = is-propβpathp (Ξ» i β Hom-set _ _ (pβ β p i p1 p2) p2) (x .Οβββ¨β©) (y .Οβββ¨β©) i
    q i .unique {p1 = p1} {p2} {other} cβ cβ = is-propβpathp (Ξ» i β Hom-set _ _ other (p i p1 p2)) (x .unique cβ cβ) (y .unique cβ cβ) i
  instance
    H-Level-is-product : β {X Y P} {pβ : Hom P X} {pβ : Hom P Y} {n} β H-Level (is-product C pβ pβ) (suc n)
    H-Level-is-product = prop-instance is-product-is-prop
unquoteDecl Product-path = declare-record-path Product-path (quote Product)
module _ {o β} {C : Precategory o β} where
  open Cat.Reasoning C
  open Product hiding (β¨_,_β© ; Οβ ; Οβ ; β¨β©β)
  private variable
    A B a b c d : ObUniquenessπ
Products, when they exist, are unique. Itβs easiest to see this with a diagrammatic argument: If we have product diagrams and we can fit them into a βcommutative diamondβ like the one below:
Since both and are products, we know that the dashed arrows in the diagram below exist, so the overall diagram commutes: hence we have an isomorphism
We construct the map as the pairing of the projections from and symmetrically for
  Γ-Unique : (p1 p2 : Product C A B) β apex p1 β
 apex p2
  Γ-Unique p1 p2 = make-iso p1βp2 p2βp1 p1βp2βp1 p2βp1βp2 where
    module p1 = Product p1
    module p2 = Product p2
    p1βp2 : Hom (apex p1) (apex p2)
    p1βp2 = p2.β¨ p1.Οβ , p1.Οβ β©
    p2βp1 : Hom (apex p2) (apex p1)
    p2βp1 = p1.β¨ p2.Οβ , p2.Οβ β©These are unique because they are maps into products which commute with the projections.
    p1βp2βp1 : p1βp2 β p2βp1 β‘ id
    p1βp2βp1 = p2.uniqueβ
      (assoc _ _ _ ββ ap (_β _) p2.Οβββ¨β© ββ p1.Οβββ¨β©)
      (assoc _ _ _ ββ ap (_β _) p2.Οβββ¨β© ββ p1.Οβββ¨β©)
      (idr _) (idr _)
    p2βp1βp2 : p2βp1 β p1βp2 β‘ id
    p2βp1βp2 = p1.uniqueβ
      (assoc _ _ _ ββ ap (_β _) p1.Οβββ¨β© ββ p2.Οβββ¨β©)
      (assoc _ _ _ ββ ap (_β _) p1.Οβββ¨β© ββ p2.Οβββ¨β©)
      (idr _) (idr _)
  is-product-iso
    : β {A A' B B' P} {Οβ : Hom P A} {Οβ : Hom P B}
        {f : Hom A A'} {g : Hom B B'}
    β is-invertible f
    β is-invertible g
    β is-product C Οβ Οβ
    β is-product C (f β Οβ) (g β Οβ)
  is-product-iso f-iso g-iso prod = prod' where
    module fi = is-invertible f-iso
    module gi = is-invertible g-iso
    open is-product
    prod' : is-product _ _ _
    prod' .β¨_,_β© qa qb = prod .β¨_,_β© (fi.inv β qa) (gi.inv β qb)
    prod' .Οβββ¨β© = pullr (prod .Οβββ¨β©) β cancell fi.invl
    prod' .Οβββ¨β© = pullr (prod .Οβββ¨β©) β cancell gi.invl
    prod' .unique p q = prod .unique
      (sym (ap (_ β_) (sym p) β pulll (cancell fi.invr)))
      (sym (ap (_ β_) (sym q) β pulll (cancell gi.invr)))
  is-product-iso-apex
    : β {A B P P'} {Οβ : Hom P A} {Οβ : Hom P B}
        {Οβ' : Hom P' A} {Οβ' : Hom P' B}
        {f : Hom P' P}
    β is-invertible f
    β Οβ β f β‘ Οβ'
    β Οβ β f β‘ Οβ'
    β is-product C Οβ Οβ
    β is-product C Οβ' Οβ'
  is-product-iso-apex {f = f} f-iso f-Οβ f-Οβ prod = prod' where
    module fi = is-invertible f-iso
    open is-product
    prod' : is-product _ _ _
    prod' .β¨_,_β© qa qb = fi.inv β prod .β¨_,_β© qa qb
    prod' .Οβββ¨β© = pulll (rswizzle (sym f-Οβ) fi.invl) β prod .Οβββ¨β©
    prod' .Οβββ¨β© = pulll (rswizzle (sym f-Οβ) fi.invl) β prod .Οβββ¨β©
    prod' .unique p q = sym $ lswizzle
      (sym (prod .unique (pulll f-Οβ β p) (pulll f-Οβ β q))) fi.invr  Product-is-prop
    : β {A B}
    β is-category C
    β is-prop (Product C A B)
  Product-is-prop cat p1 p2 = Product-path
    (cat .to-path (Γ-Unique p1 p2))
    (Univalent.Hom-pathp-refll-iso cat (p1 .Οβββ¨β©))
    (Univalent.Hom-pathp-refll-iso cat (p1 .Οβββ¨β©))Categories with all binary productsπ
Categories with all binary products are quite common, so we define a module for working with them.
has-products : β {o β} β Precategory o β β Type _
has-products C = β a b β Product C a b
module Binary-products
  {o β} (C : Precategory o β) (all-products : has-products C) where  open Cat.Reasoning C
  private variable
    A B a b c d : Ob
  -- Note: here and below we have to open public the aliases in a module
  -- with parameters so Agda picks up the display forms.
  module _ {a b} where
    open Product (all-products a b)
      renaming (unique to β¨β©-unique; uniqueβ to β¨β©-uniqueβ)
      hiding (apex)
      public
  open Functor
  infix 50 _ββ_We start by defining a βglobalβ product-assigning operation.
  module _ a b where
    open Product (all-products a b)
      renaming (apex to infixr 7 _ββ_)
      using () publicThis operation extends to a bifunctor
  _ββ_ : β {a b x y} β Hom a x β Hom b y β Hom (a ββ b) (x ββ y)
  f ββ g = β¨ f β Οβ , g β Οβ β©
  Γ-functor : Functor (C ΓαΆ C) C
  Γ-functor .Fβ (a , b) = a ββ b
  Γ-functor .Fβ (f , g) = f ββ g
  Γ-functor .F-id = sym $ β¨β©-unique id-comm id-comm
  Γ-functor .F-β (f , g) (h , i) =
    sym $ β¨β©-unique
      (pulll Οβββ¨β© β extendr Οβββ¨β©)
      (pulll Οβββ¨β© β extendr Οβββ¨β©)We also define a handful of common morphisms.
  Ξ΄ : Hom a (a ββ a)
  Ξ΄ = β¨ id , id β©
  swap : Hom (a ββ b) (b ββ a)
  swap = β¨ Οβ , Οβ β©
  Γ-assoc : Hom (a ββ (b ββ c)) ((a ββ b) ββ c)
  Γ-assoc = β¨ β¨ Οβ , Οβ β Οβ β© , Οβ β Οβ β©  Ξ΄-natural : is-natural-transformation Id (Γ-functor Fβ Catβ¨ Id , Id β©) Ξ» _ β Ξ΄
  Ξ΄-natural x y f = β¨β©-uniqueβ
    (cancell Οβββ¨β©) (cancell Οβββ¨β©)
    (pulll Οβββ¨β© β cancelr Οβββ¨β©) (pulll Οβββ¨β© β cancelr Οβββ¨β©)
  swap-is-iso : β {a b} β is-invertible (swap {a} {b})
  swap-is-iso = make-invertible swap
    (β¨β©-uniqueβ (pulll Οβββ¨β© β Οβββ¨β©) ((pulll Οβββ¨β© β Οβββ¨β©)) (idr _) (idr _))
    (β¨β©-uniqueβ (pulll Οβββ¨β© β Οβββ¨β©) ((pulll Οβββ¨β© β Οβββ¨β©)) (idr _) (idr _))
  swap-natural
    : β {A B C D} ((f , g) : Hom A C Γ Hom B D)
    β (g ββ f) β swap β‘ swap β (f ββ g)
  swap-natural (f , g) =
    (g ββ f) β swap                       β‘β¨ β¨β©β _ β©
    β¨ (g β Οβ) β swap , (f β Οβ) β swap β© β‘β¨ apβ β¨_,_β© (pullr Οβββ¨β©) (pullr Οβββ¨β©) β©
    β¨ g β Οβ , f β Οβ β©                   β‘Λβ¨ apβ β¨_,_β© Οβββ¨β© Οβββ¨β© β©
    β¨ Οβ β (f ββ g) , Οβ β (f ββ g) β©     β‘Λβ¨ β¨β©β _ β©
    swap β (f ββ g)                       β
  swap-Ξ΄ : β {A} β swap β Ξ΄ β‘ Ξ΄ {A}
  swap-Ξ΄ = β¨β©-unique (pulll Οβββ¨β© β Οβββ¨β©) (pulll Οβββ¨β© β Οβββ¨β©)
  assoc-Ξ΄ : β {a} β Γ-assoc β (id ββ Ξ΄ {a}) β Ξ΄ {a} β‘ (Ξ΄ ββ id) β Ξ΄
  assoc-Ξ΄ = β¨β©-uniqueβ
    (pulll Οβββ¨β© β β¨β©-uniqueβ
      (pulll Οβββ¨β© β pulll Οβββ¨β© β pullr Οβββ¨β©)
      (pulll Οβββ¨β© β pullr (pulll Οβββ¨β©) β pulll (pulll Οβββ¨β©) β pullr Οβββ¨β©)
      (pulll (pulll Οβββ¨β©) β pullr Οβββ¨β©)
      (pulll (pulll Οβββ¨β©) β pullr Οβββ¨β©)
    β pushl (sym Οβββ¨β©))
    (pulll Οβββ¨β© β pullr (pulll Οβββ¨β©) β pulll (pulll Οβββ¨β©) β pullr Οβββ¨β©)
    refl
    (pulll Οβββ¨β© β pullr Οβββ¨β©)
  by-Οβ : β {f f' : Hom a b} {g g' : Hom a c} β β¨ f , g β© β‘ β¨ f' , g' β© β f β‘ f'
  by-Οβ p = sym Οβββ¨β© β ap (Οβ β_) p β Οβββ¨β©
  extend-Οβ : β {f : Hom a b} {g : Hom a c} {h} β β¨ f , g β© β‘ h β f β‘ Οβ β h
  extend-Οβ p = sym Οβββ¨β© β ap (Οβ β_) p
  by-Οβ : β {f f' : Hom a b} {g g' : Hom a c} β β¨ f , g β© β‘ β¨ f' , g' β© β g β‘ g'
  by-Οβ p = sym Οβββ¨β© β ap (Οβ β_) p β Οβββ¨β©
  extend-Οβ : β {f : Hom a b} {g : Hom a c} {h} β β¨ f , g β© β‘ h β g β‘ Οβ β h
  extend-Οβ p = sym Οβββ¨β© β ap (Οβ β_) p
  Οβ-inv
    : β {f : Hom (a ββ b) c} {g : Hom (a ββ b) d}
    β (β¨β©-inv : is-invertible β¨ f , g β©)
    β f β is-invertible.inv β¨β©-inv β‘ Οβ
  Οβ-inv {f = f} {g = g} β¨β©-inv =
    pushl (sym Οβββ¨β©) β elimr (is-invertible.invl β¨β©-inv)
  Οβ-inv
    : β {f : Hom (a ββ b) c} {g : Hom (a ββ b) d}
    β (β¨β©-inv : is-invertible β¨ f , g β©)
    β g β is-invertible.inv β¨β©-inv β‘ Οβ
  Οβ-inv {f = f} {g = g} β¨β©-inv =
    pushl (sym Οβββ¨β©) β elimr (is-invertible.invl β¨β©-inv)Representability of productsπ
module _ {o β} {C : Precategory o β} where
  open Cat.Reasoning CThe collection of maps into a product is equivalent to the collection of pairs of maps into and The forward direction of the equivalence is given by postcomposition of the projections, and the reverse direction by the universal property of products.
  product-repr
    : β {a b}
    β (prod : Product C a b)
    β (x : Ob)
    β Hom x (Product.apex prod) β (Hom x a Γ Hom x b)
  product-repr prod x = IsoβEquiv Ξ» where
      .fst f β Οβ β f , Οβ β f
      .snd .is-iso.from (f , g) β β¨ f , g β©
      .snd .is-iso.rinv (f , g) β Οβββ¨β© ,β Οβββ¨β©
      .snd .is-iso.linv f β sym (β¨β©β f) β eliml β¨β©-Ξ·
    where open Product prod