Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit d45708f

Browse files
committed
feat(topology/sheaves): The empty component of a sheaf is terminal (#10578)
1 parent 1e18e93 commit d45708f

File tree

2 files changed

+21
-0
lines changed

2 files changed

+21
-0
lines changed

src/category_theory/sites/sheaf.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -168,6 +168,18 @@ def Sheaf_equiv_SheafOfTypes : Sheaf J (Type w) ≌ SheafOfTypes J :=
168168
instance : inhabited (Sheaf (⊥ : grothendieck_topology C) (Type w)) :=
169169
⟨(Sheaf_equiv_SheafOfTypes _).inverse.obj (default _)⟩
170170

171+
variables {J} {A}
172+
173+
/-- If the empty sieve is a cover of `X`, then `F(X)` is terminal. -/
174+
def Sheaf.is_terminal_of_bot_cover (F : Sheaf J A) (X : C) (H : ⊥ ∈ J X) :
175+
is_terminal (F.1.obj (op X)) :=
176+
begin
177+
apply_with is_terminal.of_unique { instances := ff },
178+
intro Y,
179+
choose t h using F.2 Y _ H (by tidy) (by tidy),
180+
exact ⟨⟨t⟩, λ a, h.2 a (by tidy)⟩
181+
end
182+
171183
end category_theory
172184

173185
namespace category_theory

src/topology/sheaves/sheaf_condition/sites.lean

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -514,6 +514,15 @@ variables {C : Type u} [category.{v} C] [limits.has_products C]
514514
variables {X : Top.{v}} {ι : Type*} {B : ι → opens X}
515515
variables (F : presheaf C X) (F' : sheaf C X) (h : opens.is_basis (set.range B))
516516

517+
/-- The empty component of a sheaf is terminal -/
518+
def is_terminal_of_empty (F : sheaf C X) : limits.is_terminal (F.val.obj (op ∅)) :=
519+
((presheaf.Sheaf_spaces_to_sheaf_sites C X).obj F).is_terminal_of_bot_cover ∅ (by tidy)
520+
521+
/-- A variant of `is_terminal_of_empty` that is easier to `apply`. -/
522+
def is_terminal_of_eq_empty (F : X.sheaf C) {U : opens X} (h : U = ∅) :
523+
limits.is_terminal (F.val.obj (op U)) :=
524+
by convert F.is_terminal_of_empty
525+
517526
/-- If a family `B` of open sets forms a basis of the topology on `X`, and if `F'`
518527
is a sheaf on `X`, then a homomorphism between a presheaf `F` on `X` and `F'`
519528
is equivalent to a homomorphism between their restrictions to the indexing type

0 commit comments

Comments
 (0)