|
| 1 | +/- |
| 2 | +Copyright (c) 2021 Andrew Yang. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Andrew Yang |
| 5 | +-/ |
| 6 | +import tactic.elementwise |
| 7 | +import category_theory.limits.shapes.multiequalizer |
| 8 | +import category_theory.limits.constructions.epi_mono |
| 9 | + |
| 10 | +/-! |
| 11 | +# Gluing data |
| 12 | +
|
| 13 | +We define `glue_data` as a family of data needed to glue topological spaces, schemes, etc. We |
| 14 | +provide the API to realize it as a multispan diagram, and also states lemmas about its |
| 15 | +interaction with a functor that preserves certain pullbacks. |
| 16 | +
|
| 17 | +-/ |
| 18 | + |
| 19 | +noncomputable theory |
| 20 | + |
| 21 | +open category_theory.limits |
| 22 | +namespace category_theory |
| 23 | + |
| 24 | +universes v u₁ u₂ |
| 25 | + |
| 26 | +variables (C : Type u₁) [category.{v} C] {C' : Type u₂} [category.{v} C'] |
| 27 | + |
| 28 | +/-- |
| 29 | +A gluing datum consists of |
| 30 | +1. An index type `J` |
| 31 | +2. An object `U i` for each `i : J`. |
| 32 | +3. An object `V i j` for each `i j : J`. |
| 33 | +4. A monomorphism `f i j : V i j ⟶ U i` for each `i j : J`. |
| 34 | +5. A transition map `t i j : V i j ⟶ V j i` for each `i j : J`. |
| 35 | +such that |
| 36 | +6. `f i i` is an isomorphism. |
| 37 | +7. `t i i` is the identity. |
| 38 | +8. The pullback for `f i j` and `f i k` exists. |
| 39 | +9. `V i j ×[U i] V i k ⟶ V i j ⟶ V j i` factors through `V j k ×[U j] V j i ⟶ V j i` via some |
| 40 | + `t' : V i j ×[U i] V i k ⟶ V j k ×[U j] V j i`. |
| 41 | +10. `t' i j k ≫ t' j k i ≫ t' k i j = 𝟙 _`. |
| 42 | +-/ |
| 43 | +@[nolint has_inhabited_instance] |
| 44 | +structure glue_data := |
| 45 | +(J : Type v) |
| 46 | +(U : J → C) |
| 47 | +(V : J × J → C) |
| 48 | +(f : Π i j, V (i, j) ⟶ U i) |
| 49 | +(f_mono : ∀ i j, mono (f i j) . tactic.apply_instance) |
| 50 | +(f_has_pullback : ∀ i j k, has_pullback (f i j) (f i k) . tactic.apply_instance) |
| 51 | +(f_id : ∀ i, is_iso (f i i) . tactic.apply_instance) |
| 52 | +(t : Π i j, V (i, j) ⟶ V (j, i)) |
| 53 | +(t_id : ∀ i, t i i = 𝟙 _) |
| 54 | +(t' : Π i j k, pullback (f i j) (f i k) ⟶ pullback (f j k) (f j i)) |
| 55 | +(t_fac : ∀ i j k, t' i j k ≫ pullback.snd = pullback.fst ≫ t i j) |
| 56 | +(cocycle : ∀ i j k , t' i j k ≫ t' j k i ≫ t' k i j = 𝟙 _) |
| 57 | + |
| 58 | +attribute [simp] glue_data.t_id |
| 59 | +attribute [instance] glue_data.f_id glue_data.f_mono glue_data.f_has_pullback |
| 60 | +attribute [reassoc] glue_data.t_fac glue_data.cocycle |
| 61 | + |
| 62 | +namespace glue_data |
| 63 | + |
| 64 | +variables {C} (D : glue_data C) |
| 65 | + |
| 66 | +@[simp] lemma t'_iij (i j : D.J) : D.t' i i j = (pullback_symmetry _ _).hom := |
| 67 | +begin |
| 68 | + have eq₁ := D.t_fac i i j, |
| 69 | + have eq₂ := (is_iso.eq_comp_inv (D.f i i)).mpr (@pullback.condition _ _ _ _ _ _ (D.f i j) _), |
| 70 | + rw [D.t_id, category.comp_id, eq₂] at eq₁, |
| 71 | + have eq₃ := (is_iso.eq_comp_inv (D.f i i)).mp eq₁, |
| 72 | + rw [category.assoc, ←pullback.condition, ←category.assoc] at eq₃, |
| 73 | + exact mono.right_cancellation _ _ |
| 74 | + ((mono.right_cancellation _ _ eq₃).trans (pullback_symmetry_hom_comp_fst _ _).symm) |
| 75 | +end |
| 76 | + |
| 77 | +lemma t'_jii (i j : D.J) : D.t' j i i = pullback.fst ≫ D.t j i ≫ inv pullback.snd := |
| 78 | +by { rw [←category.assoc, ←D.t_fac], simp } |
| 79 | + |
| 80 | +lemma t'_iji (i j : D.J) : D.t' i j i = pullback.fst ≫ D.t i j ≫ inv pullback.snd := |
| 81 | +by { rw [←category.assoc, ←D.t_fac], simp } |
| 82 | + |
| 83 | +@[simp, reassoc, elementwise] lemma t_inv (i j : D.J) : |
| 84 | + D.t i j ≫ D.t j i = 𝟙 _ := |
| 85 | +begin |
| 86 | + have eq : (pullback_symmetry (D.f i i) (D.f i j)).hom = pullback.snd ≫ inv pullback.fst, |
| 87 | + { simp }, |
| 88 | + have := D.cocycle i j i, |
| 89 | + rw [D.t'_iij, D.t'_jii, D.t'_iji, fst_eq_snd_of_mono_eq, eq] at this, |
| 90 | + simp only [category.assoc, is_iso.inv_hom_id_assoc] at this, |
| 91 | + rw [←is_iso.eq_inv_comp, ←category.assoc, is_iso.comp_inv_eq] at this, |
| 92 | + simpa using this, |
| 93 | +end |
| 94 | + |
| 95 | +lemma t'_inv (i j k : D.J) : D.t' i j k ≫ (pullback_symmetry _ _).hom ≫ |
| 96 | + D.t' j i k ≫ (pullback_symmetry _ _).hom = 𝟙 _ := |
| 97 | +begin |
| 98 | + rw ← cancel_mono (pullback.fst : pullback (D.f i j) (D.f i k) ⟶ _), |
| 99 | + simp [t_fac, t_fac_assoc] |
| 100 | +end |
| 101 | + |
| 102 | +instance t_is_iso (i j : D.J) : is_iso (D.t i j) := |
| 103 | +⟨⟨D.t j i, D.t_inv _ _, D.t_inv _ _⟩⟩ |
| 104 | + |
| 105 | +instance t'_is_iso (i j k : D.J) : is_iso (D.t' i j k) := |
| 106 | +⟨⟨D.t' j k i ≫ D.t' k i j, D.cocycle _ _ _, (by simpa using D.cocycle _ _ _)⟩⟩ |
| 107 | + |
| 108 | +@[reassoc] |
| 109 | +lemma t'_comp_eq_pullback_symmetry (i j k : D.J) : |
| 110 | + D.t' j k i ≫ D.t' k i j = (pullback_symmetry _ _).hom ≫ |
| 111 | + D.t' j i k ≫ (pullback_symmetry _ _).hom := |
| 112 | +begin |
| 113 | + transitivity inv (D.t' i j k), |
| 114 | + { exact is_iso.eq_inv_of_hom_inv_id (D.cocycle _ _ _) }, |
| 115 | + { rw ← cancel_mono (pullback.fst : pullback (D.f i j) (D.f i k) ⟶ _), |
| 116 | + simp [t_fac, t_fac_assoc] } |
| 117 | +end |
| 118 | + |
| 119 | +/-- (Implementation) The disjoint union of `U i`. -/ |
| 120 | +def sigma_opens [has_coproduct D.U] : C := ∐ D.U |
| 121 | + |
| 122 | +/-- (Implementation) The diagram to take colimit of. -/ |
| 123 | +def diagram : multispan_index C := |
| 124 | +{ L := D.J × D.J, R := D.J, |
| 125 | + fst_from := _root_.prod.fst, snd_from := _root_.prod.snd, |
| 126 | + left := D.V, right := D.U, |
| 127 | + fst := λ ⟨i, j⟩, D.f i j, |
| 128 | + snd := λ ⟨i, j⟩, D.t i j ≫ D.f j i } |
| 129 | + |
| 130 | +@[simp] lemma diagram_L : D.diagram.L = (D.J × D.J) := rfl |
| 131 | +@[simp] lemma diagram_R : D.diagram.R = D.J := rfl |
| 132 | +@[simp] lemma diagram_fst_from (i j : D.J) : D.diagram.fst_from ⟨i, j⟩ = i := rfl |
| 133 | +@[simp] lemma diagram_snd_from (i j : D.J) : D.diagram.snd_from ⟨i, j⟩ = j := rfl |
| 134 | +@[simp] lemma diagram_fst (i j : D.J) : D.diagram.fst ⟨i, j⟩ = D.f i j := rfl |
| 135 | +@[simp] lemma diagram_snd (i j : D.J) : D.diagram.snd ⟨i, j⟩ = D.t i j ≫ D.f j i := rfl |
| 136 | +@[simp] lemma diagram_left : D.diagram.left = D.V := rfl |
| 137 | +@[simp] lemma diagram_right : D.diagram.right = D.U := rfl |
| 138 | + |
| 139 | +section |
| 140 | + |
| 141 | +variable [has_multicoequalizer D.diagram] |
| 142 | + |
| 143 | +/-- The glued object given a family of gluing data. -/ |
| 144 | +def glued : C := multicoequalizer D.diagram |
| 145 | + |
| 146 | +/-- The map `D.U i ⟶ D.glued` for each `i`. -/ |
| 147 | +def ι (i : D.J) : D.U i ⟶ D.glued := |
| 148 | +multicoequalizer.π D.diagram i |
| 149 | + |
| 150 | +@[simp, elementwise] |
| 151 | +lemma glue_condition (i j : D.J) : |
| 152 | + D.t i j ≫ D.f j i ≫ D.ι j = D.f i j ≫ D.ι i := |
| 153 | +(category.assoc _ _ _).symm.trans (multicoequalizer.condition D.diagram ⟨i, j⟩).symm |
| 154 | + |
| 155 | +variables [has_colimits C] |
| 156 | + |
| 157 | +/-- The projection `∐ D.U ⟶ D.glued` given by the colimit. -/ |
| 158 | +def π : D.sigma_opens ⟶ D.glued := multicoequalizer.sigma_π D.diagram |
| 159 | + |
| 160 | +instance π_epi : epi D.π := by { unfold π, apply_instance } |
| 161 | + |
| 162 | +end |
| 163 | + |
| 164 | +variables (F : C ⥤ C') [H : ∀ i j k, preserves_limit (cospan (D.f i j) (D.f i k)) F] |
| 165 | + |
| 166 | +include H |
| 167 | + |
| 168 | +instance (i j k : D.J) : has_pullback (F.map (D.f i j)) (F.map (D.f i k)) := |
| 169 | +⟨⟨⟨_, is_limit_of_has_pullback_of_preserves_limit F (D.f i j) (D.f i k)⟩⟩⟩ |
| 170 | + |
| 171 | +/-- A functor that preserves the pullbacks of `f i j` and `f i k` can map a family of glue data. -/ |
| 172 | +@[simps] def map_glue_data : |
| 173 | + glue_data C' := |
| 174 | +{ J := D.J, |
| 175 | + U := λ i, F.obj (D.U i), |
| 176 | + V := λ i, F.obj (D.V i), |
| 177 | + f := λ i j, F.map (D.f i j), |
| 178 | + f_mono := λ i j, category_theory.preserves_mono F (D.f i j), |
| 179 | + f_id := λ i, infer_instance, |
| 180 | + t := λ i j, F.map (D.t i j), |
| 181 | + t_id := λ i, by { rw D.t_id i, simp }, |
| 182 | + t' := λ i j k, (preserves_pullback.iso F (D.f i j) (D.f i k)).inv ≫ |
| 183 | + F.map (D.t' i j k) ≫ (preserves_pullback.iso F (D.f j k) (D.f j i)).hom, |
| 184 | + t_fac := λ i j k, by simpa [iso.inv_comp_eq] using congr_arg (λ f, F.map f) (D.t_fac i j k), |
| 185 | + cocycle := λ i j k, by simp only [category.assoc, iso.hom_inv_id_assoc, ← functor.map_comp_assoc, |
| 186 | + D.cocycle, iso.inv_hom_id, category_theory.functor.map_id, category.id_comp] } |
| 187 | + |
| 188 | +/-- |
| 189 | +The diagram of the image of a `glue_data` under a functor `F` is naturally isomorphic to the |
| 190 | +original diagram of the `glue_data` via `F`. |
| 191 | +-/ |
| 192 | +def diagram_iso : D.diagram.multispan ⋙ F ≅ (D.map_glue_data F).diagram.multispan := |
| 193 | +nat_iso.of_components |
| 194 | + (λ x, match x with |
| 195 | + | walking_multispan.left a := iso.refl _ |
| 196 | + | walking_multispan.right b := iso.refl _ |
| 197 | + end) |
| 198 | + (begin |
| 199 | + rintros (⟨_,_⟩|_) _ (_|_|_), |
| 200 | + { erw [category.comp_id, category.id_comp, functor.map_id], refl }, |
| 201 | + { erw [category.comp_id, category.id_comp], refl }, |
| 202 | + { erw [category.comp_id, category.id_comp, functor.map_comp], refl }, |
| 203 | + { erw [category.comp_id, category.id_comp, functor.map_id], refl }, |
| 204 | + end) |
| 205 | + |
| 206 | +@[simp] lemma diagram_iso_app_left (i : D.J × D.J) : |
| 207 | + (D.diagram_iso F).app (walking_multispan.left i) = iso.refl _ := rfl |
| 208 | + |
| 209 | +@[simp] lemma diagram_iso_app_right (i : D.J) : |
| 210 | + (D.diagram_iso F).app (walking_multispan.right i) = iso.refl _ := rfl |
| 211 | + |
| 212 | +@[simp] lemma diagram_iso_hom_app_left (i : D.J × D.J) : |
| 213 | + (D.diagram_iso F).hom.app (walking_multispan.left i) = 𝟙 _ := rfl |
| 214 | + |
| 215 | +@[simp] lemma diagram_iso_hom_app_right (i : D.J) : |
| 216 | + (D.diagram_iso F).hom.app (walking_multispan.right i) = 𝟙 _ := rfl |
| 217 | + |
| 218 | +@[simp] lemma diagram_iso_inv_app_left (i : D.J × D.J) : |
| 219 | + (D.diagram_iso F).inv.app (walking_multispan.left i) = 𝟙 _ := rfl |
| 220 | + |
| 221 | +@[simp] lemma diagram_iso_inv_app_right (i : D.J) : |
| 222 | + (D.diagram_iso F).inv.app (walking_multispan.right i) = 𝟙 _ := rfl |
| 223 | + |
| 224 | +end glue_data |
| 225 | + |
| 226 | +end category_theory |
0 commit comments