open import Cat.Displayed.Cartesian.Indexing
open import Cat.Displayed.BeckChevalley
open import Cat.Displayed.Comprehension
open import Cat.Displayed.Cocartesian
open import Cat.Displayed.Cartesian
open import Cat.Displayed.Fibre
open import Cat.Displayed.Base
open import Cat.Prelude

import Cat.Displayed.Reasoning
import Cat.Reasoning
module Cat.Displayed.Comprehension.Coproduct
  {ob β„“b od β„“d oe β„“e} {B : Precategory ob β„“b}
  {D : Displayed B od β„“d} {E : Displayed B oe β„“e}
  (D-fib : Cartesian-fibration D) (E-fib : Cartesian-fibration E)
  (P : Comprehension-structure E)
  where
private
  open Cat.Reasoning B

  module E = Displayed E

  module D where
    open Cat.Displayed.Reasoning D public
    open Cartesian-fibration D D-fib public

  module D↓ {Ξ“} = Cat.Reasoning (Fibre D Ξ“)

  open Comprehension E E-fib P

  module D* {Ξ“ Ξ” : Ob} (Οƒ : Hom Ξ“ Ξ”) = Functor (base-change D D-fib Οƒ)
  module Ο€* {Ξ“ : Ob} (x : E.Ob[ Ξ“ ]) = Functor (base-change D D-fib (Ο€αΆœ {x = x}))

  open Functor
  open _=>_

Coproducts in comprehension categoriesπŸ”—

Let be a comprehension category and a fibration over the same base. We say that has when:

  • For every and there exists an object

  • Every projection induces a cocartesian map

  • For every cubical diagram as below, if is cartesian in and are cartesian in and is cocartesian in then is cocartesian in

From a type-theoretic perspective, the first two conditions are rather straightforward. The first condition establishes that we have a type former that can quantify over objects of The second condition defines an introduction rule of the following form:

The elimination rule comes from each being cocartesian, as cocartesian maps satisfy a mapping-out universal property, exactly like a coproduct should!

However, at a glance, the last condition is somewhat mysterious. It says that cocartesian maps over projections are stable: but what does that have to do with coproducts? This, like many other questions in category theory, is elucidated by thinking type-theoretically: A cartesian map in a comprehension category is essentially a substitution: the stability condition says that the introduction rule for coproducts is actually type-theoretic: stable under substitution.

record has-comprehension-coproducts : Type (ob βŠ” β„“b βŠ” od βŠ” β„“d βŠ” oe βŠ” β„“e) where
  no-eta-equality
  field
    ∐ : βˆ€ {Ξ“} β†’ (x : E.Ob[ Ξ“ ]) (a : D.Ob[ Ξ“ β¨Ύ x ]) β†’ D.Ob[ Ξ“ ]
    ⟨_,_⟩
      : βˆ€ {Ξ“} β†’ (x : E.Ob[ Ξ“ ]) (a : D.Ob[ Ξ“ β¨Ύ x ])
      β†’ D.Hom[ Ο€αΆœ ] a (∐ x a)
    ⟨⟩-cocartesian
      : βˆ€ {Ξ“} β†’ (x : E.Ob[ Ξ“ ]) (a : D.Ob[ Ξ“ β¨Ύ x ])
      β†’ is-cocartesian D Ο€αΆœ ⟨ x , a ⟩
    ∐-beck-chevalley
      : βˆ€ {Ξ“ Ξ” x y} {Οƒ : Hom Ξ“ Ξ”} {f : E.Hom[ Οƒ ] x y}
      β†’ is-cartesian E Οƒ f
      β†’ left-beck-chevalley D Ο€αΆœ (Οƒ β¨ΎΛ’ f) Οƒ Ο€αΆœ (sub-proj f)
  module ⟨⟩-cocartesian {Ξ“} (x : E.Ob[ Ξ“ ]) (a : D.Ob[ Ξ“ β¨Ύ x ]) =
    is-cocartesian (⟨⟩-cocartesian x a)

