Skip to content

Commit 4a55c95

Browse files
committed
refactor(CategoryTheory/Sites): sheafification as an abstract left adjoint (#9012)
We define a typeclass `HasSheafify` which says that presheaves on a site with values in some category can be sheafified, i.e. that the inclusion functor from sheaves to presheaves has a left exact left adjoint. We redefine `presheafToSheaf` as an arbitrary choice of such a left adjoint.
1 parent 0470fa9 commit 4a55c95

12 files changed

+302
-103
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1297,6 +1297,7 @@ import Mathlib.CategoryTheory.Sites.RegularExtensive
12971297
import Mathlib.CategoryTheory.Sites.Sheaf
12981298
import Mathlib.CategoryTheory.Sites.SheafHom
12991299
import Mathlib.CategoryTheory.Sites.SheafOfTypes
1300+
import Mathlib.CategoryTheory.Sites.Sheafification
13001301
import Mathlib.CategoryTheory.Sites.Sieves
13011302
import Mathlib.CategoryTheory.Sites.Spaces
13021303
import Mathlib.CategoryTheory.Sites.Subsheaf

Mathlib/CategoryTheory/Sites/Adjunction.lean

Lines changed: 16 additions & 20 deletions
Original file line numberDiff line numberDiff line change
@@ -21,32 +21,26 @@ namespace CategoryTheory
2121

2222
open GrothendieckTopology CategoryTheory Limits Opposite
2323

24-
universe w₁ w₂ v u
24+
universe v u
2525

2626
variable {C : Type u} [Category.{v} C] (J : GrothendieckTopology C)
2727

28-
variable {D : Type w₁} [Category.{max v u} D]
28+
variable {D : Type*} [Category D]
2929

30-
variable {E : Type w₂} [Category.{max v u} E]
30+
variable {E : Type*} [Category E]
3131

3232
variable {F : D ⥤ E} {G : E ⥤ D}
3333

34-
variable [∀ (X : C) (S : J.Cover X) (P : Cᵒᵖ ⥤ D), PreservesLimit (S.index P).multicospan F]
35-
36-
variable [ConcreteCategory.{max v u} D] [PreservesLimits (forget D)]
34+
variable [HasWeakSheafify J D] [HasSheafCompose J F]
3735

3836
/-- The forgetful functor from `Sheaf J D` to sheaves of types, for a concrete category `D`
3937
whose forgetful functor preserves the correct limits. -/
40-
abbrev sheafForget : Sheaf J D ⥤ SheafOfTypes J :=
38+
abbrev sheafForget [ConcreteCategory D] [HasSheafCompose J (forget D)] :
39+
Sheaf J D ⥤ SheafOfTypes J :=
4140
sheafCompose J (forget D) ⋙ (sheafEquivSheafOfTypes J).functor
4241
set_option linter.uppercaseLean3 false in
4342
#align category_theory.Sheaf_forget CategoryTheory.sheafForget
4443

45-
-- We need to sheafify...
46-
variable [∀ (P : Cᵒᵖ ⥤ D) (X : C) (S : J.Cover X), HasMultiequalizer (S.index P)]
47-
[∀ X : C, HasColimitsOfShape (J.Cover X)ᵒᵖ D]
48-
[∀ X : C, PreservesColimitsOfShape (J.Cover X)ᵒᵖ (forget D)] [ReflectsIsomorphisms (forget D)]
49-
5044
namespace Sheaf
5145

5246
noncomputable section
@@ -63,21 +57,21 @@ set_option linter.uppercaseLean3 false in
6357
def composeEquiv (adj : G ⊣ F) (X : Sheaf J E) (Y : Sheaf J D) :
6458
((composeAndSheafify J G).obj X ⟶ Y) ≃ (X ⟶ (sheafCompose J F).obj Y) :=
6559
let A := adj.whiskerRight Cᵒᵖ
66-
{ toFun := fun η => ⟨A.homEquiv _ _ (J.toSheafify _ ≫ η.val)⟩
67-
invFun := fun γ => ⟨J.sheafifyLift ((A.homEquiv _ _).symm ((sheafToPresheaf _ _).map γ)) Y.2
60+
{ toFun := fun η => ⟨A.homEquiv _ _ (toSheafify J _ ≫ η.val)⟩
61+
invFun := fun γ => ⟨sheafifyLift J ((A.homEquiv _ _).symm ((sheafToPresheaf _ _).map γ)) Y.2
6862
left_inv := by
6963
intro η
7064
ext1
7165
dsimp
7266
symm
73-
apply J.sheafifyLift_unique
67+
apply sheafifyLift_unique
7468
rw [Equiv.symm_apply_apply]
7569
right_inv := by
7670
intro γ
7771
ext1
7872
dsimp
7973
-- This used to be `rw`, but we need `erw` after leanprover/lean4#2644
80-
erw [J.toSheafify_sheafifyLift, Equiv.apply_symm_apply] }
74+
erw [toSheafify_sheafifyLift, Equiv.apply_symm_apply] }
8175
set_option linter.uppercaseLean3 false in
8276
#align category_theory.Sheaf.compose_equiv CategoryTheory.Sheaf.composeEquiv
8377

@@ -113,6 +107,8 @@ instance [IsRightAdjoint F] : IsRightAdjoint (sheafCompose J F) :=
113107

114108
section ForgetToType
115109

110+
variable [ConcreteCategory D] [HasSheafCompose J (forget D)]
111+
116112
/-- This is the functor sending a sheaf of types `X` to the sheafification of `X ⋙ G`. -/
117113
abbrev composeAndSheafifyFromTypes (G : Type max v u ⥤ D) : SheafOfTypes J ⥤ Sheaf J D :=
118114
(sheafEquivSheafOfTypes J).inverse ⋙ composeAndSheafify _ G
@@ -132,7 +128,7 @@ theorem adjunctionToTypes_unit_app_val {G : Type max v u ⥤ D} (adj : G ⊣ for
132128
(Y : SheafOfTypes J) :
133129
((adjunctionToTypes J adj).unit.app Y).val =
134130
(adj.whiskerRight _).unit.app ((sheafOfTypesToPresheaf J).obj Y) ≫
135-
whiskerRight (J.toSheafify _) (forget D) := by
131+
whiskerRight (toSheafify J _) (forget D) := by
136132
dsimp [adjunctionToTypes, Adjunction.comp]
137133
simp
138134
rfl
@@ -143,12 +139,12 @@ set_option linter.uppercaseLean3 false in
143139
theorem adjunctionToTypes_counit_app_val {G : Type max v u ⥤ D} (adj : G ⊣ forget D)
144140
(X : Sheaf J D) :
145141
((adjunctionToTypes J adj).counit.app X).val =
146-
J.sheafifyLift ((Functor.associator _ _ _).hom ≫ (adj.whiskerRight _).counit.app _) X.2 := by
147-
apply J.sheafifyLift_unique
142+
sheafifyLift J ((Functor.associator _ _ _).hom ≫ (adj.whiskerRight _).counit.app _) X.2 := by
143+
apply sheafifyLift_unique
148144
dsimp only [adjunctionToTypes, Adjunction.comp, NatTrans.comp_app,
149145
instCategorySheaf_comp_val, instCategorySheaf_id_val]
150146
rw [adjunction_counit_app_val]
151-
erw [Category.id_comp, J.sheafifyMap_sheafifyLift, J.toSheafify_sheafifyLift]
147+
erw [Category.id_comp, sheafifyMap_sheafifyLift, toSheafify_sheafifyLift]
152148
ext
153149
dsimp [sheafEquivSheafOfTypes, Equivalence.symm, Equivalence.toAdjunction,
154150
NatIso.ofComponents, Adjunction.whiskerRight, Adjunction.mkOfUnitCounit]

Mathlib/CategoryTheory/Sites/CompatibleSheafification.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Adam Topaz
55
-/
66
import Mathlib.CategoryTheory.Sites.CompatiblePlus
7-
import Mathlib.CategoryTheory.Sites.ConcreteSheafification
7+
import Mathlib.CategoryTheory.Sites.LeftExact
88

99
#align_import category_theory.sites.compatible_sheafification from "leanprover-community/mathlib"@"70fd9563a21e7b963887c9360bd29b2393e6225a"
1010

Mathlib/CategoryTheory/Sites/ConcreteSheafification.lean

Lines changed: 13 additions & 37 deletions
Original file line numberDiff line numberDiff line change
@@ -3,8 +3,8 @@ Copyright (c) 2021 Adam Topaz. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Adam Topaz
55
-/
6-
import Mathlib.CategoryTheory.Adjunction.FullyFaithful
76
import Mathlib.CategoryTheory.Sites.Plus
7+
import Mathlib.CategoryTheory.Sites.Sheafification
88
import Mathlib.CategoryTheory.Limits.Shapes.ConcreteCategory
99
import Mathlib.CategoryTheory.ConcreteCategory.Elementwise
1010

@@ -525,7 +525,8 @@ noncomputable def toSheafification : 𝟭 _ ⟶ sheafification J D :=
525525
#align category_theory.grothendieck_topology.to_sheafification CategoryTheory.GrothendieckTopology.toSheafification
526526

527527
@[simp]
528-
theorem toSheafification_app (P : Cᵒᵖ ⥤ D) : (J.toSheafification D).app P = J.toSheafify P :=
528+
theorem toSheafification_app (P : Cᵒᵖ ⥤ D) :
529+
(J.toSheafification D).app P = J.toSheafify P :=
529530
rfl
530531
#align category_theory.grothendieck_topology.to_sheafification_app CategoryTheory.GrothendieckTopology.toSheafification_app
531532

@@ -550,8 +551,7 @@ theorem isoSheafify_hom {P : Cᵒᵖ ⥤ D} (hP : Presheaf.IsSheaf J P) :
550551
rfl
551552
#align category_theory.grothendieck_topology.iso_sheafify_hom CategoryTheory.GrothendieckTopology.isoSheafify_hom
552553

553-
/-- Given a sheaf `Q` and a morphism `P ⟶ Q`, construct a morphism from
554-
`J.sheafify P` to `Q`. -/
554+
/-- Given a sheaf `Q` and a morphism `P ⟶ Q`, construct a morphism from `J.sheafify P` to `Q`. -/
555555
noncomputable def sheafifyLift {P Q : Cᵒᵖ ⥤ D} (η : P ⟶ Q) (hQ : Presheaf.IsSheaf J Q) :
556556
J.sheafify P ⟶ Q :=
557557
J.plusLift (J.plusLift η hQ) hQ
@@ -613,26 +613,26 @@ variable (D)
613613

614614
/-- The sheafification functor, as a functor taking values in `Sheaf`. -/
615615
@[simps]
616-
noncomputable def presheafToSheaf : (Cᵒᵖ ⥤ D) ⥤ Sheaf J D where
616+
noncomputable def plusPlusSheaf : (Cᵒᵖ ⥤ D) ⥤ Sheaf J D where
617617
obj P := ⟨J.sheafify P, J.sheafify_isSheaf P⟩
618618
map η := ⟨J.sheafifyMap η⟩
619619
map_id _ := Sheaf.Hom.ext _ _ <| J.sheafifyMap_id _
620620
map_comp _ _ := Sheaf.Hom.ext _ _ <| J.sheafifyMap_comp _ _
621621
set_option linter.uppercaseLean3 false in
622-
#align category_theory.presheaf_to_Sheaf CategoryTheory.presheafToSheaf
622+
#align category_theory.presheaf_to_Sheaf CategoryTheory.plusPlusSheaf
623623

624-
instance presheafToSheaf_preservesZeroMorphisms [Preadditive D] :
625-
(presheafToSheaf J D).PreservesZeroMorphisms where
624+
instance plusPlusSheaf_preservesZeroMorphisms [Preadditive D] :
625+
(plusPlusSheaf J D).PreservesZeroMorphisms where
626626
map_zero F G := by
627627
ext : 3
628628
refine' colimit.hom_ext (fun j => _)
629629
erw [colimit.ι_map, comp_zero, J.plusMap_zero, J.diagramNatTrans_zero, zero_comp]
630630
set_option linter.uppercaseLean3 false in
631-
#align category_theory.presheaf_to_Sheaf_preserves_zero_morphisms CategoryTheory.presheafToSheaf_preservesZeroMorphisms
631+
#align category_theory.presheaf_to_Sheaf_preserves_zero_morphisms CategoryTheory.plusPlusSheaf_preservesZeroMorphisms
632632

633633
/-- The sheafification functor is left adjoint to the forgetful functor. -/
634634
@[simps! unit_app counit_app_val]
635-
noncomputable def sheafificationAdjunction : presheafToSheaf J D ⊣ sheafToPresheaf J D :=
635+
noncomputable def plusPlusAdjunction : plusPlusSheaf J D ⊣ sheafToPresheaf J D :=
636636
Adjunction.mkOfHomEquiv
637637
{ homEquiv := fun P Q =>
638638
{ toFun := fun e => J.toSheafify P ≫ e.val
@@ -645,10 +645,10 @@ noncomputable def sheafificationAdjunction : presheafToSheaf J D ⊣ sheafToPres
645645
homEquiv_naturality_right := fun η γ => by
646646
dsimp
647647
rw [Category.assoc] }
648-
#align category_theory.sheafification_adjunction CategoryTheory.sheafificationAdjunction
648+
#align category_theory.sheafification_adjunction CategoryTheory.plusPlusAdjunction
649649

650650
noncomputable instance sheafToPresheafIsRightAdjoint : IsRightAdjoint (sheafToPresheaf J D) :=
651-
⟨_, sheafificationAdjunction J D⟩
651+
⟨_, plusPlusAdjunction J D⟩
652652
set_option linter.uppercaseLean3 false in
653653
#align category_theory.Sheaf_to_presheaf_is_right_adjoint CategoryTheory.sheafToPresheafIsRightAdjoint
654654

@@ -666,31 +666,7 @@ set_option linter.uppercaseLean3 false in
666666
@[simps! hom_app inv_app]
667667
noncomputable
668668
def GrothendieckTopology.sheafificationIsoPresheafToSheafCompSheafToPreasheaf :
669-
J.sheafification D ≅ presheafToSheaf J D ⋙ sheafToPresheaf J D :=
669+
J.sheafification D ≅ plusPlusSheaf J D ⋙ sheafToPresheaf J D :=
670670
NatIso.ofComponents fun P => Iso.refl _
671671

672-
variable {J D}
673-
674-
/-- A sheaf `P` is isomorphic to its own sheafification. -/
675-
@[simps]
676-
noncomputable def sheafificationIso (P : Sheaf J D) : P ≅ (presheafToSheaf J D).obj P.val where
677-
hom := ⟨(J.isoSheafify P.2).hom⟩
678-
inv := ⟨(J.isoSheafify P.2).inv⟩
679-
hom_inv_id := by
680-
ext1
681-
apply (J.isoSheafify P.2).hom_inv_id
682-
inv_hom_id := by
683-
ext1
684-
apply (J.isoSheafify P.2).inv_hom_id
685-
#align category_theory.sheafification_iso CategoryTheory.sheafificationIso
686-
687-
instance isIso_sheafificationAdjunction_counit (P : Sheaf J D) :
688-
IsIso ((sheafificationAdjunction J D).counit.app P) :=
689-
isIso_of_fully_faithful (sheafToPresheaf J D) _
690-
#align category_theory.is_iso_sheafification_adjunction_counit CategoryTheory.isIso_sheafificationAdjunction_counit
691-
692-
instance sheafification_reflective : IsIso (sheafificationAdjunction J D).counit :=
693-
NatIso.isIso_of_isIso_app _
694-
#align category_theory.sheafification_reflective CategoryTheory.sheafification_reflective
695-
696672
end CategoryTheory

Mathlib/CategoryTheory/Sites/ConstantSheaf.lean

Lines changed: 1 addition & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -37,10 +37,7 @@ noncomputable def constantPresheafAdj {T : C} (hT : IsTerminal T) :
3737
simp }
3838
naturality := by intros; ext; simp /- Note: `aesop` works but is kind of slow -/ } }
3939

