open import Cat.Displayed.Diagram.Total.Terminal
open import Cat.Diagram.Product
open import Cat.Displayed.Total
open import Cat.Displayed.Base
open import Cat.Cartesian
open import Cat.Prelude

import Cat.Displayed.Reasoning
import Cat.Reasoning
module Cat.Displayed.Diagram.Total.Product where
open ∫Hom

Total productsπŸ”—

module _ {ob β„“b oe β„“e} {B : Precategory ob β„“b} (E : Displayed B oe β„“e) where
  open Cat.Displayed.Reasoning E
  open Cat.Reasoning B

  private variable
    a x y p     : Ob
    a' x' y' p' : Ob[ a ]
    f g other   : Hom a x
    f' g'       : Hom[ f ] a' x'

Let be a displayed category, and be a product diagram in A total diagram in is a diagram

satisfying a displayed version of the universal property of products.

  record is-product-over
      {π₁ : Hom p x} {Ο€β‚‚ : Hom p y} (prod : is-product B π₁ Ο€β‚‚)
      (π₁' : Hom[ π₁ ] p' x') (Ο€β‚‚' : Hom[ Ο€β‚‚ ] p' y') : Type (ob βŠ” β„“b βŠ” oe βŠ” β„“e)
    where

    no-eta-equality
    open is-product prod

More explicitly, suppose that we had a triple displayed over as in the following diagram.

is a product, so there exists a unique that commutes with and

