module Cat.Instances.Slice where
private variable
: Level
o ℓ o' ℓ' open Functor
open _=>_
open _⊣_
module _ {o ℓ} {C : Precategory o ℓ} where
private
module C = Cat.Reasoning C
variable a b c : C.Ob
Slice categories🔗
When working in there is an evident notion of family indexed by a set: a family of sets is equivalently a functor where we have equipped the set with the discrete category structure. This works essentially because of the discrete category-global sections adjunction, but in general this can not be applied to other categories, like How, then, should we work with “indexed families” in general categories?
The answer is to consider, rather than families themselves, the projections from their total spaces as the primitive objects. A family indexed by then, would consist of an object and a morphism where is considered as the “total space” object and assigns gives the “tag” of each object. By analysing how pulls back along maps we recover a notion of “fibres”: the collection with index can be recovered as the pullback
Note that, since the discussion in the preceding paragraph made no mention of the category of sets, it applies in any category! More generally, for any category and object we have a category of objects indexed by , the slice category An object of “the slice over ” is given by an object to serve as the domain, and a map
record /-Obj (c : C.Ob) : Type (o ⊔ ℓ) where
constructor cut
field
{domain} : C.Ob
: C.Hom domain c map
A map between and is given by a map such that the triangle below commutes. Since we’re thinking of and as families indexed by , commutativity of the triangle says that the map “respects reindexing”, or less obliquely “preserves fibres”.
record /-Hom (a b : /-Obj c) : Type ℓ where
no-eta-equality
private
module a = /-Obj a
module b = /-Obj b
field
: C.Hom a.domain b.domain
map : b.map C.∘ map ≡ a.map commutes
: ∀ {c} {x y : /-Obj c}
/-Obj-path → (p : x ./-Obj.domain ≡ y ./-Obj.domain)
→ PathP (λ i → C.Hom (p i) c) (x ./-Obj.map) (y ./-Obj.map)
→ x ≡ y
./-Obj.domain = p i
/-Obj-path p q i ./-Obj.map = q i
/-Obj-path p q i
: ∀ {c a a' b b'} (p : a ≡ a') (q : b ≡ b')
/-Hom-pathp {x : /-Hom {c = c} a b} {y : /-Hom a' b'}
→ PathP (λ i → C.Hom (p i ./-Obj.domain) (q i ./-Obj.domain))
(x ./-Hom.map) (y ./-Hom.map)
→ PathP (λ i → /-Hom (p i) (q i)) x y
{x} {y} r = path where
/-Hom-pathp p q open /-Hom
: PathP (λ i → /-Hom (p i) (q i)) x y
path .map = r i
path i .commutes =
path i
is-prop→pathp(λ i → C.Hom-set (p i ./-Obj.domain) _
(q i ./-Obj.map C.∘ r i) (p i ./-Obj.map))
(x .commutes) (y .commutes) i
: ∀ {c a b} {x y : /-Hom {c = c} a b}
/-Hom-path → x ./-Hom.map ≡ y ./-Hom.map
→ x ≡ y
= /-Hom-pathp refl refl
/-Hom-path
instance
Extensional-/-Hom: ∀ {c a b ℓ} ⦃ sa : Extensional (C.Hom (/-Obj.domain a) (/-Obj.domain b)) ℓ ⦄
→ Extensional (/-Hom {c = c} a b) ℓ
= injection→extensional! (/-Hom-pathp refl refl) sa
Extensional-/-Hom ⦃ sa ⦄
unquoteDecl H-Level-/-Hom = declare-record-hlevel 2 H-Level-/-Hom (quote /-Hom)
The slice category
is given by the /-Obj
and /-Hom
s.
: (C : Precategory o ℓ) → ⌞ C ⌟ → Precategory _ _
Slice = precat where
Slice C c import Cat.Reasoning C as C
open Precategory
open /-Hom
open /-Obj
: Precategory _ _
precat .Ob = /-Obj {C = C} c
precat .Hom = /-Hom
precat .Hom-set x y = hlevel 2
precat .id .map = C.id
precat .id .commutes = C.idr _ precat
For composition in the slice over note that if the triangle (the commutativity condition for and the rhombus (the commutativity condition for both commute, then so does the larger triangle (the commutativity for
._∘_ {x} {y} {z} f g = fog where
precat module f = /-Hom f
module g = /-Hom g
: /-Hom _ _
fog .map = f.map C.∘ g.map
fog .commutes =
fog .map C.∘ f.map C.∘ g.map ≡⟨ C.pulll f.commutes ⟩
z .map C.∘ g.map ≡⟨ g.commutes ⟩
y .map ∎
x .idr f = ext (C.idr _)
precat .idl f = ext (C.idl _)
precat .assoc f g h = ext (C.assoc _ _ _) precat
There is an evident projection functor from to that only remembers the domains.
module _ {o ℓ} {C : Precategory o ℓ} {c} where
open /-Hom
open /-Obj
private
module C = Cat.Reasoning C
module C/c = Cat.Reasoning (Slice C c)
: Functor (Slice C c) C
Forget/ .F₀ o = o .domain
Forget/ .F₁ f = f .map
Forget/ .F-id = refl
Forget/ .F-∘ _ _ = refl Forget/
Furthermore, this forgetful functor is easily seen to be faithful and conservative: if is a morphism in whose underlying map has an inverse in then clearly also makes the triangle commute, so that is invertible in
: is-faithful Forget/
Forget/-is-faithful = ext p
Forget/-is-faithful p
: is-conservative Forget/
Forget/-is-conservative {f = f} i =
Forget/-is-conservative .make-invertible f⁻¹ (ext i.invl) (ext i.invr)
C/cwhere
module i = C.is-invertible i
: /-Hom _ _
f⁻¹ .map = i.inv
f⁻¹ .commutes = C.rswizzle (sym (f .commutes)) i.invl f⁻¹
Finite limits🔗
We discuss the construction of finite limits in the slice of First, every slice category has a terminal object, given by the identity map
module _ {o ℓ} {C : Precategory o ℓ} {c : ⌞ C ⌟} where
import Cat.Reasoning C as C
import Cat.Reasoning (Slice C c) as C/c
open /-Hom
open /-Obj
: is-terminal (Slice C c) (cut C.id)
Slice-terminal-object' .centre .map = obj .map
Slice-terminal-object' obj .centre .commutes = C.idl _
Slice-terminal-object' obj .paths other =
Slice-terminal-object' obj (sym (other .commutes) ∙ C.idl _)
ext
: Terminal (Slice C c)
Slice-terminal-object .Terminal.top = _
Slice-terminal-object .Terminal.has⊤ = Slice-terminal-object' Slice-terminal-object
module _ {o ℓ} {C : Precategory o ℓ} {c : ⌞ C ⌟} where
import Cat.Reasoning C as C
import Cat.Reasoning (Slice C c) as C/c
private variable
: C.Ob
a b : C.Hom a b
f g π₁ π₂ open /-Hom
open /-Obj
Products in a slice category are slightly more complicated — but if you’ve ever heard a pullback referred to as a “fibre(d) product”, you already know what we’re up to! Indeed, in defining pullback diagrams, we noted that the pullback of and is exactly the product in the slice over
Suppose we have a pullback diagram like the one below, i.e., a limit of the diagram in the category We’ll show that it’s also a limit of the (discrete) diagram consisting of and but now in the slice category
For starters, we have to define a map to serve as the actual object in the slice. It seems like there are two choices, depending on how you trace out the square. But the square commutes, which by definition means that we don’t really have a choice. We choose i.e. following the top face then the right face, for definiteness.
module
_ {f g : /-Obj c} {Pb : C.Ob} {π₁ : C.Hom Pb (f .domain)}
{π₂ : C.Hom Pb (g .domain)}
(pb : is-pullback C {X = f .domain} {Z = c} {Y = g .domain} {P = Pb}
(map {_} {_} {C} {c} f) π₂ (map {_} {_} {C} {c} g))
π₁ where
private module pb = is-pullback pb
: C/c.Ob
is-pullback→product-over = cut (f .map C.∘ π₁) is-pullback→product-over
Let us turn to the projection maps. Again by commutativity of the square, the pullback projection maps and extend directly to maps from into and over In the nontrivial case, we have to show that which is exactly what commutativity of the square gets us.
: C/c.Hom is-pullback→product-over f
is-pullback→π₁ .map = π₁
is-pullback→π₁ .commutes = refl
is-pullback→π₁
: C/c.Hom is-pullback→product-over g
is-pullback→π₂ .map = π₂
is-pullback→π₂ .commutes = sym pb.square
is-pullback→π₂
open is-product
Now we turn to showing that these projections are universal, in the slice Given a family the data of maps and over is given precisely by evidence that meaning that they fit neatly around our pullback diagram, as shown in the square below.
And, as indicated in the diagram, since this square is a pullback, we can obtain the dashed map which we can calculate satisfies
so that it is indeed a map over as required. Reading out the rest of universal property, we see that is the unique map which satisfies and exactly as required for the pairing of a product in
is-pullback→is-fibre-product: is-product (Slice C c) is-pullback→π₁ is-pullback→π₂
.⟨_,_⟩ {Q} p q = factor where
is-pullback→is-fibre-product module p = /-Hom p
module q = /-Hom q
: C/c.Hom Q _
factor .map = pb.universal (p.commutes ∙ sym q.commutes)
factor .commutes =
factor (f .map C.∘ π₁) C.∘ pb.universal _ ≡⟨ C.pullr pb.p₁∘universal ⟩
.map C.∘ p.map ≡⟨ p.commutes ⟩
f .map ∎ Q
.π₁∘⟨⟩ = ext pb.p₁∘universal
is-pullback→is-fibre-product .π₂∘⟨⟩ = ext pb.p₂∘universal
is-pullback→is-fibre-product .unique p q =
is-pullback→is-fibre-product (pb.unique (ap map p) (ap map q))
ext
Pullback→Fibre-product: ∀ {f g : /-Obj c}
→ Pullback C (f .map) (g .map) → Product (Slice C c) f g
.Product.apex = _
Pullback→Fibre-product pb .Product.π₁ = _
Pullback→Fibre-product pb .Product.π₂ = _
Pullback→Fibre-product pb .Product.has-is-product =
Pullback→Fibre-product pb (pb .Pullback.has-is-pb) is-pullback→is-fibre-product
module _
{f g : /-Obj c} {p : /-Obj c} {π₁ : C/c.Hom p f} {π₂ : C/c.Hom p g}
(prod : is-product (Slice C c) π₁ π₂)
where
private module prod = is-product prod
We can go in the other direction as well, hence products in a slice category correspond precisely to pullbacks in the base category.
open is-pullback
: is-pullback C (π₁ .map) (f .map) (π₂ .map) (g .map)
is-fibre-product→is-pullback .square = π₁ .commutes ∙ sym (π₂ .commutes)
is-fibre-product→is-pullback .universal {P} {p₁} {p₂} square =
is-fibre-product→is-pullback .⟨ record { map = p₁ ; commutes = refl }
prodrecord { map = p₂ ; commutes = sym square } ⟩ .map
, .p₁∘universal = ap map prod.π₁∘⟨⟩
is-fibre-product→is-pullback .p₂∘universal = ap map prod.π₂∘⟨⟩
is-fibre-product→is-pullback .unique {lim' = lim'} fac₁ fac₂ = ap map $
is-fibre-product→is-pullback .unique
prod{other = record { map = lim' ; commutes = ap (C._∘ lim') (sym (π₁ .commutes)) ∙ C.pullr fac₁}}
(ext fac₁) (ext fac₂)
While products and terminal objects in do not correspond to those in pullbacks (and equalisers) are precisely equivalent. A square is a pullback in precisely if its image in forgetting the maps to is a pullback.
The “if” direction (what is pullback-above→pullback-below
) in the
code is essentially an immediate consequence of the fact that equality
of morphisms in
may be tested in
but we do have to take some care when extending the “universal” morphism
back down to the slice category (see the calculation marked {- * -}
).
module _ {o ℓ} {C : Precategory o ℓ} {X : ⌞ C ⌟}
{P A B c} {p1 f p2 g}
where
open Cat.Reasoning C
open is-pullback
open /-Obj
open /-Hom
pullback-above→pullback-below: is-pullback C (p1 .map) (f .map) (p2 .map) (g .map)
→ is-pullback (Slice C X) {P} {A} {B} {c} p1 f p2 g
= pb' where
pullback-above→pullback-below pb : is-pullback (Slice _ _) _ _ _ _
pb' .square = ext (pb .square)
pb' .universal p .map = pb .universal (ap map p)
pb' .universal {P'} {p₁' = p₁'} p .commutes =
pb' (c .map ∘ pb .universal (ap map p)) ≡˘⟨ pulll (p1 .commutes) ⟩
(P .map ∘ p1 .map ∘ pb .universal (ap map p)) ≡⟨ ap (_ ∘_) (pb .p₁∘universal) ⟩
(P .map ∘ p₁' .map) ≡⟨ p₁' .commutes ⟩
.map ∎ {- * -}
P' .p₁∘universal = ext (pb .p₁∘universal)
pb' .p₂∘universal = ext (pb .p₂∘universal)
pb' .unique p q = ext (pb .unique (ap map p) (ap map q))
pb'
pullback-below→pullback-above: is-pullback (Slice C X) {P} {A} {B} {c} p1 f p2 g
→ is-pullback C (p1 .map) (f .map) (p2 .map) (g .map)
= pb' where
pullback-below→pullback-above pb : is-pullback _ _ _ _ _
pb' .square = ap map (pb .square)
pb' .universal p = pb .universal
pb' {p₁' = record { commutes = refl }}
{p₂' = record { commutes = sym (pulll (g .commutes))
(ap (_ ∘_) p)
·· sym (f .commutes) }}
·· pulll (ext p) .map
.p₁∘universal = ap map $ pb .p₁∘universal
pb' .p₂∘universal = ap map $ pb .p₂∘universal
pb' .unique p q = ap map $ pb .unique
pb' {lim' = record { commutes = sym (pulll (p1 .commutes)) ∙ ap (_ ∘_) p }}
(ext p) (ext q)
It follows that any slice of a category with pullbacks is finitely complete. Note that we could have obtained the products abstractly, since any category with pullbacks and a terminal object has products, but expanding on the definition in components helps clarify both their construction and the role of pullbacks.
module _ {o ℓ} {C : Precategory o ℓ} where
open Cat.Reasoning C
open Pullback
open /-Obj
open /-Hom
: ∀ {b} → has-pullbacks C → has-pullbacks (Slice C b)
Slice-pullbacks : ∀ {b} → has-pullbacks C → has-products (Slice C b)
Slice-products : ∀ {b} → has-pullbacks C → Finitely-complete (Slice C b) Slice-lex
This is one of the rare moments when the sea of theory has already risen
to meet our prose — put another way, the proofs of the statements above
are just putting things together. We leave them in this
<details>
block for the curious reader.
{A = A} f g = pb where
Slice-pullbacks pullbacks : Pullback (Slice C _) _ _
pb .apex = cut (A .map ∘ pullbacks _ _ .p₁)
pb .p₁ = record { commutes = refl }
pb .p₂ = record { commutes =
pb (pushl (sym (f .commutes))
sym _∘_ refl (pullbacks _ _ .square)
·· ap₂ (g .commutes)) }
·· pulll .has-is-pb = pullback-above→pullback-below $
pb (f .map) (g .map) .has-is-pb
pullbacks
= Pullback→Fibre-product (pullbacks _ _)
Slice-products pullbacks f g
= with-pullbacks (Slice C _)
Slice-lex pb
Slice-terminal-object(Slice-pullbacks pb)
Slices of Sets🔗
module _ {I : Set ℓ} where
open /-Obj
open /-Hom
Having tackled some properties of slice categories in general, we now turn to characterising the slice as the category of families of indexed by , by formalising the aforementioned equivalence Let us trace our route before we depart, and write down an outline of the proof.
A functor is sent to the first projection functorially. Since is the total space of we refer to this as the
Total-space
functor.We define a procedure that, given a morphism over recovers a natural transformation We then calculate that this procedure is an inverse to the action on morphisms of
Total-space
, so that it is fully faithful.Finally, we show that, given the assignment sending an index to the fibre of over it, gives a functor and that over so that
Total-space
is a split essential surjection, hence an equivalence of categories.
Let’s begin. The Total-space
functor gets out of our way very fast:
: Functor Cat[ Disc' I , Sets ℓ ] (Slice (Sets ℓ) I)
Total-space .F₀ F .domain = el! (Σ _ (∣_∣ ⊙ F₀ F))
Total-space .F₀ F .map = fst
Total-space
.F₁ nt .map (i , x) = i , nt .η _ x
Total-space .F₁ nt .commutes = refl
Total-space
.F-id = trivial!
Total-space .F-∘ _ _ = trivial! Total-space
Since the construction of the functor itself is straightforward, we turn to showing it is fully faithful. The content of a morphism over is a function
which preserves the first projection, i.e. we have This means that it corresponds to a dependent function and, since the morphisms in are trivial, this dependent function is automatically a natural transformation.
: is-fully-faithful Total-space
Total-space-is-ff {F} {G} = is-iso→is-equiv $
Total-space-is-ff (λ x → ext λ _ _ → transport-refl _) where
iso from linv
: /-Hom (Total-space .F₀ F) (Total-space .F₀ G) → F => G
from = nt where
from mp : ∀ i → F ʻ i → G ʻ i
eta = subst (G ʻ_) (mp .commutes # _) (mp .map (i , j) .snd)
eta i j
: F => G
nt .η = eta
nt .is-natural _ _ = J (λ _ p → eta _ ⊙ F .F₁ p ≡ G .F₁ p ⊙ eta _) $
nt _ ⊙ F .F₁ refl ≡⟨ ap (eta _ ⊙_) (F .F-id) ⟩
eta _ ≡˘⟨ ap (_⊙ eta _) (G .F-id) ⟩
eta .F₁ refl ⊙ eta _ ∎ G
: is-left-inverse (F₁ Total-space) from
linv = ext λ y s → Σ-pathp (sym (x .commutes $ₚ _)) (to-pathp⁻ refl) linv x
All that remains is showing that sending a map
to the total space of its family of fibres gets us all the way back
around to
Fortunately, our proof that universes are object classifiers
grappled with many of the same concerns, so we have a reusable
equivalence Total-equiv
which slots
right in. By univalence, we can finish in style: not only is
isomorphic to
in
it’s actually identical to
: is-split-eso Total-space
Total-space-is-eso = functor , path→iso path where
Total-space-is-eso fam : Functor _ _
functor .F₀ i = el! (fibre (fam .map) i)
functor .F₁ p = subst (fibre (fam .map)) p
functor .F-id = funext transport-refl
functor .F-∘ f g = funext (subst-∙ (fibre (fam .map)) _ _)
functor
: Total-space .F₀ functor ≡ fam
path = /-Obj-path (n-ua (Total-equiv _ e⁻¹)) (ua→ λ a → sym (a .snd .snd)) path
Slices preserve univalence🔗
In passing, we can put together the following result: any slice of a univalent category is univalent again. That’s because an identification is given by an identification and a proof that over We already have a characterisation of dependent identifications in a space of morphisms, in terms of composition with the induced isomorphisms, so that this latter condition reduces to showing
module _ {C : Precategory o ℓ} {o : ⌞ C ⌟} (isc : is-category C) where
private
module C = Cat.Reasoning C
module C/o = Cat.Reasoning (Slice C o)
open /-Obj
open /-Hom
open C/o._≅_
open C._≅_
Therefore, if we have an isomorphism over and its component in corresponds to an identification then the commutativity of is exactly the data we need to extend this to an identification
: is-category (Slice C o)
slice-is-category .to-path x = /-Obj-path (isc .to-path $
slice-is-category .make-iso (x .to .map) (x .from .map)
C(ap map (C/o._≅_.invl x)) (ap map (C/o._≅_.invr x)))
(Univalent.Hom-pathp-refll-iso isc (x .from .commutes))
.to-path-over x = C/o.≅-pathp refl _ $
slice-is-category _ _ (Univalent.Hom-pathp-reflr-iso isc (C.idr _)) /-Hom-pathp
open /-Obj
open /-Hom
Constant families🔗
If is the category of families of indexed by , it stands to reason that we should be able to consider any object as a family over where each fibre of the family is isomorphic to Type-theoretically, this would correspond to taking a closed type and trivially regarding it as living in a context by ignoring the context entirely.
From the perspective of slice categories, the constant families functor, sends an object to the projection morphism
module _ {o ℓ} {C : Precategory o ℓ} {B} (prod : has-products C) where
open Binary-products C prod
open Cat.Reasoning C
: Functor C (Slice C B)
constant-family .F₀ A = cut (π₂ {a = A})
constant-family .F₁ f .map = ⟨ f ∘ π₁ , π₂ ⟩
constant-family .F₁ f .commutes = π₂∘⟨⟩
constant-family .F-id = ext (sym (⟨⟩-unique id-comm (idr _)))
constant-family .F-∘ f g = ext $ sym $
constant-family (pulll π₁∘⟨⟩ ∙ extendr π₁∘⟨⟩) (pulll π₂∘⟨⟩ ∙ π₂∘⟨⟩) ⟨⟩-unique
We can observe that this really is a constant families functor by performing the following little calculation: If we have a map then the fibre of over given by the pullback
is isomorphic to The extra generality makes it a bit harder to see the constancy, but if were a point the fibre over would correspondingly be isomorphic to
constant-family-fibre: (pb : has-pullbacks C)
→ ∀ {A Y} (h : Hom Y B)
→ pb (constant-family .F₀ A .map) h .Pullback.apex ≅ (A ⊗₀ Y)
{A} h = make-iso
constant-family-fibre pb (universal {p₁' = ⟨ π₁ , h ∘ π₂ ⟩} {p₂' = π₂} π₂∘⟨⟩)
⟨ π₁ ∘ p₁ , p₂ ⟩ (⟨⟩∘ _ ∙ sym (Product.unique (prod _ _)
(idr _ ∙ sym (pullr p₁∘universal ∙ π₁∘⟨⟩))
(idr _ ∙ sym p₂∘universal)))
(Pullback.unique₂ (pb _ _) {p = π₂∘⟨⟩ ∙ square}
(pulll p₁∘universal ∙ ⟨⟩∘ _ ∙ ap₂ ⟨_,_⟩ π₁∘⟨⟩ (pullr π₂∘⟨⟩ ∙ sym square))
(pulll p₂∘universal ∙ π₂∘⟨⟩)
(idr _ ∙ Product.unique (prod _ _) refl refl)
(idr _))
where open Pullback (pb (constant-family .F₀ A .map) h)
The constant families functor is a right adjoint to the
projection
This can be understood in terms of base change: if
has a terminal
object
then the slice
is equivalent to
and the unique map
induces a pullback functor
that is just the constant families functor. On the other hand, the
“dependent sum” functor sends a map
to the unique composite
it simply Forget/
s the map. Thus
the following adjunction is a special case of the adjunction between
dependent sum and base change.
: Forget/ ⊣ constant-family
Forget⊣constant-family .unit .η X .map = ⟨ id , X .map ⟩
Forget⊣constant-family .unit .η X .commutes = π₂∘⟨⟩
Forget⊣constant-family .unit .is-natural _ _ f = ext (⟨⟩-unique₂
Forget⊣constant-family (pulll π₁∘⟨⟩ ∙ id-comm-sym)
(pulll π₂∘⟨⟩ ∙ f .commutes)
(pulll π₁∘⟨⟩ ∙ pullr π₁∘⟨⟩)
(pulll π₂∘⟨⟩ ∙ π₂∘⟨⟩))
.counit .η x = π₁
Forget⊣constant-family .counit .is-natural _ _ f = π₁∘⟨⟩
Forget⊣constant-family .zig = π₁∘⟨⟩
Forget⊣constant-family .zag = ext (⟨⟩-unique₂
Forget⊣constant-family (pulll π₁∘⟨⟩ ∙ pullr π₁∘⟨⟩)
(pulll π₂∘⟨⟩ ∙ π₂∘⟨⟩)
refl(idr _))