module Cat.Site.Instances.Atomic where
Atomic sites🔗
The atomic topology on a small category is the coverage where a sieve covers an object as soon as is inhabited. If is a completely arbitrary category, then the pullback of an inhabited sieve need not be inhabited; a necessary condition is that every pair of maps and can be completed to a commutative square
in which and may be completely arbitrary. Following Johnstone (2002, A2.1.11(h)), we refer to this condition on as the right Ore condition.
module _ {o ℓ} (C : Precategory o ℓ) where
open Precategory C
: ∀ {a b c} (f : Hom a c) (g : Hom b c) → Type _
Squaring {a = a} {b} {c} f g = Σ[ d ∈ C ]
Squaring (f ∘ p1 ≡ g ∘ p2)
Σ[ p1 ∈ Hom d a ] Σ[ p2 ∈ Hom d b ]
: Type _
Right-ore = ∀ {a b c} (f : Hom a c) (g : Hom b c) → ∥ Squaring f g ∥ Right-ore
module _ {o ℓ} (C : Precategory o ℓ) (square : Right-ore C) where
open Cat C
Note that if this condition is satisfied, the atomic coverage is automatically a Grothendieck coverage.
: Topology C (o ⊔ ℓ)
Atomic-topology .covering U R = ∃[ V ∈ C ] Σ[ f ∈ Hom V U ] f ∈ R
Atomic-topology .has-is-prop = hlevel 1
Atomic-topology .stable {U} {V} f {S} = rec! λ W g hg → do
Atomic-topology (d , p1 , p2 , q) ← square f g
(d , p1 , subst (_∈ S) (sym q) (S .closed hg p2))
pure .maximal r = inc (_ , _ , r)
Atomic-topology .local = elim! λ V h pt {S} α → case α h pt of λ
Atomic-topology → inc (_ , _ , β) U g β
module Atomic = Topology Atomic-topology
open Atomic using ()
renaming (from-topology to Atomic-coverage) public
Sheaves for the atomic topology🔗
We now want to characterise when a functor is a sheaf for the atomic topology. A morphism is said to support a section if maps act identically on as soon as they are equal under precomposition with
module _ {ℓ'} (A : Functor (C ^op) (Sets ℓ')) where
: ∀ {V U} (f : Hom V U) (y : A ʻ V) → Type _
is-support {V} f y =
is-support ∀ {W} (g h : Hom W V)
→ f ∘ g ≡ f ∘ h
→ A ⟪ g ⟫ y ≡ A ⟪ h ⟫ y
module _ {ℓ'} (A : Functor (C ^op) (Sets ℓ')) where
private module A = PSh A using (collapse)
The importance of this notion is that, if a sieve is inhabited by some then a patch of over is precisely an element supported by
patch→support: ∀ {U V} {S : Sieve C U} {f : Hom V U} (hf : f ∈ S) (p : Patch A S)
→ is-support A f (p .part f hf)
{S = S} {f = f} hf p g h β =
patch→support .part f hf ≡⟨ p .patch f hf g (S .closed hf _) ⟩
A ⟪ g ⟫ p .part (f ∘ g) _ ≡⟨ app p β ⟩
p .part (f ∘ h) _ ≡˘⟨ p .patch f hf h (S .closed hf _) ⟩
p .part f hf ∎ A ⟪ h ⟫ p
In the other direction, we prefer to say that any supported by can be extended to a patch for the sieve generated by
support→patch: ∀ {U V} {f : Hom V U} {x : A ʻ V}
→ is-support A f x
→ Patch {C = C} A (map→sieve f)
{f = f} {x} supp .part g hg =
support→patch
∥-∥-rec-set!(λ (h , p) → A ⟪ h ⟫ x)
(λ (h , α) (h' , β) → supp h h' (α ∙ sym β))
(□-tr hg)
Note that, since we have to produce an inhabitant of given only a truncated factorization of through we must use the support property to show that we always get the same inhabitant.
{f = f} {x} supp .patch = elim! λ g h p i j sq →
support→patch (A ⟪ h ⟫ x) ≡⟨ A.collapse refl ⟩
A ⟪ i ⟫ (h ∘ i) j (pulll p ∙ sym sq) ⟩
A ⟪ h ∘ i ⟫ x ≡⟨ supp A ⟪ j ⟫ x ∎
With this correspondence, we can define a simpler notion of sheaf for the atomic topology: is a sheaf if every supported by admits a unique lifting with
: ∀ {ℓ'} (A : Functor (C ^op) (Sets ℓ')) → Type _
is-atomic-sheaf =
is-atomic-sheaf A ∀ {U V : ⌞ C ⌟} (y : A ʻ V) (f : Hom V U)
→ is-support A f y
→ is-contr (Σ[ x ∈ A ʻ U ] (y ≡ A ⟪ f ⟫ x))
The easy direction is showing that this is implied by the sheaf condition. Since the sieve generated by any is covering in the atomic topology, and we’ve already shown that supported elements correspond to patches, all we have to do is a bit of algebra.
to-is-atomic-sheaf: ∀ {ℓ'} (P : Functor (C ^op) (Sets ℓ'))
→ is-sheaf Atomic-coverage P
→ is-atomic-sheaf P
{U} {V} y f loc = done where
to-is-atomic-sheaf A sheaf module A = PSh A
: Atomic-coverage .covers U
cov = map→sieve f , inc (V , f , inc (id , elimr refl))
cov
: Patch A (map→sieve f)
x = support→patch A loc
x
: y ≡ A ⟪ f ⟫ sheaf .whole cov x
rem₁ = sym (sheaf .glues cov x f (inc (id , elimr refl)) ∙ A.elim refl)
rem₁
: is-contr _
done .centre = record { snd = rem₁ }
done .paths (z , q) = Σ-prop-path! $ sheaf .separate cov $ elim! λ g h r →
done .whole cov x ≡⟨ A.expand (sym r) ⟩
A ⟪ g ⟫ sheaf (A ⟪ f ⟫ sheaf .whole cov x) ≡⟨ A.ap (sym rem₁) ⟩
A ⟪ h ⟫ .ap q ⟩
A ⟪ h ⟫ y ≡⟨ A(A ⟪ f ⟫ z) ≡⟨ A.collapse r ⟩
A ⟪ h ⟫
A ⟪ g ⟫ z ∎
from-is-atomic-sheaf: ∀ {ℓ'} (A : Functor (C ^op) (Sets ℓ'))
→ is-atomic-sheaf A → is-sheaf Atomic-coverage A
= from-is-sheaf₁ λ {U} → elim! done where
from-is-atomic-sheaf A at module A = PSh A
The converse direction requires slightly more work. We assume that we have a sieve covering inhabited by some and a patch of Since we can extract an element from by the “atomic sheaf” condition we obtain that there is a contractible space of lifts of through
module _ {U} (S : Sieve C U) V (f₀ : Hom V U) (hf₀ : f₀ ∈ S) (p : Patch A S) where
: A ʻ V
p₀ = p .part f₀ hf₀
p₀
: is-contr (Σ[ x ∈ A ʻ U ] (p .part f₀ hf₀ ≡ A ⟪ f₀ ⟫ x))
sec = at p₀ f₀ (patch→support A hf₀ p)
sec
open Σ (sec .centre) renaming (fst to e ; snd to q)
We must then show that agrees with at any other To this end, suppose that we have such an and fit it into a commutative square
using that satisfies the right Ore condition.
rem₁': ∀ {W} (f : Hom W U) (hf : f ∈ S)
→ Squaring C f₀ f
→ A ⟪ f ⟫ e ≡ p .part f hf
{W = W} f hf (X , v , u , sq) = done where rem₁'
First, a short calculation shows that is supported by which implies that there is exactly one element in on which acts to give
: is-support A u (A ⟪ v ⟫ p₀)
cohy =
cohy g h q (A ⟪ v ⟫ p₀) ≡⟨ A.collapse refl ⟩
A ⟪ g ⟫ _ _ (pulll sq ·· pullr q ·· extendl (sym sq)) ⟩
A ⟪ v ∘ g ⟫ p₀ ≡⟨ patch→support A hf₀ p .expand refl ⟩
A ⟪ v ∘ h ⟫ p₀ ≡⟨ A(A ⟪ v ⟫ p₀) ∎
A ⟪ h ⟫
: is-contr (Σ[ x ∈ A ʻ W ] A ⟪ v ⟫ p₀ ≡ A ⟪ u ⟫ x)
y = at (A ⟪ v ⟫ p₀) u cohy y
But since comes from a patch of and we can also calculate that is which is by commutativity of the square above, and this in turn is
: A ⟪ v ⟫ p₀ ≡ A ⟪ u ⟫ p .part f hf
rem₂ =
rem₂ .patch f₀ hf₀ v (S .closed hf₀ v) ⟩
A ⟪ v ⟫ p₀ ≡⟨ p .part (f₀ ∘ v) _ ≡⟨ ap₂ (p .part) sq prop! ⟩
p .part (f ∘ u) _ ≡˘⟨ p .patch f hf u (S .closed hf u) ⟩
p .part f hf ∎ A ⟪ u ⟫ p
On the other hand, we can also calculate that by the definition of and that so we must have Therefore is a section of and we are essentially done.
=
rem₃ .ap q ⟩
A ⟪ v ⟫ p₀ ≡⟨ A(A ⟪ f₀ ⟫ e) ≡⟨ A.weave sq ⟩
A ⟪ v ⟫ (A ⟪ f ⟫ e) ∎
A ⟪ u ⟫
= ap fst (is-contr→is-prop y (_ , rem₃) (_ , rem₂))
done
: ∀ {W} (f : Hom W U) (hf : f ∈ S) → A ⟪ f ⟫ e ≡ p .part f hf
rem₁ = ∥-∥-rec (hlevel 1) (rem₁' f hf) (square f₀ f)
rem₁ f hf
: is-contr (Section A p)
done .centre = record { glues = λ g hg → rem₁ g hg }
done .paths x = ext $ ap fst $ is-contr→is-prop sec
done (_ , q) (_ , sym (x .glues f₀ hf₀))
The most immediate consequence of this simplification is that, in the atomic topology, every constant functor is a sheaf.
: ∀ {ℓ'} (X : Set ℓ') → is-sheaf Atomic-coverage (Const X)
Const-is-sheaf = from-is-atomic-sheaf _ (λ _ _ _ → hlevel 0) Const-is-sheaf X
Atomic topoi🔗
A characteristic property of topoi of sheaves on atomic sites is that the subobject classifier sheaf is given by the constant object on the usual set of propositions. Because we know that constant presheaves are sheaves in an atomic topos, we do not need to sheafify to obtain this latter sheaf, which simplifies the calculation significantly.
module _ {ℓ} {C : Precategory ℓ ℓ} (sq : Right-ore C) where
open Cat.Site.Closure
open Cat C
open _=>_
private
= Atomic-coverage C sq
cov
module Sh = Cat (Sheaves cov ℓ)
: Sheaf cov ℓ
ΩJ' = record
ΩJ' { fst = Const (el! (Lift ℓ Ω))
; snd = Const-is-sheaf C sq _
}
It’s easy to construct this isomorphism by hand. In one direction, we send a closed sieve to the proposition (which is equivalent to In the other direction, we send a proposition to the sieve which contains any iff
private interleaved mutual
: ΩJ cov .fst => ΩJ' .fst
m1 .η U (S , _) = lift (S .arrows id)
m1
: ΩJ' .fst => ΩJ cov .fst
m2 .η x P .fst .arrows h = P .lower m2
.is-natural x y f = ext λ S cl → biimp
m1 (λ hf → cl id (inc (pullback id S , inc (y , f , subst (_∈ S) id-comm hf))))
(λ hid → subst (_∈ S) id-comm-sym (S .closed hid f))
.η x P .fst .closed hf _ = hf
m2 .η x P .snd h y = case Atomic.unsaturate C sq y of λ _ _ α → α
m2 .is-natural x y f = ext λ P → Σ-prop-path! trivial! m2
Finally, we can show that these maps are inverses. Note that one direction is definitional, and the other is not much more complicated.
: ΩJ cov Sh.≅ ΩJ'
ΩJ-is-constant =
ΩJ-is-constant let
= ext λ i X cl → Σ-prop-path! $ ext λ x → biimp
q (λ p → subst (_∈ X) (idl _) (X .closed p _))
(λ p → cl id (inc (_ , inc (_ , _ , subst (_∈ X) id-comm (X .closed p id)))))
in Sh.make-iso m1 m2 trivial! q
References
- Johnstone, Peter T. 2002. Sketches of an Elephant: a Topos Theory Compendium. Oxford Logic Guides. New York, NY: Oxford Univ. Press. https://cds.cern.ch/record/592033.