This leaves a conspicuous gap in the upstairs portion of the diagram between and is a total product precisely when we have a unique lift of that commutes with and

    field
      ⟨_,_⟩'
        : (f' : Hom[ f ] a' x') (g' : Hom[ g ] a' y')
        β†’ Hom[ ⟨ f , g ⟩ ] a' p'

      Ο€β‚βˆ˜βŸ¨βŸ©' : π₁' ∘' ⟨ f' , g' ⟩' ≑[ Ο€β‚βˆ˜βŸ¨βŸ© ] f'
      Ο€β‚‚βˆ˜βŸ¨βŸ©' : Ο€β‚‚' ∘' ⟨ f' , g' ⟩' ≑[ Ο€β‚‚βˆ˜βŸ¨βŸ© ] g'

      unique'
        : {p1 : π₁ ∘ other ≑ f} {p2 : Ο€β‚‚ ∘ other ≑ g}
        β†’ {other' : Hom[ other ] a' p'}
        β†’ (p1' : (π₁' ∘' other') ≑[ p1 ] f')
        β†’ (p2' : (Ο€β‚‚' ∘' other') ≑[ p2 ] g')
        β†’ other' ≑[ unique p1 p2 ] ⟨ f' , g' ⟩'
    ⟨_,_βŸ©β‚š
      : (f' : Hom[ π₁ ∘ f ] a' x') (g' : Hom[ Ο€β‚‚ ∘ f ] a' y')
      β†’ Hom[ f ] a' p'
    ⟨ f' , g' βŸ©β‚š = hom[ sym (unique refl refl) ] ⟨ f' , g' ⟩'

    abstract
      Ο€β‚βˆ˜βŸ¨βŸ©β‚š : π₁' ∘' ⟨ f' , g' βŸ©β‚š ≑ f'
      Ο€β‚βˆ˜βŸ¨βŸ©β‚š = whisker-r _ βˆ™ reindex _ _ βˆ™ from-pathp[] Ο€β‚βˆ˜βŸ¨βŸ©'

      Ο€β‚‚βˆ˜βŸ¨βŸ©β‚š : Ο€β‚‚' ∘' ⟨ f' , g' βŸ©β‚š ≑ g'
      Ο€β‚‚βˆ˜βŸ¨βŸ©β‚š = whisker-r _ βˆ™ reindex _ _ βˆ™ from-pathp[] Ο€β‚‚βˆ˜βŸ¨βŸ©'

A total product of and in consists of a choice of a total product diagram.

  record
    ProductP {x y} (prod : Product B x y) (x' : Ob[ x ]) (y' : Ob[ y ])
      : Type (ob βŠ” β„“b βŠ” oe βŠ” β„“e) where

    no-eta-equality
    open Product prod

    field
      apex' : Ob[ apex ]
      π₁'   : Hom[ π₁ ] apex' x'
      Ο€β‚‚'   : Hom[ Ο€β‚‚ ] apex' y'
      has-is-product'
        : is-product-over has-is-product π₁' Ο€β‚‚'

    open is-product-over has-is-product' public
has-products-over
  : βˆ€ {o β„“ o' β„“'} {B : Precategory o β„“}
  β†’ Displayed B o' β„“'
  β†’ has-products B
  β†’ Type _
has-products-over {B = B} E prod = βˆ€ {a b : ⌞ B ⌟} (x : E Κ» a) (y : E Κ» b) β†’ ProductP E (prod a b) x y

Total products and total categoriesπŸ”—

module _ {ob β„“b oe β„“e} {B : Precategory ob β„“b} {E : Displayed B oe β„“e} where
  open Cat.Reasoning B
  open Displayed E

  private module ∫E = Cat.Reasoning (∫ E)

As the name suggests, a total product diagram in induces to a product diagram in the total category of

  is-total-product→total-is-product
    : βˆ€ {x y p} {x' : Ob[ x ]} {y' : Ob[ y ]} {p' : Ob[ p ]}
    β†’ {π₁ : ∫E.Hom (p , p') (x , x')} {Ο€β‚‚ : ∫E.Hom (p , p') (y , y')}
    β†’ {prod : is-product B (π₁ .fst) (Ο€β‚‚ .fst)}
    β†’ is-product-over E prod (π₁ .snd) (Ο€β‚‚ .snd)
    β†’ is-product (∫ E) π₁ Ο€β‚‚
The proof is largely shuffling data about, so we elide the details.
  is-total-productβ†’total-is-product {π₁ = π₁} {Ο€β‚‚ = Ο€β‚‚} {prod = prod} total-prod = ∫prod where
    open is-product-over total-prod
    open is-product prod

    ∫prod : is-product (∫ E) π₁ Ο€β‚‚
    ∫prod .is-product.⟨_,_⟩ f g =
      ∫hom ⟨ f .fst , g .fst ⟩ ⟨ f .snd , g .snd ⟩'
    ∫prod .is-product.Ο€β‚βˆ˜βŸ¨βŸ© =
      ∫Hom-path E Ο€β‚βˆ˜βŸ¨βŸ© Ο€β‚βˆ˜βŸ¨βŸ©'
    ∫prod .is-product.Ο€β‚‚βˆ˜βŸ¨βŸ© =
      ∫Hom-path E Ο€β‚‚βˆ˜βŸ¨βŸ© Ο€β‚‚βˆ˜βŸ¨βŸ©'
    ∫prod .is-product.unique p1 p2 =
      ∫Hom-path E
        (unique (ap fst p1) (ap fst p2))
        (unique' (ap snd p1) (ap snd p2))
\ Warning

Note that a product diagram in a total category does not necessarily yield a product diagram in the base category. For a counterexample, consider the following displayed category:

The total category is equivalent to the terminal category, and thus has products. However, the base category does not have products, as the uniqueness condition fails!

module
  Binary-products'
    {o β„“ o' β„“'} {B : Precategory o β„“} {E : Displayed B o' β„“'} {fp : has-products B}
    (fp' : has-products-over E fp)
  where

  open Precategory B

  module _ {a b : Ob} (a' : E Κ» a) (b' : E Κ» b) where open ProductP (fp' a' b') renaming (apex' to _βŠ—β‚€'_) using () public
  module _ {a b : Ob} {a' : E ʻ a} {b' : E ʻ b} where open ProductP (fp' a' b') renaming (unique' to ⟨⟩'-unique) hiding (apex') public

record
  Cartesian-over
    {o β„“ o' β„“'} {B : Precategory o β„“} (E : Displayed B o' β„“') (cart : Cartesian-category B)
    : Type (o βŠ” β„“ βŠ” o' βŠ” β„“')
  where

  private module cart = Cartesian-category cart

  field
    terminal' : TerminalP E cart.terminal
    products' : has-products-over E cart.products

  open TerminalP terminal' hiding (has⊀') public
  open Binary-products' products' public