|
| 1 | +/- |
| 2 | +Copyright (c) 2024 Joël Riou. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Joël Riou |
| 5 | +-/ |
| 6 | +import Mathlib.CategoryTheory.Sites.Sheaf |
| 7 | + |
| 8 | +/-! |
| 9 | +# 1-hypercovers |
| 10 | +
|
| 11 | +Given a Grothendieck topology `J` on a category `C`, we define the type of |
| 12 | +`1`-hypercovers of an object `S : C`. They consists of a covering family |
| 13 | +of morphisms `X i ⟶ S` indexed by a type `I₀` and, for each tuple `(i₁, i₂)` |
| 14 | +of elements of `I₀`, a "covering `Y j` of the fibre product of `X i₁` and |
| 15 | +`X i₂` over `S`", a condition which is phrased here without assuming that |
| 16 | +the fibre product actually exist. |
| 17 | +
|
| 18 | +The definition `OneHypercover.isLimitMultifork` shows that if `E` is a |
| 19 | +`1`-hypercover of `S`, and `F` is a sheaf, then `F.obj (op S)` |
| 20 | +identifies to the multiequalizer of suitable maps |
| 21 | +`F.obj (op (E.X i)) ⟶ F.obj (op (E.Y j))`. |
| 22 | +
|
| 23 | +-/ |
| 24 | + |
| 25 | +universe w v u |
| 26 | + |
| 27 | +namespace CategoryTheory |
| 28 | + |
| 29 | +open Category Limits |
| 30 | + |
| 31 | +variable {C : Type u} [Category.{v} C] {A : Type*} [Category A] |
| 32 | + |
| 33 | +/-- The categorical data that is involved in a `1`-hypercover of an object `S`. This |
| 34 | +consists of a family of morphisms `f i : X i ⟶ S` for `i : I₀`, and for each |
| 35 | +tuple `(i₁, i₂)` of elements in `I₀`, a family of objects `Y j` indexed by |
| 36 | +a type `I₁ i₁ i₂`, which are equipped with a map to the fibre product of `X i₁` |
| 37 | +and `X i₂`, which is phrased here as the data of the two projections |
| 38 | +`p₁ : Y j ⟶ X i₁`, `p₂ : Y j ⟶ X i₂` and the relation `p₁ j ≫ f i₁ = p₂ j ≫ f i₂`. |
| 39 | +(See `GrothendieckTopology.OneHypercover` for the topological conditions.) -/ |
| 40 | +structure PreOneHypercover (S : C) where |
| 41 | + /-- the index type of the covering of `S` -/ |
| 42 | + I₀ : Type w |
| 43 | + /-- the objects in the covering of `S` -/ |
| 44 | + X (i : I₀) : C |
| 45 | + /-- the morphisms in the covering of `S` -/ |
| 46 | + f (i : I₀) : X i ⟶ S |
| 47 | + /-- the index type of the coverings of the fibre products -/ |
| 48 | + I₁ (i₁ i₂ : I₀) : Type w |
| 49 | + /-- the objects in the coverings of the fibre products -/ |
| 50 | + Y ⦃i₁ i₂ : I₀⦄ (j : I₁ i₁ i₂) : C |
| 51 | + /-- the first projection `Y j ⟶ X i₁` -/ |
| 52 | + p₁ ⦃i₁ i₂ : I₀⦄ (j : I₁ i₁ i₂) : Y j ⟶ X i₁ |
| 53 | + /-- the second projection `Y j ⟶ X i₂` -/ |
| 54 | + p₂ ⦃i₁ i₂ : I₀⦄ (j : I₁ i₁ i₂) : Y j ⟶ X i₂ |
| 55 | + w ⦃i₁ i₂ : I₀⦄ (j : I₁ i₁ i₂) : p₁ j ≫ f i₁ = p₂ j ≫ f i₂ |
| 56 | + |
| 57 | +namespace PreOneHypercover |
| 58 | + |
| 59 | +variable {S : C} (E : PreOneHypercover.{w} S) |
| 60 | + |
| 61 | +/-- The assumption that the pullback of `X i₁` and `X i₂` over `S` exists |
| 62 | +for any `i₁` and `i₂`. -/ |
| 63 | +abbrev HasPullbacks := ∀ (i₁ i₂ : E.I₀), HasPullback (E.f i₁) (E.f i₂) |
| 64 | + |
| 65 | +/-- The sieve of `S` that is generated by the morphisms `f i : X i ⟶ S`. -/ |
| 66 | +abbrev sieve₀ : Sieve S := Sieve.ofArrows _ E.f |
| 67 | + |
| 68 | +/-- Given an object `W` equipped with morphisms `p₁ : W ⟶ E.X i₁`, `p₂ : W ⟶ E.X i₂`, |
| 69 | +this is the sieve of `W` which consists of morphisms `g : Z ⟶ W` such that there exists `j` |
| 70 | +and `h : Z ⟶ E.Y j` such that `g ≫ p₁ = h ≫ E.p₁ j` and `g ≫ p₂ = h ≫ E.p₂ j`. |
| 71 | +See lemmas `sieve₁_eq_pullback_sieve₁'` and `sieve₁'_eq_sieve₁` for equational lemmas |
| 72 | +regarding this sieve. -/ |
| 73 | +@[simps] |
| 74 | +def sieve₁ {i₁ i₂ : E.I₀} {W : C} (p₁ : W ⟶ E.X i₁) (p₂ : W ⟶ E.X i₂) : Sieve W where |
| 75 | + arrows Z g := ∃ (j : E.I₁ i₁ i₂) (h : Z ⟶ E.Y j), g ≫ p₁ = h ≫ E.p₁ j ∧ g ≫ p₂ = h ≫ E.p₂ j |
| 76 | + downward_closed := by |
| 77 | + rintro Z Z' g ⟨j, h, fac₁, fac₂⟩ φ |
| 78 | + exact ⟨j, φ ≫ h, by simpa using φ ≫= fac₁, by simpa using φ ≫= fac₂⟩ |
| 79 | + |
| 80 | +section |
| 81 | + |
| 82 | +variable {i₁ i₂ : E.I₀} [HasPullback (E.f i₁) (E.f i₂)] |
| 83 | + |
| 84 | +/-- The obvious morphism `E.Y j ⟶ pullback (E.f i₁) (E.f i₂)` given by `E : PreOneHypercover S`. -/ |
| 85 | +noncomputable abbrev toPullback (j : E.I₁ i₁ i₂) [HasPullback (E.f i₁) (E.f i₂)] : |
| 86 | + E.Y j ⟶ pullback (E.f i₁) (E.f i₂) := |
| 87 | + pullback.lift (E.p₁ j) (E.p₂ j) (E.w j) |
| 88 | + |
| 89 | +variable (i₁ i₂) in |
| 90 | +/-- The sieve of `pullback (E.f i₁) (E.f i₂)` given by `E : PreOneHypercover S`. -/ |
| 91 | +def sieve₁' : Sieve (pullback (E.f i₁) (E.f i₂)) := |
| 92 | + Sieve.ofArrows _ (fun (j : E.I₁ i₁ i₂) => E.toPullback j) |
| 93 | + |
| 94 | +lemma sieve₁_eq_pullback_sieve₁' {W : C} (p₁ : W ⟶ E.X i₁) (p₂ : W ⟶ E.X i₂) |
| 95 | + (w : p₁ ≫ E.f i₁ = p₂ ≫ E.f i₂) : |
| 96 | + E.sieve₁ p₁ p₂ = (E.sieve₁' i₁ i₂).pullback (pullback.lift _ _ w) := by |
| 97 | + ext Z g |
| 98 | + constructor |
| 99 | + · rintro ⟨j, h, fac₁, fac₂⟩ |
| 100 | + exact ⟨_, h, _, ⟨j⟩, by aesop_cat⟩ |
| 101 | + · rintro ⟨_, h, w, ⟨j⟩, fac⟩ |
| 102 | + exact ⟨j, h, by simpa using fac.symm =≫ pullback.fst, |
| 103 | + by simpa using fac.symm =≫ pullback.snd⟩ |
| 104 | + |
| 105 | +variable (i₁ i₂) in |
| 106 | +lemma sieve₁'_eq_sieve₁ : E.sieve₁' i₁ i₂ = E.sieve₁ pullback.fst pullback.snd := by |
| 107 | + rw [← Sieve.pullback_id (S := E.sieve₁' i₁ i₂), |
| 108 | + sieve₁_eq_pullback_sieve₁' _ _ _ pullback.condition] |
| 109 | + congr |
| 110 | + aesop_cat |
| 111 | + |
| 112 | +end |
| 113 | + |
| 114 | +/-- The sigma type of all `E.I₁ i₁ i₂` for `⟨i₁, i₂⟩ : E.I₀ × E.I₀`. -/ |
| 115 | +abbrev I₁' : Type w := Sigma (fun (i : E.I₀ × E.I₀) => E.I₁ i.1 i.2) |
| 116 | + |
| 117 | +/-- The diagram of the multifork attached to a presheaf |
| 118 | +`F : Cᵒᵖ ⥤ A`, `S : C` and `E : PreOneHypercover S`. -/ |
| 119 | +@[simps] |
| 120 | +def multicospanIndex (F : Cᵒᵖ ⥤ A) : MulticospanIndex A where |
| 121 | + L := E.I₀ |
| 122 | + R := E.I₁' |
| 123 | + fstTo j := j.1.1 |
| 124 | + sndTo j := j.1.2 |
| 125 | + left i := F.obj (Opposite.op (E.X i)) |
| 126 | + right j := F.obj (Opposite.op (E.Y j.2)) |
| 127 | + fst j := F.map ((E.p₁ j.2).op) |
| 128 | + snd j := F.map ((E.p₂ j.2).op) |
| 129 | + |
| 130 | +/-- The multifork attached to a presheaf `F : Cᵒᵖ ⥤ A`, `S : C` and `E : PreOneHypercover S`. -/ |
| 131 | +def multifork (F : Cᵒᵖ ⥤ A) : |
| 132 | + Multifork (E.multicospanIndex F) := |
| 133 | + Multifork.ofι _ (F.obj (Opposite.op S)) (fun i₀ => F.map (E.f i₀).op) (by |
| 134 | + rintro ⟨⟨i₁, i₂⟩, (j : E.I₁ i₁ i₂)⟩ |
| 135 | + dsimp |
| 136 | + simp only [← F.map_comp, ← op_comp, E.w]) |
| 137 | + |
| 138 | +end PreOneHypercover |
| 139 | + |
| 140 | +namespace GrothendieckTopology |
| 141 | + |
| 142 | +variable (J : GrothendieckTopology C) |
| 143 | + |
| 144 | +/-- The type of `1`-hypercovers of an object `S : C` in a category equipped with a |
| 145 | +Grothendieck topology `J`. This can be constructed from a covering of `S` and |
| 146 | +a covering of the fibre products of the objects in this covering (see `OneHypercover.mk'`). -/ |
| 147 | +structure OneHypercover (S : C) extends PreOneHypercover S where |
| 148 | + mem₀ : toPreOneHypercover.sieve₀ ∈ J S |
| 149 | + mem₁ (i₁ i₂ : I₀) ⦃W : C⦄ (p₁ : W ⟶ X i₁) (p₂ : W ⟶ X i₂) (w : p₁ ≫ f i₁ = p₂ ≫ f i₂) : |
| 150 | + toPreOneHypercover.sieve₁ p₁ p₂ ∈ J W |
| 151 | + |
| 152 | +variable {J} |
| 153 | + |
| 154 | +lemma OneHypercover.mem_sieve₁' {S : C} (E : J.OneHypercover S) |
| 155 | + (i₁ i₂ : E.I₀) [HasPullback (E.f i₁) (E.f i₂)] : |
| 156 | + E.sieve₁' i₁ i₂ ∈ J _ := by |
| 157 | + rw [E.sieve₁'_eq_sieve₁] |
| 158 | + exact mem₁ _ _ _ _ _ pullback.condition |
| 159 | + |
| 160 | +namespace OneHypercover |
| 161 | + |
| 162 | +/-- In order to check that a certain data is a `1`-hypercover of `S`, it suffices to |
| 163 | +check that the data provides a covering of `S` and of the fibre products. -/ |
| 164 | +@[simps toPreOneHypercover] |
| 165 | +def mk' {S : C} (E : PreOneHypercover S) [E.HasPullbacks] |
| 166 | + (mem₀ : E.sieve₀ ∈ J S) (mem₁' : ∀ (i₁ i₂ : E.I₀), E.sieve₁' i₁ i₂ ∈ J _) : |
| 167 | + J.OneHypercover S where |
| 168 | + toPreOneHypercover := E |
| 169 | + mem₀ := mem₀ |
| 170 | + mem₁ i₁ i₂ W p₁ p₂ w := by |
| 171 | + rw [E.sieve₁_eq_pullback_sieve₁' _ _ w] |
| 172 | + exact J.pullback_stable' _ (mem₁' i₁ i₂) |
| 173 | + |
| 174 | +section |
| 175 | + |
| 176 | +variable {S : C} (E : J.OneHypercover S) (F : Sheaf J A) |
| 177 | + |
| 178 | +section |
| 179 | + |
| 180 | +variable {E F} |
| 181 | +variable (c : Multifork (E.multicospanIndex F.val)) |
| 182 | + |
| 183 | +/-- Auxiliary definition of `isLimitMultifork`. -/ |
| 184 | +noncomputable def multiforkLift : c.pt ⟶ F.val.obj (Opposite.op S) := |
| 185 | + F.cond.amalgamateOfArrows _ E.mem₀ c.ι (fun W i₁ i₂ p₁ p₂ w => by |
| 186 | + apply F.cond.hom_ext ⟨_, E.mem₁ _ _ _ _ w⟩ |
| 187 | + rintro ⟨T, g, j, h, fac₁, fac₂⟩ |
| 188 | + dsimp |
| 189 | + simp only [assoc, ← Functor.map_comp, ← op_comp, fac₁, fac₂] |
| 190 | + simp only [op_comp, Functor.map_comp] |
| 191 | + simpa using c.condition ⟨⟨i₁, i₂⟩, j⟩ =≫ F.val.map h.op) |
| 192 | + |
| 193 | +@[reassoc] |
| 194 | +lemma multiforkLift_map (i₀ : E.I₀) : multiforkLift c ≫ F.val.map (E.f i₀).op = c.ι i₀ := by |
| 195 | + simp [multiforkLift] |
| 196 | + |
| 197 | +end |
| 198 | + |
| 199 | +/-- If `E : J.OneHypercover S` and `F : Sheaf J A`, then `F.obj (op S)` is |
| 200 | +a multiequalizer of suitable maps `F.obj (op (E.X i)) ⟶ F.obj (op (E.Y j))` |
| 201 | +induced by `E.p₁ j` and `E.p₂ j`. -/ |
| 202 | +noncomputable def isLimitMultifork : IsLimit (E.multifork F.1) := |
| 203 | + Multifork.IsLimit.mk _ (fun c => multiforkLift c) (fun c => multiforkLift_map c) (by |
| 204 | + intro c m hm |
| 205 | + apply F.cond.hom_ext_ofArrows _ E.mem₀ |
| 206 | + intro i₀ |
| 207 | + dsimp only |
| 208 | + rw [multiforkLift_map] |
| 209 | + exact hm i₀) |
| 210 | + |
| 211 | +end |
| 212 | + |
| 213 | +end OneHypercover |
| 214 | + |
| 215 | +end GrothendieckTopology |
| 216 | + |
| 217 | +end CategoryTheory |
0 commit comments