module Cat.Instances.Sheaf.Omega {ℓ} {C : Precategory ℓ ℓ} (J : Coverage C ℓ) where
Closed sieves and the subobject classifier🔗
open is-generic-subobject
open is-pullback-along
open is-pullback
open Subobject
open Functor
open Soc C
open Cat C
private
module ΩPSh = Subobject-classifier PSh-omega
module PSh = Cat (PSh ℓ C)
open Coverage J using (Membership-covers ; Sem-covers)
The category of sheaves on a small site is a topos, which means that — in addition to finite limits and exponential objects — it has a subobject classifier, a sheaf which plays the role of the “universe of propositions”. We can construct the sheaf explicitly, as the sheaf of sieves.
A sieve is if it contains every morphism it covers. This notion is evidently closed under pullback, so the closed sieves form a presheaf on
: ∀ {U} → Sieve C U → Type ℓ
is-closed {U} S = ∀ {V} (h : Hom V U) → J ∋ pullback h S → h ∈ S
is-closed
abstract
is-closed-pullback: ∀ {U V} (f : Hom V U) (S : Sieve C U)
→ is-closed S → is-closed (pullback f S)
= c (f ∘ h) (subst (J ∋_) (sym pullback-∘) p) is-closed-pullback f S c h p
instance
: ∀ {ℓr U} ⦃ _ : Extensional (Sieve C U) ℓr ⦄ → Extensional (Σ[ R ∈ Sieve C U ] is-closed R) ℓr
Extensional-closed-sieve = injection→extensional! Σ-prop-path! e Extensional-closed-sieve ⦃ e ⦄
: Sheaf J ℓ
ΩJ .fst = pre where
ΩJ : Functor (C ^op) (Sets ℓ)
pre .F₀ U = el! (Σ[ R ∈ Sieve C U ] is-closed R)
pre .F₁ f (R , c) = pullback f R , is-closed-pullback f R c
pre .F-id = funext λ _ → Σ-prop-path! pullback-id
pre .F-∘ f g = funext λ _ → Σ-prop-path! pullback-∘ pre
It remains to show that this is a sheaf. We start by showing that it is separated. Suppose we have two closed sieves which agree everywhere on some We want to show so fix some we’ll show iff
.snd = from-is-separated sep mk where ΩJ
It appears that we don’t know much about how and behave outside of their agreement on but knowing that they’re closed will be enough to show that they agree everywhere. First, let’s codify that they actually agree on their intersection with
: is-separated J (ΩJ .fst)
sep {U} R {S , cS} {T , cT} α = ext λ {V} h →
sep let
: S ∩S ⟦ R ⟧ ≡ T ∩S ⟦ R ⟧
rem₁ = ext λ {V} h → Ω-ua
rem₁ (λ (h∈S , h∈R) → cT h (subst (J ∋_) (ap fst (α h h∈R)) (max (S .closed h∈S id))) , h∈R)
(λ (h∈T , h∈R) → cS h (subst (J ∋_) (ap fst (sym (α h h∈R))) (max (T .closed h∈T id))) , h∈R)
Then, assuming w.l.o.g. that we know that the pullback is a covering sieve. And since is a subsieve of we conclude that if then is and since is closed, this implies also that
: h ∈ S → J ∋ pullback h (S ∩S ⟦ R ⟧)
rem₂ = local (pull h (inc R)) λ f hf∈R → max
rem₂ h∈S ( S .closed h∈S (f ∘ id)
(_∈ R) (ap (h ∘_) (intror refl)) hf∈R
, subst )
: h ∈ S → J ∋ pullback h T
rem₂' = incl (λ _ → fst) (subst (J ∋_) (ap (pullback h) rem₁) (rem₂ h∈S)) rem₂' h∈S
We omit the symmetric converse for brevity.
: h ∈ T → J ∋ pullback h S
rem₃ = incl (λ _ → fst) (subst (J ∋_) (ap (pullback h) (sym rem₁))
rem₃ ht (local (pull h (inc R)) λ f rfh → max (T .closed ht (f ∘ id) , subst (_∈ R) (ap (h ∘_) (intror refl)) rfh)))
in Ω-ua (λ h∈S → cT h (rem₂' h∈S)) (λ h∈T → cS h (rem₃ h∈T))
Now we have to show that a family of compatible closed sieves over a sieve can be uniquely patched to a closed sieve on This is the sieve which is defined to contain whenever, for all in the part is the maximal sieve.
module _ {U : ⌞ C ⌟} (R : J ʻ U) (S : Patch (ΩJ .fst) ⟦ R ⟧) where
: Sieve C U
S' .arrows {V} g = elΩ $
S' ∀ {W} (f : Hom W V) (hf : f ∈ pullback g ⟦ R ⟧) →
∀ {V'} (i : Hom V' W) → i ∈ S .part (g ∘ f) hf .fst
.closed = elim! λ α h → inc λ {W} g hf →
S' (λ e → ∀ (h : e ∈ R) {V'} (i : Hom V' W) → i ∈ S .part e h .fst)
subst (assoc _ _ _) (α (h ∘ g)) hf
module _ {V W W'} (f : Hom V U) (hf : f ∈ ⟦ R ⟧) (g : Hom W V) (hfg : f ∘ g ∈ ⟦ R ⟧) {h : Hom W' W} where
: h ∈ S .part (f ∘ g) hfg .fst ≡ (g ∘ h) ∈ S .part f hf .fst
lemma = sym (ap (λ e → ⌞ e .fst .arrows h ⌟) (S .patch f hf g hfg))
lemma
module lemma = Equiv (path→equiv lemma)
The first thing we have to show is that this pulls back to This is, as usual, a proof of biimplication, though in this case both directions are painful — and entirely mechanical — calculations.
: ∀ {V} (f : Hom V U) (hf : f ∈ R) → pullback f S' ≡ S .part f hf .fst
restrict = ext λ {V} h → Ω-ua
restrict f hf (rec! λ α →
let
: id ∈ S .part (f ∘ h ∘ id) (⟦ R ⟧ .closed hf (h ∘ id)) .fst
step₁ = subst₂ (λ e e' → id ∈ S .part e e' .fst) (pullr refl) (to-pathp⁻ refl)
step₁ (α id _ id)
: ((h ∘ id) ∘ id) ∈ S .part f hf .fst
step₂ = lemma.to f hf (h ∘ id) (⟦ R ⟧ .closed hf (h ∘ id)) {id} step₁
step₂
in subst (λ e → ⌞ S .part f hf .fst .arrows e ⌟) (cancelr (idr _)) step₂)
(λ hh → inc λ {W} g hg {V'} i → S .part ((f ∘ h) ∘ g) hg .snd i (max
let
: i ∈ S .part (f ∘ h ∘ g) _ .fst
s1 = lemma.from f hf (h ∘ g) _
s1 (subst (_∈ S .part f hf .fst) (assoc _ _ _)
(S .part f hf .fst .closed hh (g ∘ i)))
: PathP (λ i → assoc f h g i ∈ R) _ hg
q = to-pathp⁻ refl
q in transport (λ j → ⌞ S .part (assoc f h g j) (q j) .fst .arrows (idr i (~ j)) ⌟) s1))
Finally, we can use this to show that is closed.
abstract
: is-closed S'
S'-closed {V} h hb = inc λ {W} f hf {V'} i → S .part (h ∘ f) hf .snd i $
S'-closed let
=
p (f ∘ i) (pullback h S') ≡˘⟨ pullback-∘ ⟩
pullback (h ∘ f ∘ i) S' ≡⟨ restrict (h ∘ f ∘ i) (subst (_∈ R) (sym (assoc h f i)) (⟦ R ⟧ .closed hf i)) ⟩
pullback .part (h ∘ f ∘ i) _ .fst ≡⟨ ap₂ (λ e e' → S .part e e' .fst) (assoc h f i) (to-pathp⁻ refl) ⟩
S .part ((h ∘ f) ∘ i) _ .fst ≡˘⟨ ap fst (S .patch (h ∘ f) hf i (⟦ R ⟧ .closed hf i)) ⟩
S (S .part (h ∘ f) hf .fst) ∎
pullback i in subst (J ∋_) p (pull (f ∘ i) hb)
: Section (ΩJ .fst) S
mk .whole = S' , S'-closed
mk .glues f hf = Σ-prop-path! (restrict f hf)
mk
open _=>_
: ΩJ .fst => Sieves {C = C}
ΩJ=>Ω .η U = fst
ΩJ=>Ω .is-natural x y f = refl ΩJ=>Ω
Closed subpresheaves🔗
To show that is the subobject classifier in we’ll start by showing that a subpresheaf of a sheaf is a sheaf precisely when the associated natural transformation into the presheaf of sieves on is valued in closed sieves.
First, suppose that the domain of is indeed a sheaf, and note that since the inclusion is a right adjoint, any subobject in is also a subobject in Since subobjects in are componentwise embeddings, the same is thus true of a subobject in
: {A : Sheaf J _} (A' : Subobject (Sheaves J _) A) → A .fst => ΩJ .fst
sheaf-name {A} A' = nm module sheaf-name where
sheaf-name : Subobject (PSh ℓ C) (A .fst)
sub' .domain = A' .domain .fst
sub' .map = A' .map
sub' .monic = right-adjoint→is-monic _
sub' (free-objects→left-adjoint (Small.make-free-sheaf J))
{A' .domain} {A} (λ {C} → A' .monic {C})
: ∀ {i} → is-embedding (A' .map .η i)
emb = is-monic→is-embedding-at ℓ C (sub' .monic)
emb
: A .fst => Sieves {C = C}
nm' = psh-name sub' nm'
We can then compute, in the name of resulting in a natural transformation It remains to show that this transformation is actually valued in i.e. that each of the resulting sieves is closed. To this end, suppose we have some and that the sieve is We must show that contains
By the definition of it suffices to exhibit a fibre of over and since is a sheaf, we can construct this fibre by gluing over the sieve At each we are given some fibre of over but since is an embedding everywhere, we can extract an actual element of from this mere fibre. By the usual functoriality work, this extends to a patch over
: {U : ⌞ C ⌟} (x : A ʻ U) → is-closed (nm' .η U x)
name-is-closed {U} x {V} h pb =
name-is-closed let
= sat.split (A' .domain .snd) pb record
it { part = λ {W} f hf → □-out (emb _) hf .fst
; patch = λ f hf g hgf → ap fst $ emb _
(_ , (A' .map .η _ (A' .domain .fst ⟪ g ⟫ (□-out (emb _) hf .fst)) ≡⟨ A' .map .is-natural _ _ _ $ₚ _ ⟩
(A .fst ⟪ g ⟫ (A' .map .η _ (□-out (emb _) hf .fst))) ≡⟨ ap (A .fst .F₁ g) (□-out (emb _) hf .snd) ⟩
(A .fst ⟪ g ⟫ (A .fst ⟪ h ∘ f ⟫ x)) ≡⟨ sym (A .fst .F-∘ _ _ $ₚ _) ⟩
(A .fst ⟪ (h ∘ f) ∘ g ⟫ x) ∎))
(_ , □-out (emb _) hgf .snd ∙ ap₂ (A .fst .F₁) (assoc _ _ _) refl)
}
We must then show that this is indeed a fibre of over But since being a sheaf, is separated, it suffices to do so locally along at which point we know that this is true by construction.
= sat.separate (A .snd) pb λ f hf →
prf .fst ⟪ f ⟫ A' .map .η V (it .whole) ≡˘⟨ A' .map .is-natural _ _ _ $ₚ _ ⟩
A .map .η _ (A' .domain .fst ⟪ f ⟫ it .whole) ≡⟨ ap (A' .map .η _) (it .glues f hf) ⟩
A' .map .η _ (□-out (emb _) hf .fst) ≡⟨ □-out (emb _) hf .snd ⟩
A' .fst ⟪ h ∘ f ⟫ x ≡⟨ A .fst .F-∘ _ _ $ₚ _ ⟩
A .fst ⟪ f ⟫ (A .fst ⟪ h ⟫ x) ∎
A in inc (it .whole , prf)
: A .fst => ΩJ .fst
nm .η U x = record { fst = nm' .η U x ; snd = name-is-closed x }
nm .is-natural x y f = funext λ a → Σ-prop-path! (happly (nm' .is-natural x y f) a) nm
This “if” direction turns out to be sufficient to prove that is a subobject classifier. First, the maximal sieve is obviously closed, because it contains every morphism:
: Subobject (Sheaves J ℓ) ΩJ
Sh[]-true .domain = Terminal.top (Sh[]-terminal J)
Sh[]-true .map .η _ _ = maximal' , (λ _ _ → tt)
Sh[]-true .map .is-natural x y f = trivial!
Sh[]-true .monic g h x = trivial! Sh[]-true
Now the universal property follows essentially by arguing in Since the proofs boil down to shuffling whether a given square is a pullback in or in we do not comment on them further.
: is-generic-subobject (Sheaves J ℓ) Sh[]-true
Sh[]-true-is-generic .name = sheaf-name
Sh[]-true-is-generic .classifies m = record { has-is-pb = pb' } where
Sh[]-true-is-generic : is-pullback-along (PSh ℓ C) (m .map) (ΩJ=>Ω ∘nt sheaf-name m) (ΩJ=>Ω ∘nt Sh[]-true .map)
rem = record { has-is-pb = subst-is-pullback refl trivial! refl trivial!
rem (ΩPSh.classifies (sheaf-name.sub' m) .has-is-pb) }
: is-pullback (Sheaves J ℓ) (m .map) (sheaf-name m) (rem .top) (Sh[]-true .map)
pb' .square = ext λ i h {V} x → unext (rem .square) i h x
pb' .universal α = rem .universal (PSh.extendr α)
pb' .p₁∘universal = rem .p₁∘universal
pb' .p₂∘universal = rem .p₂∘universal
pb' .unique = rem .unique
pb'
.unique {m = m} {nm} α = ext λ i x {V} h → unext path i x h where
Sh[]-true-is-generic : is-pullback (PSh ℓ C) (m .map) nm (α .top) (Sh[]-true .map)
pb = right-adjoint→is-pullback (free-objects→left-adjoint (Small.make-free-sheaf J)) (α .has-is-pb)
pb
: is-pullback (PSh ℓ C) (m .map) (ΩJ=>Ω ∘nt nm) (α .top) (ΩPSh.true .map)
pb' .square = ext λ i h {V} x → unext (pb .square) i h x
pb' .universal {p₁' = p₁'} {p₂'} α = pb .universal {p₁' = p₁'} {p₂'} (ext λ i x {V} h → unext α i x h)
pb' .p₁∘universal = pb .p₁∘universal
pb' .p₂∘universal = pb .p₂∘universal
pb' .unique = pb .unique
pb'
= ΩPSh.unique {m = sheaf-name.sub' m} record { has-is-pb = pb' } path
: Subobject-classifier (Sheaves J ℓ)
Sh[]-omega = record { generic = Sh[]-true-is-generic } Sh[]-omega
Finally, we can show the converse direction to the result we claimed in the start of this section: a subpresheaf with name valued in sieves is itself a sheaf. First, we show that it is separated, because it embeds into a sheaf:
name-is-closed→is-sheaf: {A : Sheaf J _} (A' : Subobject (PSh ℓ C) (A .fst))
→ (∀ {U : ⌞ C ⌟} (x : A ʻ U) → is-closed (psh-name A' .η U x))
→ is-sheaf J (A' .domain)
{A = A} A' cl = from-is-sheaf₁ λ {U} j p → from-is-separated₁ _ (sep j) (s j p) where
name-is-closed→is-sheaf : ∀ {i} → is-embedding (A' .map .η i)
emb = is-monic→is-embedding-at ℓ C (A' .monic)
emb
: ∀ {U} (j : J .covers U) → is-separated₁ (A' .domain) (J .cover j)
sep {x} {y} h = ap fst $ emb _ (_ , refl) $ _ , A .snd .separate j λ f hf →
sep j .fst ⟪ f ⟫ A' .map .η _ y ≡˘⟨ A' .map .is-natural _ _ _ $ₚ _ ⟩
A .map .η _ ⌜ A' .domain ⟪ f ⟫ y ⌝ ≡⟨ ap! (sym (h f hf)) ⟩
A' .map .η _ (A' .domain ⟪ f ⟫ x) ≡⟨ A' .map .is-natural _ _ _ $ₚ _ ⟩
A' .fst ⟪ f ⟫ A' .map .η _ x ∎ A
Now, given a sieve of some object and a patch along over we must put together a section in First, we lift to a patch of the sheaf, over the same sieve; and this glues to a section
module _ {U} (j : J .covers U) (p : Patch (A' .domain) (J .cover j)) where
: Patch (A .fst) (J .cover j)
p' = record
p' { part = λ {V} f hf → A' .map .η V (p .part f hf)
; patch = λ {V} {W} f hf g hgf → sym (A' .map .is-natural _ _ _ $ₚ _) ∙ ap (A' .map .η _) (p .patch f hf g hgf)
}
: Section (A .fst) p'
s' = is-sheaf.split (A .snd) p' s'
We must show that actually comes from a section This follows from having a closed name: by definition, this means that if we can show that is a sieve, then the fibre of over is contractible. To show this, it suffices to show that pulls back to a locally at every in some other sieve — we choose the we were already working against. Then we’re left with a simple computation:
abstract
: J ∋ pullback id (psh-name A' .η U (s' .whole))
has = local (inc j) λ {V} f hf →
has let
: pullback f (pullback id (psh-name A' .η U (s' .whole))) ≡ psh-name A' .η V (p' .part f hf)
α =
α (psh-name A' .η U (s' .whole)) ⌝ ≡⟨ ap! pullback-id ⟩
pullback f ⌜ pullback id (psh-name A' .η U (s' .whole)) ≡⟨ sym (psh-name A' .is-natural _ _ _ $ₚ _) ⟩
pullback f .η V (A .fst ⟪ f ⟫ s' .whole) ≡⟨ ap (psh-name A' .η _) (s' .glues f hf) ⟩
psh-name A' .η V (p' .part f hf) ∎
psh-name A' in subst (J ∋_) (sym α) (max (inc (p .part f hf , sym (A .fst .F-id $ₚ _))))
: fibre (map A' .η U) (A .fst ⟪ id ⟫ s' .whole)
it = □-out (emb _) (cl (s' .whole) id has)
it
: Section (A' .domain) p
s = record
s { whole = it .fst
; glues = λ f hf → ap fst (emb (A .fst ⟪ f ⟫ s' .whole)
(_ , A' .map .is-natural _ _ _ $ₚ _ ∙ ap (A .fst .F₁ f) (it .snd ∙ A .fst .F-id $ₚ _))
(_ , sym (s' .glues f hf)))
}