Now, some general facts about coproducts. To start, note that forming coproducts is a functorial operation. The proof is very routine β€” if you’ve seen one functoriality argument, you’ve seen them all β€” so we’ve hidden it from the page.1

  opaque
    ∐[_]
      : βˆ€ {Ξ“} {x : E.Ob[ Ξ“ ]} {a b : D.Ob[ Ξ“ β¨Ύ x ]}
      β†’ D.Hom[ id ] a b β†’ D.Hom[ id ] (∐ x a) (∐ x b)
    ∐[_] {Ξ“} {x} {a} {b} f =
      ⟨⟩-cocartesian.universal' x a id-comm-sym (⟨ x , b ⟩ D.∘' f)

    ∐[]-id : βˆ€ {Ξ“} {x : E.Ob[ Ξ“ ]} {a : D.Ob[ Ξ“ β¨Ύ x ]} β†’ ∐[ D.id' {x = a} ] ≑ D.id'
    ∐[]-∘
      : βˆ€ {Ξ“} {x : E.Ob[ Ξ“ ]} {a b c : D.Ob[ Ξ“ β¨Ύ x ]}
      β†’ (f : D.Hom[ id ] b c) (g : D.Hom[ id ] a b)
      β†’ ∐[ f D↓.∘ g ] D.≑[ sym (idl _) ] ∐[ f ] D.∘' ∐[ g ]
    ∐[]-natural
      : βˆ€ {Ξ“} {x : E.Ob[ Ξ“ ]} {a b : D.Ob[ Ξ“ β¨Ύ x ]}
      β†’ (f : D.Hom[ id ] a b)
      β†’ ∐[ f ] D.∘' ⟨ x , a ⟩ D.≑[ id-comm-sym ] ⟨ x , b ⟩ D.∘' f
    ∐[]-natural {x = x} {a} {b} f =
      ⟨⟩-cocartesian.commutesp x a _ _

    ∐[]-id {x = x} {a = a} =
      sym $ ⟨⟩-cocartesian.unique _ _ _ $ D.from-pathp[]⁻ $ D.cast[] $
      D.id' D.∘' ⟨ x , a ⟩ D.≑[]⟨ D.idl' _ ⟩
      ⟨ x , a ⟩            D.≑[]⟨ symP (D.idr' _ ) ⟩
      ⟨ x , a ⟩ D.∘' D.id' ∎

    ∐[]-∘ {x = x} {a = a} {b = b} {c = c} f g =
      symP $ ⟨⟩-cocartesian.uniquep x a _ _ _ _ $
      (∐[ f ] D.∘' ∐[ g ]) D.∘' ⟨ x , a ⟩          D.≑[]⟨ D.pullr[] _ (∐[]-natural g) ⟩
      ∐[ f ] D.∘' ⟨ x , b ⟩ D.∘' g                 D.≑[]⟨ D.extendl[] _ (∐[]-natural f) ⟩
      ⟨ x , c ⟩ D.∘' f D.∘' g                      D.≑[]⟨ D.to-pathp[] (D.unwhisker-r (ap (Ο€αΆœ ∘_) (idl _)) (idl _)) ⟩
      ⟨ x , c ⟩ D.∘' (f D↓.∘ g) ∎

We can also re-package the cocartesianness of to give a method for constructing morphisms given morphisms Type theoretically, this gives a much more familiar presentation of the elimination rule.

  opaque
    ∐-elim
      : βˆ€ {Ξ“} {x : E.Ob[ Ξ“ ]} {a : D.Ob[ Ξ“ β¨Ύ x ]} {b : D.Ob[ Ξ“ ]}
      β†’ D.Hom[ id ] a (Ο€*.Fβ‚€ x b)
      β†’ D.Hom[ id ] (∐ x a) b
    ∐-elim {x = x} {a = a} {b = b} f =
      ⟨⟩-cocartesian.universal' x a id-comm-sym (D.Ο€* Ο€αΆœ b D.∘' f)

    ∐-elim-β
      : βˆ€ {Ξ“} {x : E.Ob[ Ξ“ ]} {a : D.Ob[ Ξ“ β¨Ύ x ]} {b : D.Ob[ Ξ“ ]}
      β†’ (f : D.Hom[ id ] a (Ο€*.Fβ‚€ x b))
      β†’ ∐-elim f D.∘' ⟨ x , a ⟩ D.≑[ id-comm-sym ] D.Ο€* Ο€αΆœ b D.∘' f
    ∐-elim-β f = ⟨⟩-cocartesian.commutesp _ _ _ _

Putting on our type-theorist goggles, we’re still missing one thing to call this a proper elimination rule: stability under substitution. In the categorical world, that’s satisfied by the proof that the operation we just constructed is natural, given below.

    ∐-elim-natural
      : βˆ€ {Ξ“} {x : E.Ob[ Ξ“ ]} {a b : D.Ob[ Ξ“ β¨Ύ x ]} {c : D.Ob[ Ξ“ ]}
      β†’ (f : D.Hom[ id ] b (Ο€*.β‚€ x c)) (g : D.Hom[ id ] a b)
      β†’ ∐-elim f D.∘' ∐[ g ] D.≑[ idl _ ] ∐-elim (f D↓.∘ g)
    ∐-elim-natural {x = x} {a = a} {b = b} {c = c} f g =
      ⟨⟩-cocartesian.uniquep x a _ _ _ (∐-elim f D.∘' ∐[ g ]) $
        ((∐-elim f) D.∘' ∐[ g ]) D.∘' ⟨ x , a ⟩ D.≑[]⟨ D.pullr[] _ (∐[]-natural g) ⟩
        ∐-elim f D.∘' ⟨ x , b ⟩ D.∘' g          D.≑[]⟨ D.extendl[] id-comm-sym (⟨⟩-cocartesian.commutesp x b _ _) ⟩
        D.Ο€* Ο€αΆœ c D.∘' f D.∘' g                 D.≑[]⟨ D.to-pathp[] (D.unwhisker-r (ap (Ο€αΆœ ∘_) (idl _)) (idl _)) ⟩
        D.Ο€* Ο€αΆœ c D.∘' D.hom[] (f D.∘' g)       ∎

Conversely, we can make maps given maps This isn’t quite, type-theoretically, as the elimination principle: it’s a weird mash-up of the introduction rule for coproducts, followed by substitution.

  opaque
    ∐-transpose
      : βˆ€ {Ξ“} {x : E.Ob[ Ξ“ ]} {a : D.Ob[ Ξ“ β¨Ύ x ]} {b : D.Ob[ Ξ“ ]}
      β†’ D.Hom[ id ] (∐ x a) b
      β†’ D.Hom[ id ] a (Ο€*.β‚€ x b)
    ∐-transpose {x = x} {a = a} {b = b} f =
      D.Ο€*.universal' id-comm (f D.∘' ⟨ x , a ⟩)
  opaque
    unfolding ∐-transpose
    ∐-transpose-weaken
        : βˆ€ {Ξ“} {x : E.Ob[ Ξ“ ]} {a : D.Ob[ Ξ“ β¨Ύ x ]} {b : D.Ob[ Ξ“ ]}
        β†’ (f : D.Hom[ id ] (∐ x a) b)
        β†’ D.Ο€* Ο€αΆœ b D.∘' ∐-transpose f D.≑[ id-comm ] f D.∘' ⟨ x , a ⟩
    ∐-transpose-weaken f = D.Ο€*.commutesp id-comm _

While ∐-transpose may not play an obvious type-theoretic role, it is extremely important categorically, since it is an inverse of ∐-elim! Moreover, it’s also natural, but that proof is also hidden from the page for brevity.

  opaque
    unfolding ∐-elim ∐-transpose
    ∐-elim-transpose
      : βˆ€ {Ξ“} {x : E.Ob[ Ξ“ ]} {a : D.Ob[ Ξ“ β¨Ύ x ]} {b : D.Ob[ Ξ“ ]}
      β†’ (f : D.Hom[ id ] (∐ x a) b)
      β†’ ∐-elim (∐-transpose f) ≑ f
    ∐-elim-transpose {x = x} {a = a} {b = b} f = sym $
      ⟨⟩-cocartesian.uniquep x a _ _ _ f $ symP $
      D.Ο€*.commutesp id-comm (f D.∘' ⟨ x , a ⟩)

    ∐-transpose-elim
      : βˆ€ {Ξ“} {x : E.Ob[ Ξ“ ]} {a : D.Ob[ Ξ“ β¨Ύ x ]} {b : D.Ob[ Ξ“ ]}
      β†’ (f : D.Hom[ id ] a (Ο€*.Fβ‚€ x b))
      β†’ ∐-transpose (∐-elim f) ≑ f
    ∐-transpose-elim {x = x} {a = a} {b = b} f = sym $
      D.Ο€*.uniquep _ _ _ f $ symP $
      ⟨⟩-cocartesian.commutesp x a id-comm-sym (D.Ο€* Ο€αΆœ b D.∘' f)
  ∐-transpose-equiv
    : βˆ€ {Ξ“} {x : E.Ob[ Ξ“ ]} {a : D.Ob[ Ξ“ β¨Ύ x ]} {b : D.Ob[ Ξ“ ]}
    β†’ is-equiv (∐-transpose {a = a} {b = b})
  ∐-transpose-equiv = is-isoβ†’is-equiv $
    iso ∐-elim ∐-transpose-elim ∐-elim-transpose

  opaque
    unfolding ∐-transpose
    ∐-transpose-naturall
      : βˆ€ {Ξ“} {x : E.Ob[ Ξ“ ]} {a b : D.Ob[ Ξ“ β¨Ύ x ]} {c : D.Ob[ Ξ“ ]}
      β†’ (f : D.Hom[ id ] (∐ x b) c) (g : D.Hom[ id ] a b)
      β†’ ∐-transpose f D.∘' g D.≑[ idl _ ] ∐-transpose (f D↓.∘ ∐[ g ])
    ∐-transpose-naturall {x = x} {a = a} {b = b} {c = c} f g =
      D.Ο€*.uniquep _ _ _ (∐-transpose f D.∘' g) $
        D.Ο€* Ο€αΆœ c D.∘' ∐-transpose f D.∘' g D.≑[]⟨ D.pulll[] id-comm (D.Ο€*.commutesp _ _) ⟩
        (f D.∘' ⟨ x , b ⟩) D.∘' g           D.≑[]⟨ D.extendr[] _ (symP (∐[]-natural g)) ⟩
        (f D.∘' ∐[ g ]) D.∘' ⟨ x , a ⟩      D.≑[ ap (_∘ Ο€αΆœ) (idl _) ]⟨ D.to-pathp[] (D.unwhisker-l (ap (_∘ Ο€αΆœ) (idl _)) (idl _)) ⟩
        (f D↓.∘ ∐[ g ]) D.∘' ⟨ x , a ⟩      ∎

    ∐-transpose-naturalr
      : βˆ€ {Ξ“} {x : E.Ob[ Ξ“ ]} {a : D.Ob[ Ξ“ β¨Ύ x ]} {b c : D.Ob[ Ξ“ ]}
      β†’ (f : D.Hom[ id ] b c) (g : D.Hom[ id ] (∐ x a) b)
      β†’ Ο€*.₁ x f D.∘' ∐-transpose g D.≑[ idl _ ] ∐-transpose (f D↓.∘ g)
    ∐-transpose-naturalr {x = x} {a = a} {b = b} {c = c} f g =
      D.Ο€*.uniquep _ _ _ (Ο€*.F₁ x f D.∘' ∐-transpose g) $
        D.Ο€* Ο€αΆœ c D.∘' Ο€*.₁ x f D.∘' ∐-transpose g     D.≑[]⟨ D.pulll[] _ (D.Ο€*.commutesp id-comm _) ⟩
        (f D.∘' D.Ο€* Ο€αΆœ b) D.∘' ∐-transpose g          D.≑[]⟨ D.extendr[] id-comm (D.Ο€*.commutesp _ _) ⟩
        (f D.∘' g) D.∘' ⟨ x , a ⟩                      D.≑[ ap (_∘ Ο€αΆœ) (idl _) ]⟨ D.to-pathp[] (D.unwhisker-l (ap (_∘ Ο€αΆœ) (idl _)) (idl _)) ⟩
        D.hom[ idl id ] (f D.∘' g) D.∘' ⟨ x , a ⟩      ∎

Next, we define an introduction rule for coproducts that also lets us apply a mediating substitution:

  opaque
    ⟨_⨾_⟩
      : βˆ€ {Ξ“ Ξ” x y} {Οƒ : Hom Ξ“ Ξ”}
      β†’ (f : E.Hom[ Οƒ ] x y) (a : D.Ob[ Ξ” β¨Ύ y ])
      β†’ D.Hom[ Ο€αΆœ ] (D*.β‚€ (Οƒ β¨ΎΛ’ f) a) (D*.β‚€ Οƒ (∐ y a))

    ⟨_β¨Ύ_⟩ {x = x} {y = y} {Οƒ = Οƒ} f a =
      D.Ο€*.universal' (sym (sub-proj f)) $
      ⟨ y , a ⟩ D.∘' D.Ο€* (Οƒ β¨ΎΛ’ f) a

    ⟨⨾⟩-weaken
      : βˆ€ {Ξ“ Ξ” x y} {Οƒ : Hom Ξ“ Ξ”}
      β†’ (f : E.Hom[ Οƒ ] x y) (a : D.Ob[ Ξ” β¨Ύ y ])
      β†’ D.Ο€* Οƒ (∐ y a) D.∘' ⟨ f β¨Ύ a ⟩
      D.≑[ sym (sub-proj f) ] ⟨ y , a ⟩ D.∘' D.Ο€* (Οƒ β¨ΎΛ’ f) a

    ⟨⨾⟩-weaken {y = y} {Οƒ = Οƒ} f a =
       D.Ο€*.commutesp (symP (sub-proj f)) _

Because we have assumed that cocartesian maps are stable when pulled back along cartesian maps over projections2, this map is also cocartesian β€” and, you guessed it β€” the spiced-up introduction rule is also natural.

  opaque
    ⟨⨾⟩-cocartesian
      : βˆ€ {Ξ“ Ξ” x y} {Οƒ : Hom Ξ“ Ξ”} β†’ {f : E.Hom[ Οƒ ] x y}
      β†’ is-cartesian E Οƒ f β†’ (a : D.Ob[ Ξ” β¨Ύ y ])
      β†’ is-cocartesian D Ο€αΆœ ⟨ f β¨Ύ a ⟩
    ⟨⨾⟩-cocartesian {x = x} {y = y} {Οƒ = Οƒ} {f = f} cart a =
      ∐-beck-chevalley cart
        (symP (⟨⨾⟩-weaken f a))
        (⟨⟩-cocartesian y a)
        D.Ο€*.cartesian
        D.Ο€*.cartesian

  module ⟨⨾⟩-cocartesian
    {Ξ“ Ξ” x y} {Οƒ : Hom Ξ“ Ξ”} {f : E.Hom[ Οƒ ] x y}
    (cart : is-cartesian E Οƒ f) (a : D.Ob[ Ξ” β¨Ύ y ])
    = is-cocartesian (⟨⨾⟩-cocartesian cart a)
  opaque
    unfolding ⟨_⨾_⟩
    ⟨⨾⟩-natural
      : βˆ€ {Ξ“ Ξ” x y} {Οƒ : Hom Ξ“ Ξ”}
      β†’ {a b : D.Ob[ Ξ” β¨Ύ y ]}
      β†’ (f : D.Hom[ id ] a b) (g : E.Hom[ Οƒ ] x y)
      β†’ ⟨ g β¨Ύ b ⟩ D.∘' D*.₁ (Οƒ β¨ΎΛ’ g) f D.≑[ id-comm ] D*.₁ Οƒ ∐[ f ] D.∘' ⟨ g β¨Ύ a ⟩
    ⟨⨾⟩-natural {x = x} {y = y} {Οƒ = Οƒ} {a = a} {b = b} f g =
      D.Ο€*.uniquepβ‚‚ _ _ _ _ _
        (D.pulll[] _ (D.Ο€*.commutesp (sym (sub-proj g)) _)
         D.βˆ™[] D.pullr[] _ (D.Ο€*.commutesp id-comm _))
        (D.pulll[] _ (D.Ο€*.commutesp id-comm _)
         D.βˆ™[] D.pullr[] _ (⟨⨾⟩-weaken g a)
         D.βˆ™[] D.extendl[] _ (∐[]-natural f))

This lets us extend a substitution into a substitution

  opaque
    ∐-sub
      : βˆ€ {Ξ“ Ξ” x y} {Οƒ : Hom Ξ“ Ξ”} {f : E.Hom[ Οƒ ] x y}
      β†’ is-cartesian E Οƒ f β†’ (a : D.Ob[ Ξ” β¨Ύ y ])
      β†’ D.Hom[ id ] (D*.β‚€ Οƒ (∐ y a)) (∐ x (D*.β‚€ (Οƒ β¨ΎΛ’ f) a))
    ∐-sub {x = x} {Οƒ = Οƒ} {f = f} cart a =
      ⟨⨾⟩-cocartesian.universalv cart a ⟨ x , (D*.β‚€ (Οƒ β¨ΎΛ’ f) a) ⟩

    ∐-sub-⟨⨾⟩
      : βˆ€ {Ξ“ Ξ” x y} {Οƒ : Hom Ξ“ Ξ”} {f : E.Hom[ Οƒ ] x y}
      β†’ (cart : is-cartesian E Οƒ f) β†’ (a : D.Ob[ Ξ” β¨Ύ y ])
      β†’ ∐-sub cart a D.∘' ⟨ f β¨Ύ a ⟩ D.≑[ idl _ ] ⟨ x , (D*.β‚€ (Οƒ β¨ΎΛ’ f) a) ⟩
    ∐-sub-⟨⨾⟩ cart a =
      ⟨⨾⟩-cocartesian.commutesv cart a _

    ∐-sub-natural
      : βˆ€ {Ξ“ Ξ” x y} {Οƒ : Hom Ξ“ Ξ”} {f : E.Hom[ Οƒ ] x y}
      β†’ {a b : D.Ob[ Ξ” β¨Ύ y ]}
      β†’ (cart : is-cartesian E Οƒ f)
      β†’ (g : D.Hom[ id ] a b)
      β†’ ∐-sub cart b D.∘' D*.₁ Οƒ ∐[ g ] ≑ ∐[ D*.₁ (Οƒ β¨ΎΛ’ f) g ] D.∘' ∐-sub cart a
    ∐-sub-natural {x = x} {y = y} {Οƒ = Οƒ} {f = f} {a = a} {b = b} cart g =
      ⟨⨾⟩-cocartesian.uniquepβ‚‚ cart a _ _ _ _ _
        (D.pullr[] _ (symP (⟨⨾⟩-natural g f))
         D.βˆ™[] D.pulll[] _ (⟨⨾⟩-cocartesian.commutesv cart b _))
        (D.pullr[] _ (⟨⨾⟩-cocartesian.commutesv cart a _)
         D.βˆ™[] ∐[]-natural (D*.₁ (Οƒ β¨ΎΛ’ f) g))

  1. As a reminder, you can choose to toggle hidden code on the sidebar.β†©οΈŽ

  2. what a mouthful!β†©οΈŽ