40-
variable [ConcreteCategory D] [PreservesLimits (forget D)]
41-
[∀ (P : Cᵒᵖ ⥤ D) (X : C) (S : J.Cover X), HasMultiequalizer (S.index P)]
42-
[∀ X : C, HasColimitsOfShape (J.Cover X)ᵒᵖ D]
43-
[∀ X : C, PreservesColimitsOfShape (J.Cover X)ᵒᵖ (forget D)] [ReflectsIsomorphisms (forget D)]
40+
variable [HasWeakSheafify J D]
4441

4542
/--
4643
The functor which maps an object of `D` to the constant sheaf at that object, i.e. the

Mathlib/CategoryTheory/Sites/CoverLifting.lean

Lines changed: 9 additions & 16 deletions
Original file line numberDiff line numberDiff line change
@@ -369,13 +369,7 @@ noncomputable def Functor.sheafAdjunctionCocontinuous [G.IsCocontinuous J K]
369369
#align category_theory.sites.pullback_copullback_adjunction CategoryTheory.Functor.sheafAdjunctionCocontinuous
370370

371371
variable
372-
[ConcreteCategory.{max v u} A]
373-
[PreservesLimits (forget A)]
374-
[ReflectsIsomorphisms (forget A)]
375-
[∀ (X : C), PreservesColimitsOfShape (J.Cover X)ᵒᵖ (forget A)]
376-
[∀ (X : C), HasColimitsOfShape (J.Cover X)ᵒᵖ A]
377-
[∀ (X : D), PreservesColimitsOfShape (K.Cover X)ᵒᵖ (forget A)]
378-
[∀ (X : D), HasColimitsOfShape (K.Cover X)ᵒᵖ A]
372+
[HasWeakSheafify J A] [HasWeakSheafify K A]
379373
[G.IsCocontinuous J K] [G.IsContinuous J K]
380374

