@@ -3,8 +3,8 @@ Copyright (c) 2021 Adam Topaz. All rights reserved.
3
3
Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Adam Topaz
5
5
-/
6
- import Mathlib.CategoryTheory.Adjunction.FullyFaithful
7
6
import Mathlib.CategoryTheory.Sites.Plus
7
+ import Mathlib.CategoryTheory.Sites.Sheafification
8
8
import Mathlib.CategoryTheory.Limits.Shapes.ConcreteCategory
9
9
import Mathlib.CategoryTheory.ConcreteCategory.Elementwise
10
10
@@ -525,7 +525,8 @@ noncomputable def toSheafification : 𝟭 _ ⟶ sheafification J D :=
525
525
#align category_theory.grothendieck_topology.to_sheafification CategoryTheory.GrothendieckTopology.toSheafification
526
526
527
527
@[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 :=
529
530
rfl
530
531
#align category_theory.grothendieck_topology.to_sheafification_app CategoryTheory.GrothendieckTopology.toSheafification_app
531
532
@@ -550,8 +551,7 @@ theorem isoSheafify_hom {P : Cᵒᵖ ⥤ D} (hP : Presheaf.IsSheaf J P) :
550
551
rfl
551
552
#align category_theory.grothendieck_topology.iso_sheafify_hom CategoryTheory.GrothendieckTopology.isoSheafify_hom
552
553
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`. -/
555
555
noncomputable def sheafifyLift {P Q : Cᵒᵖ ⥤ D} (η : P ⟶ Q) (hQ : Presheaf.IsSheaf J Q) :
556
556
J.sheafify P ⟶ Q :=
557
557
J.plusLift (J.plusLift η hQ) hQ
@@ -613,26 +613,26 @@ variable (D)
613
613
614
614
/-- The sheafification functor, as a functor taking values in `Sheaf`. -/
615
615
@[simps]
616
- noncomputable def presheafToSheaf : (Cᵒᵖ ⥤ D) ⥤ Sheaf J D where
616
+ noncomputable def plusPlusSheaf : (Cᵒᵖ ⥤ D) ⥤ Sheaf J D where
617
617
obj P := ⟨J.sheafify P, J.sheafify_isSheaf P⟩
618
618
map η := ⟨J.sheafifyMap η⟩
619
619
map_id _ := Sheaf.Hom.ext _ _ <| J.sheafifyMap_id _
620
620
map_comp _ _ := Sheaf.Hom.ext _ _ <| J.sheafifyMap_comp _ _
621
621
set_option linter.uppercaseLean3 false in
622
- #align category_theory.presheaf_to_Sheaf CategoryTheory.presheafToSheaf
622
+ #align category_theory.presheaf_to_Sheaf CategoryTheory.plusPlusSheaf
623
623
624
- instance presheafToSheaf_preservesZeroMorphisms [Preadditive D] :
625
- (presheafToSheaf J D).PreservesZeroMorphisms where
624
+ instance plusPlusSheaf_preservesZeroMorphisms [Preadditive D] :
625
+ (plusPlusSheaf J D).PreservesZeroMorphisms where
626
626
map_zero F G := by
627
627
ext : 3
628
628
refine' colimit.hom_ext (fun j => _)
629
629
erw [colimit.ι_map, comp_zero, J.plusMap_zero, J.diagramNatTrans_zero, zero_comp]
630
630
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
632
632
633
633
/-- The sheafification functor is left adjoint to the forgetful functor. -/
634
634
@[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 :=
636
636
Adjunction.mkOfHomEquiv
637
637
{ homEquiv := fun P Q =>
638
638
{ toFun := fun e => J.toSheafify P ≫ e.val
@@ -645,10 +645,10 @@ noncomputable def sheafificationAdjunction : presheafToSheaf J D ⊣ sheafToPres
645
645
homEquiv_naturality_right := fun η γ => by
646
646
dsimp
647
647
rw [Category.assoc] }
648
- #align category_theory.sheafification_adjunction CategoryTheory.sheafificationAdjunction
648
+ #align category_theory.sheafification_adjunction CategoryTheory.plusPlusAdjunction
649
649
650
650
noncomputable instance sheafToPresheafIsRightAdjoint : IsRightAdjoint (sheafToPresheaf J D) :=
651
- ⟨_, sheafificationAdjunction J D⟩
651
+ ⟨_, plusPlusAdjunction J D⟩
652
652
set_option linter.uppercaseLean3 false in
653
653
#align category_theory.Sheaf_to_presheaf_is_right_adjoint CategoryTheory.sheafToPresheafIsRightAdjoint
654
654
@@ -666,31 +666,7 @@ set_option linter.uppercaseLean3 false in
666
666
@[simps! hom_app inv_app]
667
667
noncomputable
668
668
def GrothendieckTopology.sheafificationIsoPresheafToSheafCompSheafToPreasheaf :
669
- J.sheafification D ≅ presheafToSheaf J D ⋙ sheafToPresheaf J D :=
669
+ J.sheafification D ≅ plusPlusSheaf J D ⋙ sheafToPresheaf J D :=
670
670
NatIso.ofComponents fun P => Iso.refl _
671
671
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
-
696
672
end CategoryTheory
0 commit comments