module Cat.Displayed.Cartesian
{o ℓ o' ℓ'} {B : Precategory o ℓ} (E : Displayed B o' ℓ') where
open Cat.Reasoning B
open Displayed E
open Cat.Displayed.Morphism E
open DR E
Cartesian morphisms and fibrations🔗
While displayed categories give the essential framework we need to express the idea of families of categories indexed by a category, they do not quite capture the concept precisely. The problem is that while a category displayed over does indeed give a collection of fibre categories this assignment isn’t necessarily functorial!
More precisely, we should have that a collection of categories, indexed by a category, should be a pseudofunctor where is a bicategory of categories. It turns out that we can characterise this assignment entirely in terms of the displayed objects and morphisms in
Fix an arrow in the base category an object over (resp. over and an arrow over We say that is cartesian if, up to very strong handwaving, it fits into a “pullback diagram”. The barred arrows with triangle tips here indicate the “projection” from a displayed object to its base, and so are implicit in the type dependency.
record
{a b a' b'} (f : Hom a b) (f' : Hom[ f ] a' b') : Type (o ⊔ ℓ ⊔ o' ⊔ ℓ')
is-cartesian where
no-eta-equality
More specifically, let and be over and suppose that we have a map (below, in violet), and a map lying over the composite (in red). The universal property of a Cartesian map says that we have a universal factorisation of through a map (in green, marked
field
: ∀ {u u'} (m : Hom u a) (h' : Hom[ f ∘ m ] u' b') → Hom[ m ] u' a'
universal : ∀ {u u'} (m : Hom u a) (h' : Hom[ f ∘ m ] u' b')
commutes → f' ∘' universal m h' ≡ h'
: ∀ {u u'} {m : Hom u a} {h' : Hom[ f ∘ m ] u' b'}
unique → (m' : Hom[ m ] u' a') → f' ∘' m' ≡ h' → m' ≡ universal m h'
Given a “right corner” like that of the diagram below, and note that the input data consists of and over
We also provide some helper functions for working with morphisms that are displayed over something that is propositionally equal to a composite, rather than displayed directly over a composite.
: ∀ {u u'} {m : Hom u a} {k : Hom u b}
universal' → (p : f ∘ m ≡ k) (h' : Hom[ k ] u' b')
→ Hom[ m ] u' a'
{u' = u'} p h' =
universal' _ (coe1→0 (λ i → Hom[ p i ] u' b') h')
universal
: ∀ {u u'} {m : Hom u a} {k : Hom u b}
commutesp → (p : f ∘ m ≡ k) (h' : Hom[ k ] u' b')
→ f' ∘' universal' p h' ≡[ p ] h'
{u' = u'} p h' =
commutesp _ (coe1→0 (λ i → Hom[ p i ] u' b') h')
to-pathp⁻ $ commutes
: ∀ {u u'} {m₁ m₂ : Hom u a} {k : Hom u b}
universalp → (p : f ∘ m₁ ≡ k) (q : m₁ ≡ m₂) (r : f ∘ m₂ ≡ k)
→ (h' : Hom[ k ] u' b')
→ universal' p h' ≡[ q ] universal' r h'
{u = u} p q r h' i =
universalp (is-set→squarep (λ _ _ → Hom-set u b) (ap (f ∘_) q) p r refl i) h'
universal'
: ∀ {u u'} {m₁ m₂ : Hom u a} {k : Hom u b}
uniquep → (p : f ∘ m₁ ≡ k) (q : m₁ ≡ m₂) (r : f ∘ m₂ ≡ k)
→ {h' : Hom[ k ] u' b'}
→ (m' : Hom[ m₁ ] u' a')
→ f' ∘' m' ≡[ p ] h' → m' ≡[ q ] universal' r h'
{h' = h'} m' s =
uniquep p q r (unique m' (from-pathp⁻ s) ∙ from-pathp⁻ (universalp p q r h'))
to-pathp⁻
uniquep₂: ∀ {u u'} {m₁ m₂ : Hom u a} {k : Hom u b}
→ (p : f ∘ m₁ ≡ k) (q : m₁ ≡ m₂) (r : f ∘ m₂ ≡ k)
→ {h' : Hom[ k ] u' b'} (m₁' : Hom[ m₁ ] u' a') (m₂' : Hom[ m₂ ] u' a')
→ f' ∘' m₁' ≡[ p ] h'
→ f' ∘' m₂' ≡[ r ] h'
→ m₁' ≡[ q ] m₂'
{u' = u'} p q r m₁' m₂' α β = to-pathp⁻ $
uniquep₂ (from-pathp⁻ α)
unique m₁' (universalp p q r _)
·· from-pathp⁻ (coe1→0 (λ i → Hom[ q i ] u' a')) (sym (unique m₂' (from-pathp⁻ β))) ·· ap
Furthermore, if is also displayed over there’s a unique vertical map This witnesses the fact that every cartesian map is weakly cartesian.
: ∀ {a''} (f'' : Hom[ f ] a'' b') → Hom[ id ] a'' a'
universalv = universal' (idr _) f''
universalv f''
commutesv: ∀ {x'} → (g' : Hom[ f ] x' b')
→ f' ∘' universalv g' ≡[ idr _ ] g'
= commutesp (idr _)
commutesv
uniquev: ∀ {x'} {g' : Hom[ f ] x' b'}
→ (h' : Hom[ id ] x' a')
→ f' ∘' h' ≡[ idr _ ] g'
→ h' ≡ universalv g'
= uniquep (idr f) refl (idr f) h' p
uniquev h' p
uniquev₂: ∀ {x'} {g' : Hom[ f ] x' b'}
→ (h' h'' : Hom[ id ] x' a')
→ f' ∘' h' ≡[ idr _ ] g'
→ f' ∘' h'' ≡[ idr _ ] g'
→ h' ≡ h''
=
uniquev₂ h' h'' p q (idr f) refl (idr f) h' h'' p q uniquep₂
As the name suggests, being cartesian is a property of a morphism.
is-cartesian-is-prop: ∀ {x y x' y'} {f : Hom x y} {f' : Hom[ f ] x' y'}
→ is-prop (is-cartesian f f')
The proof of this fact is a bunch of cubical nonsense.
{f' = f'} cart cart' = worker where
is-cartesian-is-prop open is-cartesian
: cart ≡ cart'
worker .universal m h' =
worker i .unique (cart .universal m h') (cart .commutes _ _) i
cart' .commutes m h' =
worker i (λ _ _ → Hom[ _ ]-set _ _)
is-set→squarep (ap (f' ∘'_) (cart' .unique _ _))
(cart .commutes m h')
(cart' .commutes m h')
refl i.unique m' p =
worker i (λ _ _ → Hom[ _ ]-set _ _)
is-set→squarep
refl(cart .unique m' p)
(cart' .unique m' p)
(cart' .unique _ _) i
We also provide a bundled form of cartesian morphisms.
record Cartesian-morphism
{x y : Ob} (f : Hom x y) (x' : Ob[ x ]) (y' : Ob[ y ])
: Type (o ⊔ ℓ ⊔ o' ⊔ ℓ') where
no-eta-equality
field
: Hom[ f ] x' y'
hom' : is-cartesian f hom' cartesian
Cartesian-morphism-pathp: ∀ {x y x' y'} {f g : Hom x y}
→ {f' : Cartesian-morphism f x' y'} {g' : Cartesian-morphism g x' y'}
→ {p : f ≡ g}
→ PathP (λ i → Hom[ p i ] x' y') (Cartesian-morphism.hom' f') (Cartesian-morphism.hom' g')
→ PathP (λ i → Cartesian-morphism (p i) x' y') f' g'
.Cartesian-morphism.hom' = q i
Cartesian-morphism-pathp q i {f' = f'} {g' = g'} {p = p} q i .Cartesian-morphism.cartesian =
Cartesian-morphism-pathp (λ i → is-cartesian-is-prop {f = p i} {f' = q i})
is-prop→pathp (Cartesian-morphism.cartesian f')
(Cartesian-morphism.cartesian g') i
Cartesian-morphism-is-set: ∀ {x y x' y'} {f : Hom x y}
→ is-set (Cartesian-morphism f x' y')
= Iso→is-hlevel 2 eqv $
Cartesian-morphism-is-set 2 (Hom[ _ ]-set _ _) λ _ →
Σ-is-hlevel 1 is-cartesian-is-prop
is-hlevel-suc where unquoteDecl eqv = declare-record-iso eqv (quote Cartesian-morphism)
Properties of cartesian morphisms🔗
The composite of 2 cartesian morphisms is in turn cartesian.
cartesian-∘: ∀ {x y z} {f : Hom y z} {g : Hom x y}
→ ∀ {x' y' z'} {f' : Hom[ f ] y' z'} {g' : Hom[ g ] x' y'}
→ is-cartesian f f' → is-cartesian g g'
→ is-cartesian (f ∘ g) (f' ∘' g')
{f = f} {g = g} {f' = f'} {g' = g'} f-cart g-cart = fg-cart where
cartesian-∘
module f' = is-cartesian f-cart
module g' = is-cartesian g-cart
: is-cartesian (f ∘ g) (f' ∘' g')
fg-cart .is-cartesian.universal m h' =
fg-cart .universal m (f'.universal' (assoc f g m) h')
g'.is-cartesian.commutes m h' =
fg-cart (f' ∘' g') ∘' g'.universal m (f'.universal' (assoc f g m) h') ≡⟨ shiftr (sym $ assoc _ _ _) (pullr' refl (g'.commutes m _)) ⟩
(f' ∘' f'.universal' (assoc f g m) h') ≡⟨ hom[]⟩⟨ f'.commutes _ _ ⟩
hom[] (hom[] h') ≡⟨ hom[]-∙ _ _ ∙ liberate _ ⟩
hom[]
h' ∎.is-cartesian.unique {m = m} {h' = h'} m' p =
fg-cart .unique m' $ f'.unique (g' ∘' m') $
g'(assoc' f' g' m') ⟩
f' ∘' g' ∘' m' ≡⟨ from-pathp⁻ ((f' ∘' g') ∘' m') ≡⟨ weave _ _ _ p ⟩
hom[]
hom[] h' ∎
_∘cart_
: ∀ {x y z x' y' z'} {f : Hom y z} {g : Hom x y}
→ Cartesian-morphism f y' z' → Cartesian-morphism g x' y'
→ Cartesian-morphism (f ∘ g) x' z'
= fg' where
f' ∘cart g' open Cartesian-morphism
: Cartesian-morphism _ _ _
fg' .hom' = f' .hom' ∘' g' .hom'
fg' .cartesian = cartesian-∘ (f' .cartesian) (g' .cartesian) fg'
Furthermore, the identity morphism is cartesian.
: ∀ {x x'} → is-cartesian id (id' {x} {x'})
cartesian-id .is-cartesian.universal m h' = hom[ idl m ] h'
cartesian-id .is-cartesian.commutes m h' =
cartesian-id (idl' _) ∙ hom[]-∙ _ _ ∙ liberate _
from-pathp⁻ .is-cartesian.unique m' p =
cartesian-id (symP $ idl' _) ∙ weave _ _ _ p
from-pathp⁻
: ∀ {x} {x' : Ob[ x ]} → Cartesian-morphism id x' x'
idcart .Cartesian-morphism.hom' = id'
idcart .Cartesian-morphism.cartesian = cartesian-id idcart
In fact, every invertible map is also cartesian, as we can use the inverse to construct the requisite factorisation.
invertible→cartesian: ∀ {x y} {f : Hom x y} {x' y'} {f' : Hom[ f ] x' y'}
→ (f-inv : is-invertible f)
→ is-invertible[ f-inv ] f'
→ is-cartesian f f'
invertible→cartesian{f = f} {f' = f'} f-inv f'-inv = f-cart where
module f-inv = is-invertible f-inv
module f'-inv = is-invertible[_] f'-inv
: is-cartesian f f'
f-cart .is-cartesian.universal m h' =
f-cart .invr ] (f'-inv.inv' ∘' h')
hom[ cancell f-inv.is-cartesian.commutes m h' =
f-cart .invr ] (f'-inv.inv' ∘' h') ≡⟨ whisker-r _ ⟩
f' ∘' hom[ cancell f-inv(f' ∘' f'-inv.inv' ∘' h') ≡⟨ revive₁ (cancell' f-inv.invl f'-inv.invl' {q = cancell f-inv.invl}) ⟩
hom[] _ ⟩
hom[] h' ≡⟨ liberate
h' ∎.is-cartesian.unique {h' = h'} m' p =
f-cart _ ⟩
m' ≡˘⟨ liberate (insertl f-inv.invr) (cancell f-inv.invr) (insertl' _ f'-inv.invr') ⟩
hom[] m' ≡⟨ weave refl (f'-inv.inv' ∘' f' ∘' m') ≡⟨ apr' p ⟩
hom[] (f'-inv.inv' ∘' h') ∎ hom[]
If is cartesian, it’s also a weak monomorphism.
cartesian→weak-monic: ∀ {x y} {f : Hom x y}
→ ∀ {x' y'} {f' : Hom[ f ] x' y'}
→ is-cartesian f f'
→ is-weak-monic f'
{f' = f'} f-cart g' g'' p =
cartesian→weak-monic
g' ≡⟨ unique g' p ⟩_ (f' ∘' g'') ≡˘⟨ unique g'' refl ⟩
universal
g'' ∎where
open is-cartesian f-cart
We can use this fact to show that 2 cartesian lifts over the same morphisms must have their domains related by a vertical isomorphism. Suppose they’re called and and fit into a diagram like the one below.
Since and are both Cartesian morphisms, we can factor through by a map and conversely, through by
Since we’re trying to prove that is an isomorphism, we want to show that We know that factors through its own domain, via the identity map. We will show that it also factors through to show that the two are equal, by the universal property of being Cartesian. Consider the following diagram:
We have A symmetric argument shows that is also the identity, so
cartesian-domain-unique: ∀ {x y} {f : Hom x y}
→ ∀ {x' x'' y'} {f' : Hom[ f ] x' y'} {f'' : Hom[ f ] x'' y'}
→ is-cartesian f f'
→ is-cartesian f f''
→ x' ≅↓ x''
{f' = f'} {f'' = f''} f'-cart f''-cart =
cartesian-domain-unique
make-iso[ id-iso ] to* from* invl* invr*where
open is-cartesian
= universal' f''-cart (B .Precategory.idr _) f'
to* = universal' f'-cart (B .Precategory.idr _) f''
from*
: to* ∘' from* ≡[ idl id ] id'
invl* = to-pathp⁻ $ cartesian→weak-monic f''-cart _ _ $
invl* (assoc _ _ _) (pulll' _ (f''-cart .commutes _ _)) ⟩
f'' ∘' to* ∘' from* ≡⟨ shiftr (hom[] f' ∘' from*) ≡⟨ smashl _ _ ⟩
hom[] (f' ∘' from*) ≡⟨ (hom[]⟩⟨ f'-cart .commutes _ _) ∙ hom[]-∙ _ _ ⟩
hom[] _ (sym $ idr _) (ap (_ ∘_) (sym $ idl _)) (symP $ idr' f'') ⟩
hom[] f'' ≡⟨ weave (f'' ∘' id') ≡˘⟨ whisker-r _ ⟩
hom[]
f'' ∘' hom[] id' ∎
: from* ∘' to* ≡[ idl id ] id'
invr* = to-pathp⁻ $ cartesian→weak-monic f'-cart _ _ $
invr* (assoc _ _ _) (pulll' _ (f'-cart .commutes _ _)) ⟩
f' ∘' from* ∘' to* ≡⟨ shiftr (hom[] f'' ∘' to*) ≡⟨ smashl _ _ ⟩
hom[] (f'' ∘' to*) ≡⟨ (hom[]⟩⟨ f''-cart .commutes _ _) ∙ hom[]-∙ _ _ ⟩
hom[] _ (sym $ idr _) (ap (_ ∘_) (sym $ idl _)) (symP $ idr' f') ⟩
hom[] f' ≡⟨ weave (f' ∘' id') ≡˘⟨ whisker-r _ ⟩
hom[] f' ∘' hom[] id' ∎
Cartesian morphisms are also stable under vertical retractions.
cartesian-vertical-retraction-stable: ∀ {x y} {f : Hom x y}
→ ∀ {x' x'' y'} {f' : Hom[ f ] x' y'} {f'' : Hom[ f ] x'' y'} {ϕ : Hom[ id ] x' x''}
→ is-cartesian f f'
→ has-section↓ ϕ
→ f'' ∘' ϕ ≡[ idr _ ] f'
→ is-cartesian f f''
{f' = f'} {f''} {ϕ} f-cart ϕ-sect factor = f''-cart where
cartesian-vertical-retraction-stable open is-cartesian f-cart
module ϕ-sect = has-section[_] ϕ-sect
: is-cartesian _ f''
f''-cart .is-cartesian.universal m h' =
f''-cart (ϕ ∘' universal m h')
hom[ idl m ] .is-cartesian.commutes m h' =
f''-cart (ϕ ∘' universal m h') ≡⟨ whisker-r _ ⟩
f'' ∘' hom[] (f'' ∘' ϕ ∘' universal m h') ≡⟨ revive₁ {p = ap (_ ∘_) (idl m)} (pulll' (idr _) factor) ⟩
hom[] (f' ∘' universal m h') ≡⟨ (hom[]⟩⟨ commutes m h') ∙ liberate _ ⟩
hom[]
h' ∎.is-cartesian.unique {m = m} {h' = h'} m' p =
f''-cart (sym (eliml (idl _))) (introl' (idl _) ϕ-sect.is-section') ⟩
m' ≡⟨ shiftr ((ϕ ∘' ϕ-sect.section') ∘' m') ≡⟨ weave _ (pullr (idl _)) _ (pullr' (idl _) (to-pathp (unique _ unique-path))) ⟩
hom[] (ϕ ∘' universal m h') ∎
hom[] where
: f' ∘' ϕ-sect.section' ≡[ idr _ ] f''
sect-commute = to-pathp⁻ $
sect-commute .section' ≡⟨ shiftr _ (λ i → factor (~ i) ∘' ϕ-sect.section') ⟩
f' ∘' ϕ-sect((f'' ∘' ϕ) ∘' ϕ-sect.section') ≡⟨ weave _ (idr _ ∙ idr _) _ (cancelr' (idl _) ϕ-sect.is-section') ⟩
hom[]
hom[] f'' ∎
: f' ∘' hom[ idl m ] (ϕ-sect.section' ∘' m') ≡ h'
unique-path =
unique-path (ϕ-sect.section' ∘' m') ≡⟨ whisker-r _ ⟩
f' ∘' hom[ idl m ] (f' ∘' ϕ-sect.section' ∘' m') ≡⟨ shiftl _ (pulll' (idr _) sect-commute) ⟩
hom[]
f'' ∘' m' ≡⟨ p ⟩ h' ∎
We also have the following extremely useful pasting lemma, which generalizes the pasting law for pullbacks.
cartesian-pasting: ∀ {x y z} {f : Hom y z} {g : Hom x y}
→ ∀ {x' y' z'} {f' : Hom[ f ] y' z'} {g' : Hom[ g ] x' y'}
→ is-cartesian f f'
→ is-cartesian (f ∘ g) (f' ∘' g')
→ is-cartesian g g'
{f = f} {g = g} {f' = f'} {g' = g'} f-cart fg-cart = g-cart where
cartesian-pasting open is-cartesian
: is-cartesian g g'
g-cart .universal m h' =
g-cart (sym (assoc _ _ _)) (f' ∘' h')
universal' fg-cart .commutes m h' =
g-cart (sym (assoc _ _ _)) (f' ∘' h') ≡⟨ f-cart .unique _ (from-pathp⁻ (assoc' _ _ _) ∙ from-pathp (commutesp fg-cart _ _)) ⟩
g' ∘' universal' fg-cart .universal _ (f' ∘' h') ≡˘⟨ f-cart .unique h' refl ⟩
f-cart
h' ∎.unique {m = m} {h' = h'} m' p =
g-cart (sym (assoc _ _ _)) refl (sym (assoc _ _ _)) m' (pullr' refl p) uniquep fg-cart
We can prove a similar fact for bundled cartesian morphisms.
cart-paste: ∀ {x y z x' y' z'} {f : Hom y z} {g : Hom x y}
→ Cartesian-morphism f y' z'
→ Cartesian-morphism (f ∘ g) x' z'
→ Cartesian-morphism g x' y'
{x' = x'} {y' = y'} {f = f} {g = g} f' fg' = g' where
cart-paste open Cartesian-morphism
open is-cartesian
module f' = is-cartesian (f' .cartesian)
module fg' = is-cartesian (fg' .cartesian)
: Cartesian-morphism g x' y'
g' .hom' = f'.universal g (fg' .hom')
g' .cartesian .universal m h' =
g' .universal' (sym (assoc _ _ _)) (f' .hom' ∘' h')
fg'.cartesian .commutes m h' =
g' .uniquep₂ _ _ (assoc _ _ _) _ _
f'(pulll[] _ (f'.commutes _ _) ∙[] fg'.commutes _ _)
(to-pathp refl)
.cartesian .unique m' p =
g' .uniquep _ refl (sym (assoc _ _ _)) m'
fg'(ap (_∘' m') (symP (f'.commutes _ _)) ∙[] pullr[] _ p)
If a morphism is both vertical and cartesian, then it must be an isomorphism. We can construct the inverse by factorizing the identity morphism, which is possible due to the fact that is vertical.
vertical+cartesian→invertible: ∀ {x} {x' x'' : Ob[ x ]} {f' : Hom[ id ] x' x''}
→ is-cartesian id f'
→ is-invertible↓ f'
{x' = x'} {x'' = x''} {f' = f'} f-cart =
vertical+cartesian→invertible where
make-invertible↓ f⁻¹' f'-invl f'-invr open is-cartesian f-cart
: Hom[ id ] x'' x'
f⁻¹' = universal' (idl _) id'
f⁻¹'
: f' ∘' f⁻¹' ≡[ idl _ ] id'
f'-invl = commutesp _ id'
f'-invl
: f' ∘' f⁻¹' ∘' f' ≡[ elimr (idl _) ] f'
path = cancell' (idl _) (commutesp (idl _) id')
path
: f⁻¹' ∘' f' ≡[ idl _ ] id'
f'-invr = to-pathp⁻ $
f'-invr _ (uniquep _ (idl _) (idl _) (f⁻¹' ∘' f') path) ⟩
f⁻¹' ∘' f' ≡⟨ shiftr (universal' (idl _) f') ≡⟨ weave _ _ _ (symP $ uniquep (idr _) refl (idl _) id' (idr' _)) ⟩
hom[] hom[] id' ∎
Furthermore, is cartesian if and only if the function is an equivalence.
postcompose-equiv→cartesian: ∀ {x y x' y'} {f : Hom x y}
→ (f' : Hom[ f ] x' y')
→ (∀ {w w'} {g : Hom w x} → is-equiv {A = Hom[ g ] w' x'} (f' ∘'_))
→ is-cartesian f f'
.is-cartesian.universal m h' =
postcompose-equiv→cartesian f' eqv
equiv→inverse eqv h'.is-cartesian.commutes m h' =
postcompose-equiv→cartesian f' eqv
equiv→counit eqv h'.is-cartesian.unique m' p =
postcompose-equiv→cartesian f' eqv (equiv→unit eqv m') ∙ ap (equiv→inverse eqv) p
sym
cartesian→postcompose-equiv: ∀ {x y z x' y' z'} {f : Hom y z} {g : Hom x y} {f' : Hom[ f ] y' z'}
→ is-cartesian f f'
→ is-equiv {A = Hom[ g ] x' y'} (f' ∘'_)
=
cartesian→postcompose-equiv cart
is-iso→is-equiv $(universal _)
iso (commutes _)
(λ g' → sym (unique g' refl))
where open is-cartesian cart
Cartesian lifts🔗
We call an object over together with a Cartesian arrow a Cartesian lift of Cartesian lifts, defined by universal property as they are, are unique when they exist, so that “having Cartesian lifts” is a property, not a structure.
record
{x y} (f : Hom x y) (y' : Ob[ y ]) : Type (o ⊔ ℓ ⊔ o' ⊔ ℓ')
Cartesian-lift where
no-eta-equality
field
{x'} : Ob[ x ]
: Hom[ f ] x' y'
lifting : is-cartesian f lifting
cartesian open is-cartesian cartesian public
We note that the classical literature often differentiates between fibrations — those displayed categories for which there exist Cartesian lifts for every right corner — and cloven fibrations, those for which the Cartesian lifts are “algebraic” in a sense. This is because, classically, essentially unique means that there are still some choices to be made, and invoking the axiom of choice makes an “arbitrary” set of such choices. But, in the presence of univalence, there is exactly one choice to be made, that is, no choice at all. Thus, we do not dwell on the distinction.
record Cartesian-fibration : Type (o ⊔ ℓ ⊔ o' ⊔ ℓ') where
no-eta-equality
field
: ∀ {x y} (f : Hom x y) (y' : Ob[ y ]) → Cartesian-lift f y'
has-lift
module has-lift {x y} (f : Hom x y) (y' : Ob[ y ]) =
(has-lift f y') Cartesian-lift
Note that if is a fibration, we can define an operation that allows us to move vertical morphisms between fibres. This actually extends to a collection of functors, called base change functors. This operation is also definable for weak fibrations, as it only uses the universal property that yields a vertical morphism.
: ∀ {x y y' y''} → (f : Hom x y)
rebase → Hom[ id ] y' y''
→ Hom[ id ] (has-lift.x' f y') (has-lift.x' f y'')
=
rebase f vert .universal' f _ id-comm (vert ∘' has-lift.lifting f _) has-lift
A Cartesian fibration is a displayed category having Cartesian lifts for every right corner.
Why?🔗
Admittedly, the notion of Cartesian morphism is slightly artificial. It arises from studying the specific properties of the pullback functors which exist in a category of pullbacks, hence the similarity in universal property!
In fact, the quintessential example of a Cartesian fibration is the codomain fibration, which is a category displayed over where the fibre over is the slice category We may investigate further (to uncover the name “codomain”): the total space of this fibration is the arrow category and the projection functor is the codomain functor
This displayed category extends to a pseudofunctor exactly when has all pullbacks, because in a world where the vertical arrows are “just” arrows, a Cartesian morphism is exactly a pullback square.
Other examples exist:
- The family fibration exhibits any category as displayed over The fibres are functor categories (with discrete domains), reindexing is given by composition.
- The category of modules is fibred over the category of rings. The fibre over a ring is the category of Cartesian lifts are given by restriction of scalars.