381375
/-- The natural isomorphism exhibiting compatibility between pushforward and sheafification. -/
@@ -393,35 +387,34 @@ def Functor.pushforwardContinuousSheafificationCompatibility :
393387
/- Implementation: This is primarily used to prove the lemma
394388
`pullbackSheafificationCompatibility_hom_app_val`. -/
395389
lemma Functor.toSheafify_pullbackSheafificationCompatibility (F : Dᵒᵖ ⥤ A) :
396-
J.toSheafify (G.op ⋙ F) ≫
390+
toSheafify J (G.op ⋙ F) ≫
397391
((G.pushforwardContinuousSheafificationCompatibility A J K).hom.app F).val =
398-
whiskerLeft _ (K.toSheafify _) := by
392+
whiskerLeft _ (toSheafify K _) := by
399393
dsimp [pushforwardContinuousSheafificationCompatibility, Adjunction.leftAdjointUniq]
400394
apply Quiver.Hom.op_inj
401395
apply coyoneda.map_injective
402396
ext E : 2
403397
dsimp [Functor.preimage, Full.preimage, coyoneda, Adjunction.leftAdjointsCoyonedaEquiv]
404398
erw [Adjunction.homEquiv_unit, Adjunction.homEquiv_counit]
405399
dsimp [Adjunction.comp]
406-
simp only [sheafificationAdjunction_unit_app, Category.comp_id, Functor.map_id,
407-
whiskerLeft_id', GrothendieckTopology.sheafifyMap_comp,
408-
GrothendieckTopology.sheafifyMap_sheafifyLift, Category.id_comp,
409-
Category.assoc, GrothendieckTopology.toSheafify_sheafifyLift]
400+
simp only [Category.comp_id, map_id, whiskerLeft_id', map_comp, Sheaf.instCategorySheaf_comp_val,
401+
sheafificationAdjunction_counit_app_val, sheafifyMap_sheafifyLift,
402+
Category.id_comp, Category.assoc, toSheafify_sheafifyLift]
410403
ext t s : 3
411404
dsimp [sheafPushforwardContinuous]
412405
congr 1
413406
simp only [← Category.assoc]
414407
convert Category.id_comp (obj := A) _
415408
have := (Ran.adjunction A G.op).left_triangle
416-
apply_fun (fun e => (e.app (K.sheafify F)).app s) at this
409+
apply_fun (fun e => (e.app (sheafify K F)).app s) at this
417410
exact this
418411

