module Cat.Diagram.Sieve where
module _ {o κ : _} (C : Precategory o κ) (c : ⌞ C ⌟) where
private module C = Precategory C
Sieves🔗
Given a category a sieve on an object Is a subset of the maps closed under composition: If then The data of a sieve on corresponds to the data of a subobject of considered as an object of
Here, the subset is represented as the function arrows
, which, given an arrow
(and its domain), yields a proposition representing inclusion in the
subset.
record Sieve : Type (o ⊔ κ) where
no-eta-equality
field
: ∀ {y} → ℙ (C.Hom y c)
arrows : ∀ {y z f} (hf : f ∈ arrows) (g : C.Hom y z) → (f C.∘ g) ∈ arrows
closed open Sieve public
The maximal
sieve on an object
is the collection of all maps
It represents the identity map
as a subfunctor. A
family of sieves can be intersected (the underlying predicate is the
“
conjunction”
— the universal quantifier), and this represents a wide pullback of
subobjects.
module _ {o ℓ : _} {C : Precategory o ℓ} where
private
module C = Cat.Reasoning C
module PSh = Cat.Reasoning (PSh ℓ C)
open Precategory C
: ∀ {c} {x y : Sieve C c} → Path (∀ {y} → ℙ (C.Hom y c)) (x .arrows) (y .arrows) → x ≡ y
Sieve-path {x = x} {y} p i .arrows = p i
Sieve-path {x = x} {y} p i .closed {f = f} hf g =
Sieve-path (λ i → fun-is-hlevel {A = ⌞ p i f ⌟} 1 (p i (f ∘ g) .is-tr)) (λ w → x .closed w g) (λ w → y .closed w g) i hf
is-prop→pathp
instance
: ∀ {c d} → Membership (C.Hom d c) (Sieve C c) _
hom∈Sieve = record { _∈_ = λ x S → x ∈ S .Sieve.arrows }
hom∈Sieve
: ∀ {c} → Membership (/-Obj {C = C} c) (Sieve C c) _
slice∈Sieve = record { _∈_ = λ x S → x .map ∈ S }
slice∈Sieve
: ∀ {U} → Inclusion (Sieve C U) _
Inclusion-sieve {U} = record { _⊆_ = λ S T → ∀ {V} (h : Hom V U) → h ∈ S → h ∈ T }
Inclusion-sieve
: ∀ {ℓr c} ⦃ _ : Extensional (∀ {y} → C.Hom y c → Ω) ℓr ⦄ → Extensional (Sieve C c) ℓr
Extensional-sieve = injection→extensional! Sieve-path e
Extensional-sieve ⦃ e ⦄
: ∀ {c n} → H-Level (Sieve C c) (2 + n)
H-Level-Sieve = basic-instance 2 $
H-Level-Sieve 1 (injective→is-embedding! Sieve-path) (hlevel 2)
embedding→is-hlevel
open PSh._↪_
open _=>_
open Functor
: ∀ {c} → Sieve C c
maximal' .arrows x = ⊤Ω
maximal' .closed g x = tt
maximal'
: ∀ {c} {I : Type ℓ} (F : I → Sieve C c) → Sieve C c
intersect {I = I} F .arrows h = elΩ ((x : I) → h ∈ F x)
intersect {I = I} F .closed x g = inc λ i → F i .closed (□-out! x i) g intersect
_∩S_ : ∀ {U} → Sieve C U → Sieve C U → Sieve C U
(S ∩S T) .arrows f = S .arrows f ∧Ω T .arrows f
(S ∩S T) .closed (Sf , Tf) g = S .closed Sf g , T .closed Tf g
Representing subfunctors🔗
Let be a sieve on We show that it determines a presheaf and that this presheaf admits a monic natural transformation The functor determined by a sieve sends each object to the set of arrows s.t. The functorial action is given by composition, as with the functor.
: ∀ {c} → Sieve C c → PSh.Ob
to-presheaf {c} sieve .F₀ d = el! (Σ[ f ∈ C.Hom d c ] (f ∈ sieve))
to-presheaf .F₁ f (g , s) = g C.∘ f , sieve .closed s _ to-presheaf sieve
.F-id = funext λ _ → Σ-prop-path! (C.idr _)
to-presheaf sieve .F-∘ f g = funext λ _ → Σ-prop-path! (C.assoc _ _ _) to-presheaf sieve
That this functor is a subobject of follows straightforwardly: The injection map is given by projecting out the first component, which is an embedding (since “being in a sieve” is a proposition). Since natural transformations are monic if they are componentwise monic, and embeddings are monic, the result follows.
: ∀ {c} {S : Sieve C c} → to-presheaf S PSh.↪ よ₀ C c
to-presheaf↪よ {S} .mor .η x (f , _) = f
to-presheaf↪よ {S} .mor .is-natural x y f = refl
to-presheaf↪よ {S} .monic g h path = ext λ i x → Σ-prop-path! (unext path i x) to-presheaf↪よ
Pullback of sieves🔗
If we have a sieve on and any morphism then there is a natural way to restrict the to a sieve on a morphism belongs to the restriction if the composite belongs to We refer to the resulting sieve as the pullback of along , and write it
: ∀ {u v} → C.Hom v u → Sieve C u → Sieve C v
pullback .arrows h = el (f C.∘ h ∈ s) (hlevel 1)
pullback f s .closed hf g = subst (_∈ s) (sym (C.assoc f _ g)) (s .closed hf g) pullback f s
If we consider the collection of “sieves on ” as varying along as a parameter, the pullback operation becomes functorial. Since it is contravariant, this means that the assignment is itself a presheaf on
abstract
: ∀ {c} {s : Sieve C c} → pullback C.id s ≡ s
pullback-id {s = s} = ext λ h → Ω-ua
pullback-id (subst (_∈ s) (C.idl h))
(subst (_∈ s) (sym (C.idl h)))
pullback-∘: ∀ {u v w} {f : C.Hom w v} {g : C.Hom v u} {R : Sieve C u}
→ pullback (g C.∘ f) R ≡ pullback f (pullback g R)
{f = f} {g} {R = R} = ext λ h → Ω-ua
pullback-∘ (subst (_∈ R) (sym (C.assoc g f h)))
(subst (_∈ R) (C.assoc g f h))
This presheaf has an important universal property: the natural transformations correspond naturally to the subobjects of Categorically, we say that is a subobject classifier in the category
: Functor (C ^op) (Sets (o ⊔ ℓ))
Sieves .F₀ U .∣_∣ = Sieve C U
Sieves .F₀ U .is-tr = hlevel 2
Sieves .F₁ = pullback
Sieves .F-id = funext λ x → pullback-id
Sieves .F-∘ f g = funext λ x → pullback-∘ Sieves
Generated sieves🔗
Often, it’s more practical to define a family of maps, and
to obtain a sieve from this family after the fact. To this end, we
define a type Cover
for families of
maps into a common codomain, paired with their indexing type.
module _ {o ℓ} (C : Precategory o ℓ) where
open Precategory C using (Hom)
record Cover (U : ⌞ C ⌟) ℓ' : Type (o ⊔ ℓ ⊔ lsuc ℓ') where
field
{index} : Type ℓ'
{domain} : index → ⌞ C ⌟
: ∀ i → Hom (domain i) U map
open Cover
module _ {o ℓ} {C : Precategory o ℓ} where
private module C = Cat.Reasoning C
instance
: ∀ {ℓ' U} → Underlying (Cover C U ℓ')
Underlying-Cover = record { ⌞_⌟ = index } Underlying-Cover
The sieve generated by a cover is the collection of maps that factor through at least one of the i.e., for a map it is the proposition
: ∀ {ℓ' U} → Cover C U ℓ' → Sieve C U
cover→sieve {U = U} f .arrows {W} g = elΩ do
cover→sieve .Hom W (f .domain i) ] (f .map i C.∘ h ≡ g)
Σ[ i ∈ f ] Σ[ h ∈ C.closed = rec! λ i h p g → inc (i , h C.∘ g , C.pulll p) cover→sieve f
Since the primary purpose of Cover
is to present sieves, we register
an instance of the ⟦⟧-notation
class, so that we can write ⟦ cov ⟧
instead of
cover→sieve cov
.
: ∀ {V U} → C.Hom V U → Sieve C U
map→sieve .arrows g = elΩ (Σ[ h ∈ C.Hom _ _ ] (f C.∘ h ≡ g))
map→sieve f .closed = rec! λ g p h → inc (g C.∘ h , C.pulll p)
map→sieve f
instance
: ∀ {ℓ' U} → ⟦⟧-notation (Cover C U ℓ')
⟦⟧-Cover = brackets _ cover→sieve ⟦⟧-Cover