Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

Sites, sheaves and sheafification as a QIT #1031

Merged
merged 59 commits into from
Nov 21, 2023
Merged
Show file tree
Hide file tree
Changes from all commits
Commits
Show all changes
59 commits
Select commit Hold shift + click to select a range
aa333d9
stubs for Coverage, Cover, Sieve
MatthiasHu Apr 19, 2023
14c9cb3
covers as families, sieves
MatthiasHu Apr 19, 2023
c37ac8a
pulledBackSieve
MatthiasHu Apr 20, 2023
ba9f430
finish Coverage
MatthiasHu Apr 20, 2023
44dad75
slightly simplify generatedSieve
MatthiasHu May 19, 2023
32c144a
universal property of generatedSieve
MatthiasHu May 19, 2023
960feee
make C implicit in sieve operations
MatthiasHu May 20, 2023
0c2770f
make C implicit in patchArr and patchObj
MatthiasHu May 20, 2023
fe4438e
definition of sheaves on a site
MatthiasHu May 20, 2023
e0f1ad8
tiny improvement (qualified import)
MatthiasHu May 28, 2023
9b86341
Merge branch 'master' into attempt-at-sheafification
MatthiasHu Jul 27, 2023
76dfda4
remove unnecessary parenthesis
MatthiasHu Jul 28, 2023
bdff467
indentation in equational reasoning
MatthiasHu Jul 28, 2023
49d1d06
sheafification QIT
MatthiasHu Jul 31, 2023
d31a55f
correct silly mistake
MatthiasHu Aug 1, 2023
e9008c3
rename CompatibleFamilies -> CompatibleFamily
MatthiasHu Aug 1, 2023
343170f
sheafification is a sheaf
MatthiasHu Aug 1, 2023
2d7666f
rename
MatthiasHu Aug 1, 2023
1811b11
stub of the universal property
MatthiasHu Aug 2, 2023
fbd1c26
small renaming
MatthiasHu Aug 2, 2023
363f17a
(wip)
MatthiasHu Aug 2, 2023
2b84caa
(wip) (termination checker problems)
MatthiasHu Aug 3, 2023
eb4937f
(wip) solved termination issue
MatthiasHu Aug 3, 2023
dc55c39
finish inducedMap of sheafification
MatthiasHu Aug 8, 2023
9e0e274
comment with reference
MatthiasHu Aug 8, 2023
12017ec
rename i -> patch
MatthiasHu Aug 9, 2023
4ab79ee
diagram comment and a bit of preparation
MatthiasHu Aug 9, 2023
9f5bbbc
elimProp with restrictPreservesB
MatthiasHu Aug 9, 2023
1bb1887
introduce isSeparated
MatthiasHu Aug 10, 2023
cd7469d
uniqueness part of universal property
MatthiasHu Aug 11, 2023
4ab1418
package up as universal element
MatthiasHu Aug 11, 2023
1694fdb
name constructor arguments
MatthiasHu Aug 15, 2023
fa75467
elimProp without restrictPreservesB assumption
MatthiasHu Aug 16, 2023
0cbcbda
prepare for merging module
MatthiasHu Aug 16, 2023
d55d39a
merge module
MatthiasHu Aug 16, 2023
70a8dc8
properly name elimProp assumptions
MatthiasHu Aug 16, 2023
1be4618
fix misnomer (now we have (η ⟦ _ ⟧) x = η⟦⟧ x)
MatthiasHu Aug 16, 2023
87bf43a
structure elimPropInduction more nicely
MatthiasHu Aug 16, 2023
26a8551
fix indentation
MatthiasHu Aug 16, 2023
48cb7a1
use improved elimProp in uniqueness
MatthiasHu Aug 21, 2023
7ece10f
split sheafification in several files
MatthiasHu Aug 21, 2023
833859b
clean up imports
MatthiasHu Aug 21, 2023
3e76720
eliminate Families by inlining
MatthiasHu Aug 21, 2023
e74cede
comments in ElimProp
MatthiasHu Aug 22, 2023
9e6634a
references for coverage
MatthiasHu Aug 22, 2023
b8d2c8d
make two arguments implicit
MatthiasHu Aug 22, 2023
af0e3d7
a few more comments
MatthiasHu Aug 22, 2023
d53c9fb
Merge branch 'master' into sheafification-QIT
MatthiasHu Aug 22, 2023
258e8b1
fix unresolved meta (with Agda 2.6.3)
MatthiasHu Aug 22, 2023
8f0c3e1
Merge branch 'master' into sheafification-QIT
MatthiasHu Aug 28, 2023
06a6fb9
make helpers private
MatthiasHu Aug 28, 2023
551387c
let isSheaf be an hProp
MatthiasHu Aug 31, 2023
b41e529
fix naming typo
MatthiasHu Aug 31, 2023
2eef94d
un-bundle `hProp`s in `Sheaf.agda`
MatthiasHu Nov 16, 2023
fa6a3a3
un-bundle `hProp`s in `Sieve.agda`
MatthiasHu Nov 17, 2023
24a7c5c
rename `F` -> `sheafification`
MatthiasHu Nov 18, 2023
74b0ae6
also rename the HIT `⟨F⟅_⟆⟩` and the local `F`
MatthiasHu Nov 18, 2023
6a8f149
elaborate universe level comment
MatthiasHu Nov 18, 2023
0f521b3
Merge branch 'master' into sheafification-QIT
MatthiasHu Nov 20, 2023
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
44 changes: 44 additions & 0 deletions Cubical/Categories/Site/Cover.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,44 @@
{-# OPTIONS --safe #-}
module Cubical.Categories.Site.Cover where

-- A cover of an object is just a family of arrows into that object.
-- This notion is used in the definition of a coverage.

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Structure
open import Cubical.HITs.PropositionalTruncation

open import Cubical.Categories.Category
open import Cubical.Categories.Constructions.Slice


module _
{ℓ ℓ' : Level}
(C : Category ℓ ℓ')
where

open Category

Cover : (ℓpat : Level) → ob C → Type (ℓ-max (ℓ-max ℓ ℓ') (ℓ-suc ℓpat))
Cover ℓpat c = TypeWithStr ℓpat λ patches → patches → ob (SliceCat C c)

module _
{ℓ ℓ' : Level}
{C : Category ℓ ℓ'}
where

open Category

-- Extracting the members (patches) from a cover.
module _
{ℓpat : Level}
{c : ob C}
(cov : Cover C ℓpat c)
(patch : ⟨ cov ⟩)
where

patchObj : ob C
patchObj = S-ob (str cov patch)

patchArr : C [ patchObj , c ]
patchArr = S-arr (str cov patch)
43 changes: 43 additions & 0 deletions Cubical/Categories/Site/Coverage.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,43 @@
{-# OPTIONS --safe #-}
module Cubical.Categories.Site.Coverage where

-- Definition of a coverage on a category, also called a Grothendieck topology.
-- A coverage on a category turns it into a site
-- and enables us to formulate a sheaf condition.

-- We stay close to the definitions given in the nLab and the Elephant:
-- * https://ncatlab.org/nlab/show/coverage
-- * Peter Johnstone, "Sketches of an Elephant" (Definition C2.1.1)

-- While the covers are just families of arrows,
-- we use the notion of sieves to express the "pullback stability" property of the coverage.

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Structure
open import Cubical.HITs.PropositionalTruncation
open import Cubical.Data.Sigma

open import Cubical.Categories.Category
open import Cubical.Categories.Site.Cover
open import Cubical.Categories.Site.Sieve

module _
{ℓ ℓ' : Level}
(C : Category ℓ ℓ')
where

open Category C

record Coverage (ℓcov ℓpat : Level) : Type (ℓ-max ℓ (ℓ-max ℓ' (ℓ-suc (ℓ-max ℓcov ℓpat)))) where
no-eta-equality
field
covers : (c : ob) → TypeWithStr ℓcov λ Cov → Cov → (Cover C ℓpat c)
pullbackStability :
{c : ob} →
(cov : ⟨ covers c ⟩) →
{d : ob} →
(f : Hom[ d , c ]) →
∃[ cov' ∈ ⟨ covers d ⟩ ]
coverRefinesSieve
(str (covers d) cov')
(pulledBackSieve f (generatedSieve (str (covers c) cov)))
138 changes: 138 additions & 0 deletions Cubical/Categories/Site/Sheaf.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,138 @@
{-# OPTIONS --safe #-}
module Cubical.Categories.Site.Sheaf where

-- Definition of sheaves on a site.

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Structure
open import Cubical.Foundations.HLevels
open import Cubical.Foundations.Function using (_∘_)
open import Cubical.Foundations.Equiv

open import Cubical.Data.Sigma

open import Cubical.Functions.Embedding

open import Cubical.Categories.Category
open import Cubical.Categories.Site.Cover
open import Cubical.Categories.Site.Sieve
open import Cubical.Categories.Site.Coverage
open import Cubical.Categories.Presheaf
open import Cubical.Categories.Functor
open import Cubical.Categories.Constructions.FullSubcategory

module _
{ℓ ℓ' : Level}
{C : Category ℓ ℓ'}
{ℓP : Level}
(P : Presheaf C ℓP)
where

open Category C hiding (_∘_)

module _
{c : ob}
{ℓ'' : Level}
(cov : Cover C ℓ'' c)
where

FamilyOnCover : Type (ℓ-max ℓP ℓ'')
FamilyOnCover = (i : ⟨ cov ⟩) → ⟨ P ⟅ patchObj cov i ⟆ ⟩

isCompatibleFamily : FamilyOnCover → Type (ℓ-max (ℓ-max (ℓ-max ℓ ℓ') ℓP) ℓ'')
isCompatibleFamily fam =
(i : ⟨ cov ⟩) →
(j : ⟨ cov ⟩) →
(d : ob) →
(f : Hom[ d , patchObj cov i ]) →
(g : Hom[ d , patchObj cov j ]) →
f ⋆ patchArr cov i ≡ g ⋆ patchArr cov j →
(P ⟪ f ⟫) (fam i) ≡ (P ⟪ g ⟫) (fam j)

isPropIsCompatibleFamily : (fam : FamilyOnCover) → isProp (isCompatibleFamily fam)
isPropIsCompatibleFamily fam =
isPropΠ6 λ _ _ d _ _ _ → str (P ⟅ d ⟆) _ _

CompatibleFamily : Type (ℓ-max (ℓ-max (ℓ-max ℓ ℓ') ℓP) ℓ'')
CompatibleFamily = Σ FamilyOnCover isCompatibleFamily

isSetCompatibleFamily : isSet CompatibleFamily
isSetCompatibleFamily =
isSetΣSndProp
(isSetΠ (λ i → str (P ⟅ patchObj cov i ⟆)))
isPropIsCompatibleFamily

elementToCompatibleFamily : ⟨ P ⟅ c ⟆ ⟩ → CompatibleFamily
elementToCompatibleFamily x =
(λ i → (P ⟪ patchArr cov i ⟫) x) ,
λ i j d f g eq → cong (λ f → f x) (
P ⟪ f ⟫ ∘ P ⟪ patchArr cov i ⟫ ≡⟨ sym (F-seq (patchArr cov i) f ) ⟩
P ⟪ f ⋆ patchArr cov i ⟫ ≡⟨ cong (P ⟪_⟫) eq ⟩
P ⟪ g ⋆ patchArr cov j ⟫ ≡⟨ F-seq (patchArr cov j) g ⟩
P ⟪ g ⟫ ∘ P ⟪ patchArr cov j ⟫ ∎ )
where
open Functor P

hasAmalgamationPropertyForCover : Type (ℓ-max (ℓ-max (ℓ-max ℓ ℓ') ℓP) ℓ'')
hasAmalgamationPropertyForCover =
isEquiv elementToCompatibleFamily

isPropHasAmalgamationPropertyForCover : isProp hasAmalgamationPropertyForCover
isPropHasAmalgamationPropertyForCover =
isPropIsEquiv _

module _
{ℓ ℓ' ℓcov ℓpat : Level}
{C : Category ℓ ℓ'}
(J : Coverage C ℓcov ℓpat)
{ℓP : Level}
(P : Presheaf C ℓP)
where

open Coverage J

isSeparated : Type (ℓ-max (ℓ-max (ℓ-max ℓ ℓcov) ℓpat) ℓP)
isSeparated =
(c : _) →
(cov : ⟨ covers c ⟩) →
(x : ⟨ P ⟅ c ⟆ ⟩) →
(y : ⟨ P ⟅ c ⟆ ⟩) →
( (patch : _) →
let f = patchArr (str (covers c) cov) patch
in (P ⟪ f ⟫) x ≡ (P ⟪ f ⟫) y ) →
x ≡ y

isPropIsSeparated : isProp isSeparated
isPropIsSeparated = isPropΠ5 (λ c _ _ _ _ → str (P ⟅ c ⟆) _ _)

isSheaf : Type (ℓ-max (ℓ-max (ℓ-max (ℓ-max ℓ ℓ') ℓcov) ℓpat) ℓP)
isSheaf =
(c : _) →
(cov : ⟨ covers c ⟩) →
hasAmalgamationPropertyForCover P (str (covers c) cov)

isPropIsSheaf : isProp isSheaf
isPropIsSheaf = isPropΠ2 λ c cov → isPropHasAmalgamationPropertyForCover P (str (covers c) cov)

isSheaf→isSeparated : isSheaf → isSeparated
isSheaf→isSeparated isSheafP c cov x y locallyEqual =
isEmbedding→Inj (isEquiv→isEmbedding (isSheafP c cov)) x y
(Σ≡Prop
(isPropIsCompatibleFamily {C = C} P _)
(funExt locallyEqual))

module _
{ℓ ℓ' ℓcov ℓpat : Level}
{C : Category ℓ ℓ'}
(J : Coverage C ℓcov ℓpat)
(ℓF : Level)
where

Sheaf : Type (ℓ-max (ℓ-max (ℓ-max (ℓ-max ℓ ℓ') ℓcov) ℓpat) (ℓ-suc ℓF))
Sheaf = Σ[ P ∈ Presheaf C ℓF ] isSheaf J P

SheafCategory :
Category
(ℓ-max (ℓ-max (ℓ-max (ℓ-max ℓ ℓ') ℓcov) ℓpat) (ℓ-suc ℓF))
(ℓ-max (ℓ-max ℓ ℓ') ℓF)
SheafCategory = FullSubcategory (PresheafCategory C ℓF) (isSheaf J)
6 changes: 6 additions & 0 deletions Cubical/Categories/Site/Sheafification.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,6 @@
{-# OPTIONS --safe #-}
module Cubical.Categories.Site.Sheafification where

open import Cubical.Categories.Site.Sheafification.Base public
open import Cubical.Categories.Site.Sheafification.ElimProp public
open import Cubical.Categories.Site.Sheafification.UniversalProperty public
126 changes: 126 additions & 0 deletions Cubical/Categories/Site/Sheafification/Base.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,126 @@
{-# OPTIONS --safe #-}
module Cubical.Categories.Site.Sheafification.Base where

-- Construction of the sheafification of a presheaf on a site
-- using a quotient inductive type (QIT).

-- This is inspired by the construction of the sheafification (for finite coverings) in:
-- * E. Palmgren, S.J. Vickers, "Partial Horn logic and cartesian categories"

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Structure
open import Cubical.Foundations.Function using (_∘_)

open import Cubical.HITs.PropositionalTruncation using (∣_∣₁)

open import Cubical.Data.Sigma

open import Cubical.Functions.Surjection
open import Cubical.Functions.Embedding

open import Cubical.Categories.Category
open import Cubical.Categories.Presheaf
open import Cubical.Categories.Functor

open import Cubical.Categories.Site.Cover
open import Cubical.Categories.Site.Coverage
open import Cubical.Categories.Site.Sheaf


module Sheafification
{ℓ ℓ' ℓcov ℓpat : Level}
{C : Category ℓ ℓ'}
(J : Coverage C ℓcov ℓpat)
{ℓP : Level}
(P : Presheaf C ℓP)
where

open Category C hiding (_∘_)
open Coverage J

data ⟨sheafification⟅_⟆⟩ : ob → Type (ℓ-max ℓ (ℓ-max ℓ' (ℓ-max ℓcov (ℓ-max ℓpat ℓP)))) where

trunc : {c : ob} → isSet ⟨sheafification⟅ c ⟆⟩

restrict : {c d : ob} → (f : Hom[ d , c ]) → ⟨sheafification⟅ c ⟆⟩ → ⟨sheafification⟅ d ⟆⟩

restrictId :
{c : ob} →
(x : ⟨sheafification⟅ c ⟆⟩) →
restrict id x ≡ x
restrictRestrict :
{c d e : ob} →
(f : Hom[ d , c ]) →
(g : Hom[ e , d ]) →
(x : ⟨sheafification⟅ c ⟆⟩) →
restrict (g ⋆ f) x ≡ restrict g (restrict f x)

η⟦⟧ : {c : ob} → (x : ⟨ P ⟅ c ⟆ ⟩) → ⟨sheafification⟅ c ⟆⟩

ηNatural :
{c d : ob} →
(f : Hom[ d , c ]) →
(x : ⟨ P ⟅ c ⟆ ⟩) →
η⟦⟧ ((P ⟪ f ⟫) x) ≡ restrict f (η⟦⟧ x)

sep :
{c : ob} →
(cover : ⟨ covers c ⟩) →
let cov = str (covers c) cover in
(x y : ⟨sheafification⟅ c ⟆⟩) →
(x~y :
(patch : ⟨ cov ⟩) →
restrict (patchArr cov patch) x ≡ restrict (patchArr cov patch) y) →
x ≡ y

amalgamate :
let
-- Is there any way to make this definition now and reuse it later?
sheafification : Presheaf C _
sheafification = record
{ F-ob = λ c → ⟨sheafification⟅ c ⟆⟩ , trunc
; F-hom = restrict
; F-id = funExt restrictId
; F-seq = λ f g → funExt (restrictRestrict f g)
}
in
{c : ob} →
(cover : ⟨ covers c ⟩) →
let cov = str (covers c) cover in
MatthiasHu marked this conversation as resolved.
Show resolved Hide resolved
(fam : CompatibleFamily sheafification cov) →
⟨sheafification⟅ c ⟆⟩
restrictAmalgamate :
let
sheafification : Presheaf C _
sheafification = record
{ F-ob = λ c → ⟨sheafification⟅ c ⟆⟩ , trunc
; F-hom = restrict
; F-id = funExt restrictId
; F-seq = λ f g → funExt (restrictRestrict f g)
}
in
{c : ob} →
(cover : ⟨ covers c ⟩) →
let cov = str (covers c) cover in
(fam : CompatibleFamily sheafification cov) →
(patch : ⟨ cov ⟩) →
restrict (patchArr cov patch) (amalgamate cover fam) ≡ fst fam patch

sheafification : Presheaf C (ℓ-max (ℓ-max (ℓ-max (ℓ-max ℓ ℓ') ℓcov) ℓpat) ℓP)
Functor.F-ob sheafification c = ⟨sheafification⟅ c ⟆⟩ , trunc
Functor.F-hom sheafification = restrict
Functor.F-id sheafification = funExt restrictId
Functor.F-seq sheafification f g = funExt (restrictRestrict f g)

isSheafSheafification : isSheaf J sheafification
isSheafSheafification c cover = isEmbedding×isSurjection→isEquiv
( injEmbedding
(isSetCompatibleFamily sheafification cov)
(λ {x} {y} x~y → sep cover x y (funExt⁻ (cong fst x~y)))
, λ fam →
∣ amalgamate cover fam
, Σ≡Prop
(isPropIsCompatibleFamily sheafification cov)
(funExt (restrictAmalgamate cover fam)) ∣₁ )
where
cov = str (covers c) cover
Loading