419412
@[simp]
420413
lemma Functor.pushforwardContinuousSheafificationCompatibility_hom_app_val (F : Dᵒᵖ ⥤ A) :
421414
((G.pushforwardContinuousSheafificationCompatibility A J K).hom.app F).val =
422-
J.sheafifyLift (whiskerLeft G.op <| K.toSheafify F)
415+
sheafifyLift J (whiskerLeft G.op <| toSheafify K F)
423416
((presheafToSheaf K A ⋙ G.sheafPushforwardContinuous A J K).obj F).cond := by
424-
apply J.sheafifyLift_unique
417+
apply sheafifyLift_unique
425418
apply toSheafify_pullbackSheafificationCompatibility
426419

427420
end CategoryTheory

Mathlib/CategoryTheory/Sites/DenseSubsite.lean

Lines changed: 1 addition & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -545,14 +545,7 @@ noncomputable def sheafEquivOfCoverPreservingCoverLifting : Sheaf J A ≌ Sheaf
545545
set_option linter.uppercaseLean3 false in
546546
#align category_theory.cover_dense.Sheaf_equiv_of_cover_preserving_cover_lifting CategoryTheory.Functor.IsCoverDense.sheafEquivOfCoverPreservingCoverLifting
547547

548-
variable
549-
[ConcreteCategory.{max v u} A]
550-
[Limits.PreservesLimits (forget A)]
551-
[ReflectsIsomorphisms (forget A)]
552-
[∀ (X : C), Limits.PreservesColimitsOfShape (J.Cover X)ᵒᵖ (forget A)]
553-
[∀ (X : C), Limits.HasColimitsOfShape (J.Cover X)ᵒᵖ A]
554-
[∀ (X : D), Limits.PreservesColimitsOfShape (K.Cover X)ᵒᵖ (forget A)]
555-
[∀ (X : D), Limits.HasColimitsOfShape (K.Cover X)ᵒᵖ A]
548+
variable [HasWeakSheafify J A] [HasWeakSheafify K A]
556549

557550
/-- The natural isomorphism exhibiting the compatibility of
558551
`sheafEquivOfCoverPreservingCoverLifting` with sheafification. -/

Mathlib/CategoryTheory/Sites/LeftExact.lean

Lines changed: 26 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -3,7 +3,6 @@ Copyright (c) 2021 Adam Topaz. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Adam Topaz
55
-/
6-
import Mathlib.CategoryTheory.Sites.ConcreteSheafification
76
import Mathlib.CategoryTheory.Sites.Limits
87
import Mathlib.CategoryTheory.Limits.FunctorCategory
98
import Mathlib.CategoryTheory.Limits.FilteredColimitCommutesFiniteLimit
@@ -243,7 +242,7 @@ variable (K : Type w')
243242
variable [SmallCategory K] [FinCategory K] [HasLimitsOfShape K D]
244243

245244
instance preservesLimitsOfShape_presheafToSheaf :
246-
PreservesLimitsOfShape K (presheafToSheaf J D) := by
245+
PreservesLimitsOfShape K (plusPlusSheaf J D) := by
247246
let e := (FinCategory.equivAsType K).symm.trans (AsSmall.equiv.{0, 0, max v u})
248247
haveI : HasLimitsOfShape (AsSmall.{max v u} (FinCategory.AsType K)) D :=
249248
Limits.hasLimitsOfShape_of_equivalence e
@@ -261,22 +260,42 @@ instance preservesLimitsOfShape_presheafToSheaf :
261260
reflectsLimitsOfShapeOfReflectsIsomorphisms
262261
-- porting note: the mathlib proof was by `apply is_limit_of_preserves (J.sheafification D) hS`
263262
have : PreservesLimitsOfShape (AsSmall.{max v u} (FinCategory.AsType K))
264-
(presheafToSheaf J D ⋙ sheafToPresheaf J D) :=
263+
(plusPlusSheaf J D ⋙ sheafToPresheaf J D) :=
265264
preservesLimitsOfShapeOfNatIso (J.sheafificationIsoPresheafToSheafCompSheafToPreasheaf D)
266-
exact isLimitOfPreserves (presheafToSheaf J D ⋙ sheafToPresheaf J D) hS
265+
exact isLimitOfPreserves (plusPlusSheaf J D ⋙ sheafToPresheaf J D) hS
267266

268267
instance preservesfiniteLimits_presheafToSheaf [HasFiniteLimits D] :
269-
PreservesFiniteLimits (presheafToSheaf J D) := by
268+
PreservesFiniteLimits (plusPlusSheaf J D) := by
270269
apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{max v u}
271270
intros
272271
infer_instance
273272

273+
instance : HasWeakSheafify J D := ⟨sheafToPresheafIsRightAdjoint J D⟩
274+
275+
variable (J D)
276+
277+
/-- `plusPlusSheaf` is isomorphic to an arbitrary choice of left adjoint. -/
278+
def plusPlusSheafIsoPresheafToSheaf : plusPlusSheaf J D ≅ presheafToSheaf J D :=
279+
(plusPlusAdjunction J D).leftAdjointUniq (sheafificationAdjunction J D)
280+
281+
/-- `plusPlusFunctor` is isomorphic to `sheafification`. -/
282+
def plusPlusFunctorIsoSheafification : J.sheafification D ≅ sheafification J D :=
283+
isoWhiskerRight (plusPlusSheafIsoPresheafToSheaf J D) (sheafToPresheaf J D)
284+
285+
/-- `plusPlus` is isomorphic to `sheafify`. -/
286+
def plusPlusIsoSheafify (P : Cᵒᵖ ⥤ D) : J.sheafify P ≅ sheafify J P :=
287+
(sheafToPresheaf J D).mapIso ((plusPlusSheafIsoPresheafToSheaf J D).app P)
288+
289+
instance [HasFiniteLimits D] : HasSheafify J D := HasSheafify.mk' J D (plusPlusAdjunction J D)
290+
291+
variable {J D}
292+
274293
instance [FinitaryExtensive D] [HasFiniteCoproducts D] [HasPullbacks D] :
275294
FinitaryExtensive (Sheaf J D) :=
276-
finitaryExtensive_of_reflective (sheafificationAdjunction _ _)
295+
finitaryExtensive_of_reflective (plusPlusAdjunction _ _)
277296

278297
instance [Adhesive D] [HasPullbacks D] [HasPushouts D] : Adhesive (Sheaf J D) :=
279-
adhesive_of_reflective (sheafificationAdjunction _ _)
298+
adhesive_of_reflective (plusPlusAdjunction _ _)
280299

281300
instance SheafOfTypes.finitary_extensive {C : Type u} [SmallCategory C]
282301
(J : GrothendieckTopology C) : FinitaryExtensive (Sheaf J (Type u)) :=

0 commit comments

Comments
 (0)