| author | nipkow | 
| Thu, 23 Feb 2023 22:04:32 +0100 | |
| changeset 77356 | 1f5428d66591 | 
| parent 73932 | fd21b4a93043 | 
| child 78322 | 74c75da4cb01 | 
| permissions | -rw-r--r-- | 
| 70095 | 1 | section\<open>Homology, III: Brouwer Degree\<close> | 
| 2 | ||
| 3 | theory Brouwer_Degree | |
| 72632 
773ad766f1b8
Multiplicative_Group now required due to Algebra restructuring
 paulson <lp15@cam.ac.uk> parents: 
70097diff
changeset | 4 | imports Homology_Groups "HOL-Algebra.Multiplicative_Group" | 
| 70095 | 5 | |
| 6 | begin | |
| 7 | ||
| 8 | subsection\<open>Reduced Homology\<close> | |
| 9 | ||
| 10 | definition reduced_homology_group :: "int \<Rightarrow> 'a topology \<Rightarrow> 'a chain set monoid" | |
| 11 | where "reduced_homology_group p X \<equiv> | |
| 12 | subgroup_generated (homology_group p X) | |
| 13 |              (kernel (homology_group p X) (homology_group p (discrete_topology {()}))
 | |
| 14 |                      (hom_induced p X {} (discrete_topology {()}) {} (\<lambda>x. ())))"
 | |
| 15 | ||
| 16 | lemma one_reduced_homology_group: "\<one>\<^bsub>reduced_homology_group p X\<^esub> = \<one>\<^bsub>homology_group p X\<^esub>" | |
| 17 | by (simp add: reduced_homology_group_def) | |
| 18 | ||
| 19 | lemma group_reduced_homology_group [simp]: "group (reduced_homology_group p X)" | |
| 20 | by (simp add: reduced_homology_group_def group.group_subgroup_generated) | |
| 21 | ||
| 22 | lemma carrier_reduced_homology_group: | |
| 23 | "carrier (reduced_homology_group p X) = | |
| 24 |     kernel (homology_group p X) (homology_group p (discrete_topology {()}))
 | |
| 25 |            (hom_induced p X {} (discrete_topology {()}) {} (\<lambda>x. ()))"
 | |
| 26 | (is "_ = kernel ?G ?H ?h") | |
| 27 | proof - | |
| 28 | interpret subgroup "kernel ?G ?H ?h" ?G | |
| 29 | by (simp add: hom_induced_empty_hom group_hom_axioms_def group_hom_def group_hom.subgroup_kernel) | |
| 30 | show ?thesis | |
| 31 | unfolding reduced_homology_group_def | |
| 32 | using carrier_subgroup_generated_subgroup by blast | |
| 33 | qed | |
| 34 | ||
| 35 | lemma carrier_reduced_homology_group_subset: | |
| 36 | "carrier (reduced_homology_group p X) \<subseteq> carrier (homology_group p X)" | |
| 37 | by (simp add: group.carrier_subgroup_generated_subset reduced_homology_group_def) | |
| 38 | ||
| 39 | lemma un_reduced_homology_group: | |
| 40 | assumes "p \<noteq> 0" | |
| 41 | shows "reduced_homology_group p X = homology_group p X" | |
| 42 | proof - | |
| 43 |   have "(kernel (homology_group p X) (homology_group p (discrete_topology {()}))
 | |
| 44 |               (hom_induced p X {} (discrete_topology {()}) {} (\<lambda>x. ())))
 | |
| 45 | = carrier (homology_group p X)" | |
| 46 | proof (rule group_hom.kernel_to_trivial_group) | |
| 47 |     show "group_hom (homology_group p X) (homology_group p (discrete_topology {()}))
 | |
| 48 |          (hom_induced p X {} (discrete_topology {()}) {} (\<lambda>x. ()))"
 | |
| 49 | by (auto simp: hom_induced_empty_hom group_hom_def group_hom_axioms_def) | |
| 50 |     show "trivial_group (homology_group p (discrete_topology {()}))"
 | |
| 51 | by (simp add: homology_dimension_axiom [OF _ assms]) | |
| 52 | qed | |
| 53 | then show ?thesis | |
| 54 | by (simp add: reduced_homology_group_def group.subgroup_generated_group_carrier) | |
| 55 | qed | |
| 56 | ||
| 57 | lemma trivial_reduced_homology_group: | |
| 58 | "p < 0 \<Longrightarrow> trivial_group(reduced_homology_group p X)" | |
| 59 | by (simp add: trivial_homology_group un_reduced_homology_group) | |
| 60 | ||
| 61 | lemma hom_induced_reduced_hom: | |
| 62 |    "(hom_induced p X {} Y {} f) \<in> hom (reduced_homology_group p X) (reduced_homology_group p Y)"
 | |
| 63 | proof (cases "continuous_map X Y f") | |
| 64 | case True | |
| 65 | have eq: "continuous_map X Y f | |
| 66 |          \<Longrightarrow> hom_induced p X {} (discrete_topology {()}) {} (\<lambda>x. ())
 | |
| 67 |            = (hom_induced p Y {} (discrete_topology {()}) {} (\<lambda>x. ()) \<circ> hom_induced p X {} Y {} f)"
 | |
| 68 | by (simp flip: hom_induced_compose_empty) | |
| 69 | interpret subgroup "kernel (homology_group p X) | |
| 70 |                        (homology_group p (discrete_topology {()}))
 | |
| 71 |                          (hom_induced p X {} (discrete_topology {()}) {} (\<lambda>x. ()))"
 | |
| 72 | "homology_group p X" | |
| 73 | by (meson group_hom.subgroup_kernel group_hom_axioms_def group_hom_def group_relative_homology_group hom_induced) | |
| 74 |   have sb: "hom_induced p X {} Y {} f ` carrier (homology_group p X) \<subseteq> carrier (homology_group p Y)"
 | |
| 75 | using hom_induced_carrier by blast | |
| 76 | show ?thesis | |
| 77 | using True | |
| 78 | unfolding reduced_homology_group_def | |
| 79 | apply (simp add: hom_into_subgroup_eq group_hom.subgroup_kernel hom_induced_empty_hom group.hom_from_subgroup_generated group_hom_def group_hom_axioms_def) | |
| 80 | unfolding kernel_def using eq sb by auto | |
| 81 | next | |
| 82 | case False | |
| 83 |   then have "hom_induced p X {} Y {} f = (\<lambda>c. one(reduced_homology_group p Y))"
 | |
| 84 | by (force simp: hom_induced_default reduced_homology_group_def) | |
| 85 | then show ?thesis | |
| 86 | by (simp add: trivial_hom) | |
| 87 | qed | |
| 88 | ||
| 89 | ||
| 90 | lemma hom_induced_reduced: | |
| 91 | "c \<in> carrier(reduced_homology_group p X) | |
| 92 |         \<Longrightarrow> hom_induced p X {} Y {} f c \<in> carrier(reduced_homology_group p Y)"
 | |
| 93 | by (meson hom_in_carrier hom_induced_reduced_hom) | |
| 94 | ||
| 95 | lemma hom_boundary_reduced_hom: | |
| 96 | "hom_boundary p X S | |
| 97 | \<in> hom (relative_homology_group p X S) (reduced_homology_group (p-1) (subtopology X S))" | |
| 98 | proof - | |
| 99 |   have *: "continuous_map X (discrete_topology {()}) (\<lambda>x. ())" "(\<lambda>x. ()) ` S \<subseteq> {()}"
 | |
| 100 | by auto | |
| 101 |   interpret group_hom "relative_homology_group p (discrete_topology {()}) {()}"
 | |
| 102 |                       "homology_group (p-1) (discrete_topology {()})"
 | |
| 103 |                       "hom_boundary p (discrete_topology {()}) {()}"
 | |
| 104 | apply (clarsimp simp: group_hom_def group_hom_axioms_def) | |
| 105 | by (metis UNIV_unit hom_boundary_hom subtopology_UNIV) | |
| 106 | have "hom_boundary p X S ` | |
| 107 | carrier (relative_homology_group p X S) | |
| 108 | \<subseteq> kernel (homology_group (p - 1) (subtopology X S)) | |
| 109 |             (homology_group (p - 1) (discrete_topology {()}))
 | |
| 110 |             (hom_induced (p - 1) (subtopology X S) {}
 | |
| 111 |               (discrete_topology {()}) {} (\<lambda>x. ()))"
 | |
| 112 | proof (clarsimp simp add: kernel_def hom_boundary_carrier) | |
| 113 | fix c | |
| 114 | assume c: "c \<in> carrier (relative_homology_group p X S)" | |
| 115 |     have triv: "trivial_group (relative_homology_group p (discrete_topology {()}) {()})"
 | |
| 116 | by (metis topspace_discrete_topology trivial_relative_homology_group_topspace) | |
| 117 |     have "hom_boundary p (discrete_topology {()}) {()}
 | |
| 118 |          (hom_induced p X S (discrete_topology {()}) {()} (\<lambda>x. ()) c)
 | |
| 119 |        = \<one>\<^bsub>homology_group (p - 1) (discrete_topology {()})\<^esub>"
 | |
| 120 | by (metis hom_induced_carrier local.hom_one singletonD triv trivial_group_def) | |
| 121 |     then show "hom_induced (p - 1) (subtopology X S) {} (discrete_topology {()}) {} (\<lambda>x. ()) (hom_boundary p X S c) =
 | |
| 122 |         \<one>\<^bsub>homology_group (p - 1) (discrete_topology {()})\<^esub>"
 | |
| 123 | using naturality_hom_induced [OF *, of p, symmetric] by (simp add: o_def fun_eq_iff) | |
| 124 | qed | |
| 125 | then show ?thesis | |
| 126 | by (simp add: reduced_homology_group_def hom_boundary_hom hom_into_subgroup) | |
| 127 | qed | |
| 128 | ||
| 129 | ||
| 130 | lemma homotopy_equivalence_reduced_homology_group_isomorphisms: | |
| 131 | assumes contf: "continuous_map X Y f" and contg: "continuous_map Y X g" | |
| 132 | and gf: "homotopic_with (\<lambda>h. True) X X (g \<circ> f) id" | |
| 133 | and fg: "homotopic_with (\<lambda>k. True) Y Y (f \<circ> g) id" | |
| 134 | shows "group_isomorphisms (reduced_homology_group p X) (reduced_homology_group p Y) | |
| 135 |                                (hom_induced p X {} Y {} f) (hom_induced p Y {} X {} g)"
 | |
| 136 | proof (simp add: hom_induced_reduced_hom group_isomorphisms_def, intro conjI ballI) | |
| 137 | fix a | |
| 138 | assume "a \<in> carrier (reduced_homology_group p X)" | |
| 139 |   then have "(hom_induced p Y {} X {} g \<circ> hom_induced p X {} Y {} f) a = a"
 | |
| 140 | apply (simp add: contf contg flip: hom_induced_compose) | |
| 141 | using carrier_reduced_homology_group_subset gf hom_induced_id homology_homotopy_empty by fastforce | |
| 142 |   then show "hom_induced p Y {} X {} g (hom_induced p X {} Y {} f a) = a"
 | |
| 143 | by simp | |
| 144 | next | |
| 145 | fix b | |
| 146 | assume "b \<in> carrier (reduced_homology_group p Y)" | |
| 147 |   then have "(hom_induced p X {} Y {} f \<circ> hom_induced p Y {} X {} g) b = b"
 | |
| 148 | apply (simp add: contf contg flip: hom_induced_compose) | |
| 149 | using carrier_reduced_homology_group_subset fg hom_induced_id homology_homotopy_empty by fastforce | |
| 150 |   then show "hom_induced p X {} Y {} f (hom_induced p Y {} X {} g b) = b"
 | |
| 151 | by (simp add: carrier_reduced_homology_group) | |
| 152 | qed | |
| 153 | ||
| 154 | lemma homotopy_equivalence_reduced_homology_group_isomorphism: | |
| 155 | assumes "continuous_map X Y f" "continuous_map Y X g" | |
| 156 | and "homotopic_with (\<lambda>h. True) X X (g \<circ> f) id" "homotopic_with (\<lambda>k. True) Y Y (f \<circ> g) id" | |
| 157 |   shows "(hom_induced p X {} Y {} f)
 | |
| 158 | \<in> iso (reduced_homology_group p X) (reduced_homology_group p Y)" | |
| 159 | proof (rule group_isomorphisms_imp_iso) | |
| 160 | show "group_isomorphisms (reduced_homology_group p X) (reduced_homology_group p Y) | |
| 161 |          (hom_induced p X {} Y {} f) (hom_induced p Y {} X {} g)"
 | |
| 162 | by (simp add: assms homotopy_equivalence_reduced_homology_group_isomorphisms) | |
| 163 | qed | |
| 164 | ||
| 165 | lemma homotopy_equivalent_space_imp_isomorphic_reduced_homology_groups: | |
| 166 | "X homotopy_equivalent_space Y | |
| 167 | \<Longrightarrow> reduced_homology_group p X \<cong> reduced_homology_group p Y" | |
| 168 | unfolding homotopy_equivalent_space_def | |
| 169 | using homotopy_equivalence_reduced_homology_group_isomorphism is_isoI by blast | |
| 170 | ||
| 171 | lemma homeomorphic_space_imp_isomorphic_reduced_homology_groups: | |
| 172 | "X homeomorphic_space Y \<Longrightarrow> reduced_homology_group p X \<cong> reduced_homology_group p Y" | |
| 173 | by (simp add: homeomorphic_imp_homotopy_equivalent_space homotopy_equivalent_space_imp_isomorphic_reduced_homology_groups) | |
| 174 | ||
| 175 | lemma trivial_reduced_homology_group_empty: | |
| 176 |    "topspace X = {} \<Longrightarrow> trivial_group(reduced_homology_group p X)"
 | |
| 177 | by (metis carrier_reduced_homology_group_subset group.trivial_group_alt group_reduced_homology_group trivial_group_def trivial_homology_group_empty) | |
| 178 | ||
| 179 | lemma homology_dimension_reduced: | |
| 180 |   assumes "topspace X = {a}"
 | |
| 181 | shows "trivial_group (reduced_homology_group p X)" | |
| 182 | proof - | |
| 183 |   have iso: "(hom_induced p X {} (discrete_topology {()}) {} (\<lambda>x. ()))
 | |
| 184 |            \<in> iso (homology_group p X) (homology_group p (discrete_topology {()}))"
 | |
| 185 | apply (rule homeomorphic_map_homology_iso) | |
| 186 | apply (force simp: homeomorphic_map_maps homeomorphic_maps_def assms) | |
| 187 | done | |
| 188 | show ?thesis | |
| 189 | unfolding reduced_homology_group_def | |
| 190 | by (rule group.trivial_group_subgroup_generated) (use iso in \<open>auto simp: iso_kernel_image\<close>) | |
| 191 | qed | |
| 192 | ||
| 193 | ||
| 194 | lemma trivial_reduced_homology_group_contractible_space: | |
| 195 | "contractible_space X \<Longrightarrow> trivial_group (reduced_homology_group p X)" | |
| 196 | apply (simp add: contractible_eq_homotopy_equivalent_singleton_subtopology) | |
| 197 | apply (auto simp: trivial_reduced_homology_group_empty) | |
| 198 | using isomorphic_group_triviality | |
| 199 | by (metis (full_types) group_reduced_homology_group homology_dimension_reduced homotopy_equivalent_space_imp_isomorphic_reduced_homology_groups path_connectedin_def path_connectedin_singleton topspace_subtopology_subset) | |
| 200 | ||
| 201 | ||
| 202 | lemma image_reduced_homology_group: | |
| 203 |   assumes "topspace X \<inter> S \<noteq> {}"
 | |
| 204 |   shows "hom_induced p X {} X S id ` carrier (reduced_homology_group p X)
 | |
| 205 |        = hom_induced p X {} X S id ` carrier (homology_group p X)"
 | |
| 206 | (is "?h ` carrier ?G = ?h ` carrier ?H") | |
| 207 | proof - | |
| 208 | obtain a where a: "a \<in> topspace X" and "a \<in> S" | |
| 209 | using assms by blast | |
| 210 |   have [simp]: "A \<inter> {x \<in> A. P x} = {x \<in> A. P x}" for A P
 | |
| 211 | by blast | |
| 212 | interpret comm_group "homology_group p X" | |
| 213 | by (rule abelian_relative_homology_group) | |
| 214 | have *: "\<exists>x'. ?h y = ?h x' \<and> | |
| 215 | x' \<in> carrier ?H \<and> | |
| 216 |              hom_induced p X {} (discrete_topology {()}) {} (\<lambda>x. ()) x'
 | |
| 217 |            = \<one>\<^bsub>homology_group p (discrete_topology {()})\<^esub>"
 | |
| 218 | if "y \<in> carrier ?H" for y | |
| 219 | proof - | |
| 220 |     let ?f = "hom_induced p (discrete_topology {()}) {} X {} (\<lambda>x. a)"
 | |
| 221 |     let ?g = "hom_induced p X {} (discrete_topology {()}) {} (\<lambda>x. ())"
 | |
| 222 | have bcarr: "?f (?g y) \<in> carrier ?H" | |
| 223 | by (simp add: hom_induced_carrier) | |
| 224 | interpret gh1: | |
| 225 |       group_hom "relative_homology_group p X S" "relative_homology_group p (discrete_topology {()}) {()}"
 | |
| 226 |                 "hom_induced p X S (discrete_topology {()}) {()} (\<lambda>x. ())"
 | |
| 227 | by (meson group_hom_axioms_def group_hom_def hom_induced_hom group_relative_homology_group) | |
| 228 | interpret gh2: | |
| 229 |       group_hom "relative_homology_group p (discrete_topology {()}) {()}" "relative_homology_group p X S"
 | |
| 230 |                 "hom_induced p (discrete_topology {()}) {()} X S (\<lambda>x. a)"
 | |
| 231 | by (meson group_hom_axioms_def group_hom_def hom_induced_hom group_relative_homology_group) | |
| 232 | interpret gh3: | |
| 233 | group_hom "homology_group p X" "relative_homology_group p X S" "?h" | |
| 234 | by (meson group_hom_axioms_def group_hom_def hom_induced_hom group_relative_homology_group) | |
| 235 | interpret gh4: | |
| 236 |       group_hom "homology_group p X" "homology_group p (discrete_topology {()})"
 | |
| 237 | "?g" | |
| 238 | by (meson group_hom_axioms_def group_hom_def hom_induced_hom group_relative_homology_group) | |
| 239 | interpret gh5: | |
| 240 |       group_hom "homology_group p (discrete_topology {()})" "homology_group p X"
 | |
| 241 | "?f" | |
| 242 | by (meson group_hom_axioms_def group_hom_def hom_induced_hom group_relative_homology_group) | |
| 243 | interpret gh6: | |
| 244 |       group_hom "homology_group p (discrete_topology {()})" "relative_homology_group p (discrete_topology {()}) {()}"
 | |
| 245 |                 "hom_induced p (discrete_topology {()}) {} (discrete_topology {()}) {()} id"
 | |
| 246 | by (meson group_hom_axioms_def group_hom_def hom_induced_hom group_relative_homology_group) | |
| 247 | show ?thesis | |
| 248 | proof (intro exI conjI) | |
| 249 | have "(?h \<circ> ?f \<circ> ?g) y | |
| 250 |           = (hom_induced p (discrete_topology {()}) {()} X S (\<lambda>x. a) \<circ>
 | |
| 251 |              hom_induced p (discrete_topology {()}) {} (discrete_topology {()}) {()} id \<circ> ?g) y"
 | |
| 252 | by (simp add: a \<open>a \<in> S\<close> flip: hom_induced_compose) | |
| 253 | also have "\<dots> = \<one>\<^bsub>relative_homology_group p X S\<^esub>" | |
| 254 |         using trivial_relative_homology_group_topspace [of p "discrete_topology {()}"]
 | |
| 255 | apply simp | |
| 256 | by (metis (full_types) empty_iff gh1.H.one_closed gh1.H.trivial_group gh2.hom_one hom_induced_carrier insert_iff) | |
| 257 | finally have "?h (?f (?g y)) = \<one>\<^bsub>relative_homology_group p X S\<^esub>" | |
| 258 | by simp | |
| 259 | then show "?h y = ?h (y \<otimes>\<^bsub>?H\<^esub> inv\<^bsub>?H\<^esub> ?f (?g y))" | |
| 260 | by (simp add: that hom_induced_carrier) | |
| 261 | show "(y \<otimes>\<^bsub>?H\<^esub> inv\<^bsub>?H\<^esub> ?f (?g y)) \<in> carrier (homology_group p X)" | |
| 262 | by (simp add: hom_induced_carrier that) | |
| 263 |       have *: "(?g \<circ> hom_induced p X {} X {} (\<lambda>x. a)) y = hom_induced p X {} (discrete_topology {()}) {} (\<lambda>a. ()) y"
 | |
| 264 | by (simp add: a \<open>a \<in> S\<close> flip: hom_induced_compose) | |
| 265 | have "?g (y \<otimes>\<^bsub>?H\<^esub> inv\<^bsub>?H\<^esub> (?f \<circ> ?g) y) | |
| 266 |           = \<one>\<^bsub>homology_group p (discrete_topology {()})\<^esub>"
 | |
| 267 | by (simp add: a \<open>a \<in> S\<close> that hom_induced_carrier flip: hom_induced_compose * [unfolded o_def]) | |
| 268 | then show "?g (y \<otimes>\<^bsub>?H\<^esub> inv\<^bsub>?H\<^esub> ?f (?g y)) | |
| 269 |           = \<one>\<^bsub>homology_group p (discrete_topology {()})\<^esub>"
 | |
| 270 | by simp | |
| 271 | qed | |
| 272 | qed | |
| 273 | show ?thesis | |
| 274 | apply (auto simp: reduced_homology_group_def carrier_subgroup_generated kernel_def image_iff) | |
| 275 | apply (metis (no_types, lifting) generate_in_carrier mem_Collect_eq subsetI) | |
| 276 | apply (force simp: dest: * intro: generate.incl) | |
| 277 | done | |
| 278 | qed | |
| 279 | ||
| 280 | ||
| 281 | lemma homology_exactness_reduced_1: | |
| 282 |   assumes "topspace X \<inter> S \<noteq> {}"
 | |
| 283 | shows "exact_seq([reduced_homology_group(p - 1) (subtopology X S), | |
| 284 | relative_homology_group p X S, | |
| 285 | reduced_homology_group p X], | |
| 286 |                     [hom_boundary p X S, hom_induced p X {} X S id])"
 | |
| 287 | (is "exact_seq ([?G1,?G2,?G3], [?h1,?h2])") | |
| 288 | proof - | |
| 289 | have *: "?h2 ` carrier (homology_group p X) | |
| 290 | = kernel ?G2 (homology_group (p - 1) (subtopology X S)) ?h1" | |
| 291 | using homology_exactness_axiom_1 [of p X S] by simp | |
| 292 | have gh: "group_hom ?G3 ?G2 ?h2" | |
| 293 | by (simp add: reduced_homology_group_def group_hom_def group_hom_axioms_def | |
| 294 | group.group_subgroup_generated group.hom_from_subgroup_generated hom_induced_hom) | |
| 295 | show ?thesis | |
| 296 | apply (simp add: hom_boundary_reduced_hom gh * image_reduced_homology_group [OF assms]) | |
| 297 | apply (simp add: kernel_def one_reduced_homology_group) | |
| 298 | done | |
| 299 | qed | |
| 300 | ||
| 301 | ||
| 302 | lemma homology_exactness_reduced_2: | |
| 303 | "exact_seq([reduced_homology_group(p - 1) X, | |
| 304 | reduced_homology_group(p - 1) (subtopology X S), | |
| 305 | relative_homology_group p X S], | |
| 306 |                 [hom_induced (p - 1) (subtopology X S) {} X {} id, hom_boundary p X S])"
 | |
| 307 | (is "exact_seq ([?G1,?G2,?G3], [?h1,?h2])") | |
| 308 | using homology_exactness_axiom_2 [of p X S] | |
| 309 | apply (simp add: group_hom_axioms_def group_hom_def hom_boundary_reduced_hom hom_induced_reduced_hom) | |
| 310 | apply (simp add: reduced_homology_group_def group_hom.subgroup_kernel group_hom_axioms_def group_hom_def hom_induced_hom) | |
| 311 | using hom_boundary_reduced_hom [of p X S] | |
| 312 | apply (auto simp: image_def set_eq_iff) | |
| 313 | by (metis carrier_reduced_homology_group hom_in_carrier set_eq_iff) | |
| 314 | ||
| 315 | ||
| 316 | lemma homology_exactness_reduced_3: | |
| 317 | "exact_seq([relative_homology_group p X S, | |
| 318 | reduced_homology_group p X, | |
| 319 | reduced_homology_group p (subtopology X S)], | |
| 320 |               [hom_induced p X {} X S id, hom_induced p (subtopology X S) {} X {} id])"
 | |
| 321 | (is "exact_seq ([?G1,?G2,?G3], [?h1,?h2])") | |
| 322 | proof - | |
| 323 | have "kernel ?G2 ?G1 ?h1 = | |
| 324 | ?h2 ` carrier ?G3" | |
| 325 | proof - | |
| 326 | obtain U where U: | |
| 327 |       "(hom_induced p (subtopology X S) {} X {} id) ` carrier ?G3 \<subseteq> U"
 | |
| 328 |       "(hom_induced p (subtopology X S) {} X {} id) ` carrier ?G3
 | |
| 329 |        \<subseteq> (hom_induced p (subtopology X S) {} X {} id) ` carrier (homology_group p (subtopology X S))"
 | |
| 330 |       "U \<inter> kernel (homology_group p X) ?G1 (hom_induced p X {} X S id)
 | |
| 331 |      = kernel ?G2 ?G1 (hom_induced p X {} X S id)"
 | |
| 332 |       "U \<inter> (hom_induced p (subtopology X S) {} X {} id) ` carrier (homology_group p (subtopology X S))
 | |
| 333 |     \<subseteq> (hom_induced p (subtopology X S) {} X {} id) ` carrier ?G3"
 | |
| 334 | proof | |
| 335 | show "?h2 ` carrier ?G3 \<subseteq> carrier ?G2" | |
| 336 | by (simp add: hom_induced_reduced image_subset_iff) | |
| 337 | show "?h2 ` carrier ?G3 \<subseteq> ?h2 ` carrier (homology_group p (subtopology X S))" | |
| 338 | by (meson carrier_reduced_homology_group_subset image_mono) | |
| 339 |       have "subgroup (kernel (homology_group p X) (homology_group p (discrete_topology {()}))
 | |
| 340 |                              (hom_induced p X {} (discrete_topology {()}) {} (\<lambda>x. ())))
 | |
| 341 | (homology_group p X)" | |
| 342 | by (simp add: group.normal_invE(1) group_hom.normal_kernel group_hom_axioms_def group_hom_def hom_induced_empty_hom) | |
| 343 | then show "carrier ?G2 \<inter> kernel (homology_group p X) ?G1 ?h1 = kernel ?G2 ?G1 ?h1" | |
| 344 | unfolding carrier_reduced_homology_group | |
| 345 | by (auto simp: reduced_homology_group_def) | |
| 346 | show "carrier ?G2 \<inter> ?h2 ` carrier (homology_group p (subtopology X S)) | |
| 347 | \<subseteq> ?h2 ` carrier ?G3" | |
| 348 | by (force simp: carrier_reduced_homology_group kernel_def hom_induced_compose') | |
| 349 | qed | |
| 350 | with homology_exactness_axiom_3 [of p X S] show ?thesis | |
| 351 | by (fastforce simp add:) | |
| 352 | qed | |
| 353 | then show ?thesis | |
| 354 | apply (simp add: group_hom_axioms_def group_hom_def hom_boundary_reduced_hom hom_induced_reduced_hom) | |
| 355 | apply (simp add: group.hom_from_subgroup_generated hom_induced_hom reduced_homology_group_def) | |
| 356 | done | |
| 357 | qed | |
| 358 | ||
| 359 | ||
| 360 | subsection\<open>More homology properties of deformations, retracts, contractible spaces\<close> | |
| 361 | ||
| 362 | lemma iso_relative_homology_of_contractible: | |
| 363 |    "\<lbrakk>contractible_space X; topspace X \<inter> S \<noteq> {}\<rbrakk>
 | |
| 364 | \<Longrightarrow> hom_boundary p X S | |
| 365 | \<in> iso (relative_homology_group p X S) (reduced_homology_group(p - 1) (subtopology X S))" | |
| 366 | using very_short_exact_sequence | |
| 367 | [of "reduced_homology_group (p - 1) X" | |
| 368 | "reduced_homology_group (p - 1) (subtopology X S)" | |
| 369 | "relative_homology_group p X S" | |
| 370 | "reduced_homology_group p X" | |
| 371 |         "hom_induced (p - 1) (subtopology X S) {} X {} id"
 | |
| 372 | "hom_boundary p X S" | |
| 373 |         "hom_induced p X {} X S id"]
 | |
| 374 | by (meson exact_seq_cons_iff homology_exactness_reduced_1 homology_exactness_reduced_2 trivial_reduced_homology_group_contractible_space) | |
| 375 | ||
| 376 | lemma isomorphic_group_relative_homology_of_contractible: | |
| 377 |    "\<lbrakk>contractible_space X; topspace X \<inter> S \<noteq> {}\<rbrakk>
 | |
| 378 | \<Longrightarrow> relative_homology_group p X S \<cong> | |
| 379 | reduced_homology_group(p - 1) (subtopology X S)" | |
| 380 | by (meson iso_relative_homology_of_contractible is_isoI) | |
| 381 | ||
| 382 | lemma isomorphic_group_reduced_homology_of_contractible: | |
| 383 |    "\<lbrakk>contractible_space X; topspace X \<inter> S \<noteq> {}\<rbrakk>
 | |
| 384 | \<Longrightarrow> reduced_homology_group p (subtopology X S) \<cong> relative_homology_group(p + 1) X S" | |
| 385 | by (metis add.commute add_diff_cancel_left' group.iso_sym group_relative_homology_group isomorphic_group_relative_homology_of_contractible) | |
| 386 | ||
| 387 | lemma iso_reduced_homology_by_contractible: | |
| 388 |    "\<lbrakk>contractible_space(subtopology X S); topspace X \<inter> S \<noteq> {}\<rbrakk>
 | |
| 389 |       \<Longrightarrow> (hom_induced p X {} X S id) \<in> iso (reduced_homology_group p X) (relative_homology_group p X S)"
 | |
| 390 | using very_short_exact_sequence | |
| 391 | [of "reduced_homology_group (p - 1) (subtopology X S)" | |
| 392 | "relative_homology_group p X S" | |
| 393 | "reduced_homology_group p X" | |
| 394 | "reduced_homology_group p (subtopology X S)" | |
| 395 | "hom_boundary p X S" | |
| 396 |         "hom_induced p X {} X S id"
 | |
| 397 |         "hom_induced p (subtopology X S) {} X {} id"]
 | |
| 398 | by (meson exact_seq_cons_iff homology_exactness_reduced_1 homology_exactness_reduced_3 trivial_reduced_homology_group_contractible_space) | |
| 399 | ||
| 400 | lemma isomorphic_reduced_homology_by_contractible: | |
| 401 |    "\<lbrakk>contractible_space(subtopology X S); topspace X \<inter> S \<noteq> {}\<rbrakk>
 | |
| 402 | \<Longrightarrow> reduced_homology_group p X \<cong> relative_homology_group p X S" | |
| 403 | using is_isoI iso_reduced_homology_by_contractible by blast | |
| 404 | ||
| 405 | lemma isomorphic_relative_homology_by_contractible: | |
| 406 |    "\<lbrakk>contractible_space(subtopology X S); topspace X \<inter> S \<noteq> {}\<rbrakk>
 | |
| 407 | \<Longrightarrow> relative_homology_group p X S \<cong> reduced_homology_group p X" | |
| 408 | using group.iso_sym group_reduced_homology_group isomorphic_reduced_homology_by_contractible by blast | |
| 409 | ||
| 410 | lemma isomorphic_reduced_homology_by_singleton: | |
| 411 |    "a \<in> topspace X \<Longrightarrow> reduced_homology_group p X \<cong> relative_homology_group p X ({a})"
 | |
| 412 | by (simp add: contractible_space_subtopology_singleton isomorphic_reduced_homology_by_contractible) | |
| 413 | ||
| 414 | lemma isomorphic_relative_homology_by_singleton: | |
| 415 |    "a \<in> topspace X \<Longrightarrow> relative_homology_group p X ({a}) \<cong> reduced_homology_group p X"
 | |
| 416 | by (simp add: group.iso_sym isomorphic_reduced_homology_by_singleton) | |
| 417 | ||
| 418 | lemma reduced_homology_group_pair: | |
| 419 | assumes "t1_space X" and a: "a \<in> topspace X" and b: "b \<in> topspace X" and "a \<noteq> b" | |
| 420 |   shows "reduced_homology_group p (subtopology X {a,b}) \<cong> homology_group p (subtopology X {a})"
 | |
| 421 | (is "?lhs \<cong> ?rhs") | |
| 422 | proof - | |
| 423 |   have "?lhs \<cong> relative_homology_group p (subtopology X {a,b}) {b}"
 | |
| 424 | by (simp add: b isomorphic_reduced_homology_by_singleton topspace_subtopology) | |
| 425 | also have "\<dots> \<cong> ?rhs" | |
| 426 | proof - | |
| 427 |     have sub: "subtopology X {a, b} closure_of {b} \<subseteq> subtopology X {a, b} interior_of {b}"
 | |
| 428 | by (simp add: assms t1_space_subtopology closure_of_singleton subtopology_eq_discrete_topology_finite discrete_topology_closure_of) | |
| 429 | show ?thesis | |
| 430 |       using homology_excision_axiom [OF sub, of "{a,b}" p]
 | |
| 431 | by (simp add: assms(4) group.iso_sym is_isoI subtopology_subtopology) | |
| 432 | qed | |
| 433 | finally show ?thesis . | |
| 434 | qed | |
| 435 | ||
| 436 | ||
| 437 | lemma deformation_retraction_relative_homology_group_isomorphisms: | |
| 438 | "\<lbrakk>retraction_maps X Y r s; r ` U \<subseteq> V; s ` V \<subseteq> U; homotopic_with (\<lambda>h. h ` U \<subseteq> U) X X (s \<circ> r) id\<rbrakk> | |
| 439 | \<Longrightarrow> group_isomorphisms (relative_homology_group p X U) (relative_homology_group p Y V) | |
| 440 | (hom_induced p X U Y V r) (hom_induced p Y V X U s)" | |
| 441 | apply (simp add: retraction_maps_def) | |
| 442 | apply (rule homotopy_equivalence_relative_homology_group_isomorphisms) | |
| 443 | apply (auto simp: image_subset_iff continuous_map_compose homotopic_with_equal) | |
| 444 | done | |
| 445 | ||
| 446 | ||
| 447 | lemma deformation_retract_relative_homology_group_isomorphisms: | |
| 448 | "\<lbrakk>retraction_maps X Y r id; V \<subseteq> U; r ` U \<subseteq> V; homotopic_with (\<lambda>h. h ` U \<subseteq> U) X X r id\<rbrakk> | |
| 449 | \<Longrightarrow> group_isomorphisms (relative_homology_group p X U) (relative_homology_group p Y V) | |
| 450 | (hom_induced p X U Y V r) (hom_induced p Y V X U id)" | |
| 451 | by (simp add: deformation_retraction_relative_homology_group_isomorphisms) | |
| 452 | ||
| 453 | lemma deformation_retract_relative_homology_group_isomorphism: | |
| 454 | "\<lbrakk>retraction_maps X Y r id; V \<subseteq> U; r ` U \<subseteq> V; homotopic_with (\<lambda>h. h ` U \<subseteq> U) X X r id\<rbrakk> | |
| 455 | \<Longrightarrow> (hom_induced p X U Y V r) \<in> iso (relative_homology_group p X U) (relative_homology_group p Y V)" | |
| 456 | by (metis deformation_retract_relative_homology_group_isomorphisms group_isomorphisms_imp_iso) | |
| 457 | ||
| 458 | lemma deformation_retract_relative_homology_group_isomorphism_id: | |
| 459 | "\<lbrakk>retraction_maps X Y r id; V \<subseteq> U; r ` U \<subseteq> V; homotopic_with (\<lambda>h. h ` U \<subseteq> U) X X r id\<rbrakk> | |
| 460 | \<Longrightarrow> (hom_induced p Y V X U id) \<in> iso (relative_homology_group p Y V) (relative_homology_group p X U)" | |
| 461 | by (metis deformation_retract_relative_homology_group_isomorphisms group_isomorphisms_imp_iso group_isomorphisms_sym) | |
| 462 | ||
| 463 | lemma deformation_retraction_imp_isomorphic_relative_homology_groups: | |
| 464 | "\<lbrakk>retraction_maps X Y r s; r ` U \<subseteq> V; s ` V \<subseteq> U; homotopic_with (\<lambda>h. h ` U \<subseteq> U) X X (s \<circ> r) id\<rbrakk> | |
| 465 | \<Longrightarrow> relative_homology_group p X U \<cong> relative_homology_group p Y V" | |
| 466 | by (blast intro: is_isoI group_isomorphisms_imp_iso deformation_retraction_relative_homology_group_isomorphisms) | |
| 467 | ||
| 468 | lemma deformation_retraction_imp_isomorphic_homology_groups: | |
| 469 | "\<lbrakk>retraction_maps X Y r s; homotopic_with (\<lambda>h. True) X X (s \<circ> r) id\<rbrakk> | |
| 470 | \<Longrightarrow> homology_group p X \<cong> homology_group p Y" | |
| 471 | by (simp add: deformation_retraction_imp_homotopy_equivalent_space homotopy_equivalent_space_imp_isomorphic_homology_groups) | |
| 472 | ||
| 473 | lemma deformation_retract_imp_isomorphic_relative_homology_groups: | |
| 474 | "\<lbrakk>retraction_maps X X' r id; V \<subseteq> U; r ` U \<subseteq> V; homotopic_with (\<lambda>h. h ` U \<subseteq> U) X X r id\<rbrakk> | |
| 475 | \<Longrightarrow> relative_homology_group p X U \<cong> relative_homology_group p X' V" | |
| 476 | by (simp add: deformation_retraction_imp_isomorphic_relative_homology_groups) | |
| 477 | ||
| 478 | lemma deformation_retract_imp_isomorphic_homology_groups: | |
| 479 | "\<lbrakk>retraction_maps X X' r id; homotopic_with (\<lambda>h. True) X X r id\<rbrakk> | |
| 480 | \<Longrightarrow> homology_group p X \<cong> homology_group p X'" | |
| 481 | by (simp add: deformation_retraction_imp_isomorphic_homology_groups) | |
| 482 | ||
| 483 | ||
| 484 | lemma epi_hom_induced_inclusion: | |
| 485 | assumes "homotopic_with (\<lambda>x. True) X X id f" and "f ` (topspace X) \<subseteq> S" | |
| 486 |   shows "(hom_induced p (subtopology X S) {} X {} id)
 | |
| 487 | \<in> epi (homology_group p (subtopology X S)) (homology_group p X)" | |
| 488 | proof (rule epi_right_invertible) | |
| 489 |   show "hom_induced p (subtopology X S) {} X {} id
 | |
| 490 | \<in> hom (homology_group p (subtopology X S)) (homology_group p X)" | |
| 491 | by (simp add: hom_induced_empty_hom) | |
| 492 |   show "hom_induced p X {} (subtopology X S) {} f
 | |
| 493 | \<in> carrier (homology_group p X) \<rightarrow> carrier (homology_group p (subtopology X S))" | |
| 494 | by (simp add: hom_induced_carrier) | |
| 495 | fix x | |
| 496 | assume "x \<in> carrier (homology_group p X)" | |
| 497 |   then show "hom_induced p (subtopology X S) {} X {} id (hom_induced p X {} (subtopology X S) {} f x) = x"
 | |
| 498 | by (metis assms continuous_map_id_subt continuous_map_in_subtopology hom_induced_compose' hom_induced_id homology_homotopy_empty homotopic_with_imp_continuous_maps image_empty order_refl) | |
| 499 | qed | |
| 500 | ||
| 501 | ||
| 502 | lemma trivial_homomorphism_hom_induced_relativization: | |
| 503 | assumes "homotopic_with (\<lambda>x. True) X X id f" and "f ` (topspace X) \<subseteq> S" | |
| 504 | shows "trivial_homomorphism (homology_group p X) (relative_homology_group p X S) | |
| 505 |               (hom_induced p X {} X S id)"
 | |
| 506 | proof - | |
| 507 |   have "(hom_induced p (subtopology X S) {} X {} id)
 | |
| 508 | \<in> epi (homology_group p (subtopology X S)) (homology_group p X)" | |
| 509 | by (metis assms epi_hom_induced_inclusion) | |
| 510 | then show ?thesis | |
| 511 | using homology_exactness_axiom_3 [of p X S] homology_exactness_axiom_1 [of p X S] | |
| 512 | by (simp add: epi_def group.trivial_homomorphism_image group_hom.trivial_hom_iff) | |
| 513 | qed | |
| 514 | ||
| 515 | ||
| 516 | lemma mon_hom_boundary_inclusion: | |
| 517 | assumes "homotopic_with (\<lambda>x. True) X X id f" and "f ` (topspace X) \<subseteq> S" | |
| 518 | shows "(hom_boundary p X S) \<in> mon | |
| 519 | (relative_homology_group p X S) (homology_group (p - 1) (subtopology X S))" | |
| 520 | proof - | |
| 521 |   have "(hom_induced p (subtopology X S) {} X {} id)
 | |
| 522 | \<in> epi (homology_group p (subtopology X S)) (homology_group p X)" | |
| 523 | by (metis assms epi_hom_induced_inclusion) | |
| 524 | then show ?thesis | |
| 525 | using homology_exactness_axiom_3 [of p X S] homology_exactness_axiom_1 [of p X S] | |
| 526 | apply (simp add: mon_def epi_def hom_boundary_hom) | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
72632diff
changeset | 527 | by (metis (no_types, opaque_lifting) group_hom.trivial_hom_iff group_hom.trivial_ker_imp_inj group_hom_axioms_def group_hom_def group_relative_homology_group hom_boundary_hom) | 
| 70095 | 528 | qed | 
| 529 | ||
| 530 | lemma short_exact_sequence_hom_induced_relativization: | |
| 531 | assumes "homotopic_with (\<lambda>x. True) X X id f" and "f ` (topspace X) \<subseteq> S" | |
| 532 | shows "short_exact_sequence (homology_group (p-1) X) (homology_group (p-1) (subtopology X S)) (relative_homology_group p X S) | |
| 533 |                    (hom_induced (p-1) (subtopology X S) {} X {} id) (hom_boundary p X S)"
 | |
| 534 | unfolding short_exact_sequence_iff | |
| 535 | by (intro conjI homology_exactness_axiom_2 epi_hom_induced_inclusion [OF assms] mon_hom_boundary_inclusion [OF assms]) | |
| 536 | ||
| 537 | ||
| 538 | lemma group_isomorphisms_homology_group_prod_deformation: | |
| 539 | fixes p::int | |
| 540 | assumes "homotopic_with (\<lambda>x. True) X X id f" and "f ` (topspace X) \<subseteq> S" | |
| 541 | obtains H K where | |
| 542 | "subgroup H (homology_group p (subtopology X S))" | |
| 543 | "subgroup K (homology_group p (subtopology X S))" | |
| 544 | "(\<lambda>(x, y). x \<otimes>\<^bsub>homology_group p (subtopology X S)\<^esub> y) | |
| 545 | \<in> Group.iso (subgroup_generated (homology_group p (subtopology X S)) H \<times>\<times> | |
| 546 | subgroup_generated (homology_group p (subtopology X S)) K) | |
| 547 | (homology_group p (subtopology X S))" | |
| 548 | "hom_boundary (p + 1) X S | |
| 549 | \<in> Group.iso (relative_homology_group (p + 1) X S) | |
| 550 | (subgroup_generated (homology_group p (subtopology X S)) H)" | |
| 551 |     "hom_induced p (subtopology X S) {} X {} id
 | |
| 552 | \<in> Group.iso | |
| 553 | (subgroup_generated (homology_group p (subtopology X S)) K) | |
| 554 | (homology_group p X)" | |
| 555 | proof - | |
| 556 | let ?rhs = "relative_homology_group (p + 1) X S" | |
| 557 | let ?pXS = "homology_group p (subtopology X S)" | |
| 558 | let ?pX = "homology_group p X" | |
| 559 | let ?hb = "hom_boundary (p + 1) X S" | |
| 560 |   let ?hi = "hom_induced p (subtopology X S) {} X {} id"
 | |
| 561 | have x: "short_exact_sequence (?pX) ?pXS ?rhs ?hi ?hb" | |
| 562 | using short_exact_sequence_hom_induced_relativization [OF assms, of "p + 1"] by simp | |
| 563 | have contf: "continuous_map X (subtopology X S) f" | |
| 564 | by (meson assms continuous_map_in_subtopology homotopic_with_imp_continuous_maps) | |
| 565 |   obtain H K where HK: "H \<lhd> ?pXS" "subgroup K ?pXS" "H \<inter> K \<subseteq> {one ?pXS}" "set_mult ?pXS H K = carrier ?pXS"
 | |
| 566 | and iso: "?hb \<in> iso ?rhs (subgroup_generated ?pXS H)" "?hi \<in> iso (subgroup_generated ?pXS K) ?pX" | |
| 567 |     apply (rule splitting_lemma_right [OF x, where g' = "hom_induced p X {} (subtopology X S) {} f"])
 | |
| 568 | apply (simp add: hom_induced_empty_hom) | |
| 569 | apply (simp add: contf hom_induced_compose') | |
| 570 | apply (metis (full_types) assms(1) hom_induced_id homology_homotopy_empty) | |
| 571 | apply blast | |
| 572 | done | |
| 573 | show ?thesis | |
| 574 | proof | |
| 575 | show "subgroup H ?pXS" | |
| 576 | using HK(1) normal_imp_subgroup by blast | |
| 577 | then show "(\<lambda>(x, y). x \<otimes>\<^bsub>?pXS\<^esub> y) | |
| 578 | \<in> Group.iso (subgroup_generated (?pXS) H \<times>\<times> subgroup_generated (?pXS) K) (?pXS)" | |
| 579 | by (meson HK abelian_relative_homology_group group_disjoint_sum.iso_group_mul group_disjoint_sum_def group_relative_homology_group) | |
| 580 | show "subgroup K ?pXS" | |
| 581 | by (rule HK) | |
| 582 | show "hom_boundary (p + 1) X S \<in> Group.iso ?rhs (subgroup_generated (?pXS) H)" | |
| 583 | using iso int_ops(4) by presburger | |
| 584 |     show "hom_induced p (subtopology X S) {} X {} id \<in> Group.iso (subgroup_generated (?pXS) K) (?pX)"
 | |
| 585 | by (simp add: iso(2)) | |
| 586 | qed | |
| 587 | qed | |
| 588 | ||
| 589 | lemma iso_homology_group_prod_deformation: | |
| 590 | assumes "homotopic_with (\<lambda>x. True) X X id f" and "f ` (topspace X) \<subseteq> S" | |
| 591 | shows "homology_group p (subtopology X S) | |
| 592 | \<cong> DirProd (homology_group p X) (relative_homology_group(p + 1) X S)" | |
| 593 | (is "?G \<cong> DirProd ?H ?R") | |
| 594 | proof - | |
| 595 | obtain H K where HK: | |
| 596 | "(\<lambda>(x, y). x \<otimes>\<^bsub>?G\<^esub> y) | |
| 597 | \<in> Group.iso (subgroup_generated (?G) H \<times>\<times> subgroup_generated (?G) K) (?G)" | |
| 598 | "hom_boundary (p + 1) X S \<in> Group.iso (?R) (subgroup_generated (?G) H)" | |
| 599 |     "hom_induced p (subtopology X S) {} X {} id \<in> Group.iso (subgroup_generated (?G) K) (?H)"
 | |
| 600 | by (blast intro: group_isomorphisms_homology_group_prod_deformation [OF assms]) | |
| 601 | have "?G \<cong> DirProd (subgroup_generated (?G) H) (subgroup_generated (?G) K)" | |
| 602 | by (meson DirProd_group HK(1) group.group_subgroup_generated group.iso_sym group_relative_homology_group is_isoI) | |
| 603 | also have "\<dots> \<cong> DirProd ?R ?H" | |
| 604 | by (meson HK group.DirProd_iso_trans group.group_subgroup_generated group.iso_sym group_relative_homology_group is_isoI) | |
| 605 | also have "\<dots> \<cong> DirProd ?H ?R" | |
| 606 | by (simp add: DirProd_commute_iso) | |
| 607 | finally show ?thesis . | |
| 608 | qed | |
| 609 | ||
| 610 | ||
| 611 | ||
| 612 | lemma iso_homology_contractible_space_subtopology1: | |
| 613 |   assumes "contractible_space X" "S \<subseteq> topspace X" "S \<noteq> {}"
 | |
| 614 | shows "homology_group 0 (subtopology X S) \<cong> DirProd integer_group (relative_homology_group(1) X S)" | |
| 615 | proof - | |
| 616 | obtain f where "homotopic_with (\<lambda>x. True) X X id f" and "f ` (topspace X) \<subseteq> S" | |
| 617 | using assms contractible_space_alt by fastforce | |
| 618 | then have "homology_group 0 (subtopology X S) \<cong> homology_group 0 X \<times>\<times> relative_homology_group 1 X S" | |
| 619 | using iso_homology_group_prod_deformation [of X _ S 0] by auto | |
| 620 | also have "\<dots> \<cong> integer_group \<times>\<times> relative_homology_group 1 X S" | |
| 621 | using assms contractible_imp_path_connected_space group.DirProd_iso_trans group_relative_homology_group iso_refl isomorphic_integer_zeroth_homology_group by blast | |
| 622 | finally show ?thesis . | |
| 623 | qed | |
| 624 | ||
| 625 | lemma iso_homology_contractible_space_subtopology2: | |
| 626 |   "\<lbrakk>contractible_space X; S \<subseteq> topspace X; p \<noteq> 0; S \<noteq> {}\<rbrakk>
 | |
| 627 | \<Longrightarrow> homology_group p (subtopology X S) \<cong> relative_homology_group (p + 1) X S" | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
72632diff
changeset | 628 | by (metis (no_types, opaque_lifting) add.commute isomorphic_group_reduced_homology_of_contractible topspace_subtopology topspace_subtopology_subset un_reduced_homology_group) | 
| 70095 | 629 | |
| 630 | lemma trivial_relative_homology_group_contractible_spaces: | |
| 631 |    "\<lbrakk>contractible_space X; contractible_space(subtopology X S); topspace X \<inter> S \<noteq> {}\<rbrakk>
 | |
| 632 | \<Longrightarrow> trivial_group(relative_homology_group p X S)" | |
| 633 | using group_reduced_homology_group group_relative_homology_group isomorphic_group_triviality isomorphic_relative_homology_by_contractible trivial_reduced_homology_group_contractible_space by blast | |
| 634 | ||
| 635 | lemma trivial_relative_homology_group_alt: | |
| 636 | assumes contf: "continuous_map X (subtopology X S) f" and hom: "homotopic_with (\<lambda>k. k ` S \<subseteq> S) X X f id" | |
| 637 | shows "trivial_group (relative_homology_group p X S)" | |
| 638 | proof (rule trivial_relative_homology_group_gen [OF contf]) | |
| 639 | show "homotopic_with (\<lambda>h. True) (subtopology X S) (subtopology X S) f id" | |
| 640 | using hom unfolding homotopic_with_def | |
| 641 | apply (rule ex_forward) | |
| 642 | apply (auto simp: prod_topology_subtopology continuous_map_in_subtopology continuous_map_from_subtopology image_subset_iff topspace_subtopology) | |
| 643 | done | |
| 644 | show "homotopic_with (\<lambda>k. True) X X f id" | |
| 645 | using assms by (force simp: homotopic_with_def) | |
| 646 | qed | |
| 647 | ||
| 648 | ||
| 649 | lemma iso_hom_induced_relativization_contractible: | |
| 650 |   assumes "contractible_space(subtopology X S)" "contractible_space(subtopology X T)" "T \<subseteq> S" "topspace X \<inter> T \<noteq> {}"
 | |
| 651 | shows "(hom_induced p X T X S id) \<in> iso (relative_homology_group p X T) (relative_homology_group p X S)" | |
| 652 | proof (rule very_short_exact_sequence) | |
| 653 | show "exact_seq | |
| 654 | ([relative_homology_group(p - 1) (subtopology X S) T, relative_homology_group p X S, relative_homology_group p X T, relative_homology_group p (subtopology X S) T], | |
| 655 | [hom_relboundary p X S T, hom_induced p X T X S id, hom_induced p (subtopology X S) T X T id])" | |
| 656 | using homology_exactness_triple_1 [OF \<open>T \<subseteq> S\<close>] homology_exactness_triple_3 [OF \<open>T \<subseteq> S\<close>] | |
| 657 | by fastforce | |
| 658 | show "trivial_group (relative_homology_group p (subtopology X S) T)" "trivial_group (relative_homology_group(p - 1) (subtopology X S) T)" | |
| 659 | using assms | |
| 660 | by (force simp: inf.absorb_iff2 subtopology_subtopology topspace_subtopology intro!: trivial_relative_homology_group_contractible_spaces)+ | |
| 661 | qed | |
| 662 | ||
| 663 | corollary isomorphic_relative_homology_groups_relativization_contractible: | |
| 664 |   assumes "contractible_space(subtopology X S)" "contractible_space(subtopology X T)" "T \<subseteq> S" "topspace X \<inter> T \<noteq> {}"
 | |
| 665 | shows "relative_homology_group p X T \<cong> relative_homology_group p X S" | |
| 666 | by (rule is_isoI) (rule iso_hom_induced_relativization_contractible [OF assms]) | |
| 667 | ||
| 668 | lemma iso_hom_induced_inclusion_contractible: | |
| 669 |   assumes "contractible_space X" "contractible_space(subtopology X S)" "T \<subseteq> S" "topspace X \<inter> S \<noteq> {}"
 | |
| 670 | shows "(hom_induced p (subtopology X S) T X T id) | |
| 671 | \<in> iso (relative_homology_group p (subtopology X S) T) (relative_homology_group p X T)" | |
| 672 | proof (rule very_short_exact_sequence) | |
| 673 | show "exact_seq | |
| 674 | ([relative_homology_group p X S, relative_homology_group p X T, | |
| 675 | relative_homology_group p (subtopology X S) T, relative_homology_group (p+1) X S], | |
| 676 | [hom_induced p X T X S id, hom_induced p (subtopology X S) T X T id, hom_relboundary (p+1) X S T])" | |
| 677 | using homology_exactness_triple_2 [OF \<open>T \<subseteq> S\<close>] homology_exactness_triple_3 [OF \<open>T \<subseteq> S\<close>] | |
| 678 | by (metis add_diff_cancel_left' diff_add_cancel exact_seq_cons_iff) | |
| 679 | show "trivial_group (relative_homology_group (p+1) X S)" "trivial_group (relative_homology_group p X S)" | |
| 680 | using assms | |
| 681 | by (auto simp: subtopology_subtopology topspace_subtopology intro!: trivial_relative_homology_group_contractible_spaces) | |
| 682 | qed | |
| 683 | ||
| 684 | corollary isomorphic_relative_homology_groups_inclusion_contractible: | |
| 685 |   assumes "contractible_space X" "contractible_space(subtopology X S)" "T \<subseteq> S" "topspace X \<inter> S \<noteq> {}"
 | |
| 686 | shows "relative_homology_group p (subtopology X S) T \<cong> relative_homology_group p X T" | |
| 687 | by (rule is_isoI) (rule iso_hom_induced_inclusion_contractible [OF assms]) | |
| 688 | ||
| 689 | lemma iso_hom_relboundary_contractible: | |
| 690 |   assumes "contractible_space X" "contractible_space(subtopology X T)" "T \<subseteq> S" "topspace X \<inter> T \<noteq> {}"
 | |
| 691 | shows "hom_relboundary p X S T | |
| 692 | \<in> iso (relative_homology_group p X S) (relative_homology_group (p - 1) (subtopology X S) T)" | |
| 693 | proof (rule very_short_exact_sequence) | |
| 694 | show "exact_seq | |
| 695 | ([relative_homology_group (p - 1) X T, relative_homology_group (p - 1) (subtopology X S) T, relative_homology_group p X S, relative_homology_group p X T], | |
| 696 | [hom_induced (p - 1) (subtopology X S) T X T id, hom_relboundary p X S T, hom_induced p X T X S id])" | |
| 697 | using homology_exactness_triple_1 [OF \<open>T \<subseteq> S\<close>] homology_exactness_triple_2 [OF \<open>T \<subseteq> S\<close>] by simp | |
| 698 | show "trivial_group (relative_homology_group p X T)" "trivial_group (relative_homology_group (p - 1) X T)" | |
| 699 | using assms | |
| 700 | by (auto simp: subtopology_subtopology topspace_subtopology intro!: trivial_relative_homology_group_contractible_spaces) | |
| 701 | qed | |
| 702 | ||
| 703 | corollary isomorphic_relative_homology_groups_relboundary_contractible: | |
| 704 |   assumes "contractible_space X" "contractible_space(subtopology X T)" "T \<subseteq> S" "topspace X \<inter> T \<noteq> {}"
 | |
| 705 | shows "relative_homology_group p X S \<cong> relative_homology_group (p - 1) (subtopology X S) T" | |
| 706 | by (rule is_isoI) (rule iso_hom_relboundary_contractible [OF assms]) | |
| 707 | ||
| 708 | lemma isomorphic_relative_contractible_space_imp_homology_groups: | |
| 709 | assumes "contractible_space X" "contractible_space Y" "S \<subseteq> topspace X" "T \<subseteq> topspace Y" | |
| 710 |      and ST: "S = {} \<longleftrightarrow> T = {}"
 | |
| 711 | and iso: "\<And>p. relative_homology_group p X S \<cong> relative_homology_group p Y T" | |
| 712 | shows "homology_group p (subtopology X S) \<cong> homology_group p (subtopology Y T)" | |
| 713 | proof (cases "T = {}")
 | |
| 714 | case True | |
| 715 |   have "homology_group p (subtopology X {}) \<cong> homology_group p (subtopology Y {})"
 | |
| 716 | by (simp add: homeomorphic_empty_space_eq homeomorphic_space_imp_isomorphic_homology_groups) | |
| 717 | then show ?thesis | |
| 718 | using ST True by blast | |
| 719 | next | |
| 720 | case False | |
| 721 | show ?thesis | |
| 722 | proof (cases "p = 0") | |
| 723 | case True | |
| 724 | have "homology_group p (subtopology X S) \<cong> integer_group \<times>\<times> relative_homology_group 1 X S" | |
| 725 |       using assms True \<open>T \<noteq> {}\<close>
 | |
| 726 | by (simp add: iso_homology_contractible_space_subtopology1) | |
| 727 | also have "\<dots> \<cong> integer_group \<times>\<times> relative_homology_group 1 Y T" | |
| 728 | by (simp add: assms group.DirProd_iso_trans iso_refl) | |
| 729 | also have "\<dots> \<cong> homology_group p (subtopology Y T)" | |
| 730 |       by (simp add: True \<open>T \<noteq> {}\<close> assms group.iso_sym iso_homology_contractible_space_subtopology1)
 | |
| 731 | finally show ?thesis . | |
| 732 | next | |
| 733 | case False | |
| 734 | have "homology_group p (subtopology X S) \<cong> relative_homology_group (p+1) X S" | |
| 735 |       using assms False \<open>T \<noteq> {}\<close>
 | |
| 736 | by (simp add: iso_homology_contractible_space_subtopology2) | |
| 737 | also have "\<dots> \<cong> relative_homology_group (p+1) Y T" | |
| 738 | by (simp add: assms) | |
| 739 | also have "\<dots> \<cong> homology_group p (subtopology Y T)" | |
| 740 |       by (simp add: False \<open>T \<noteq> {}\<close> assms group.iso_sym iso_homology_contractible_space_subtopology2)
 | |
| 741 | finally show ?thesis . | |
| 742 | qed | |
| 743 | qed | |
| 744 | ||
| 745 | ||
| 746 | subsection\<open>Homology groups of spheres\<close> | |
| 747 | ||
| 748 | lemma iso_reduced_homology_group_lower_hemisphere: | |
| 749 | assumes "k \<le> n" | |
| 750 |   shows "hom_induced p (nsphere n) {} (nsphere n) {x. x k \<le> 0} id
 | |
| 751 |       \<in> iso (reduced_homology_group p (nsphere n)) (relative_homology_group p (nsphere n) {x. x k \<le> 0})"
 | |
| 752 | proof (rule iso_reduced_homology_by_contractible) | |
| 753 |   show "contractible_space (subtopology (nsphere n) {x. x k \<le> 0})"
 | |
| 754 | by (simp add: assms contractible_space_lower_hemisphere) | |
| 755 |   have "(\<lambda>i. if i = k then -1 else 0) \<in> topspace (nsphere n) \<inter> {x. x k \<le> 0}"
 | |
| 756 | using assms by (simp add: nsphere if_distrib [of "\<lambda>x. x ^ 2"] cong: if_cong) | |
| 757 |   then show "topspace (nsphere n) \<inter> {x. x k \<le> 0} \<noteq> {}"
 | |
| 758 | by blast | |
| 759 | qed | |
| 760 | ||
| 761 | ||
| 762 | lemma topspace_nsphere_1: | |
| 763 | assumes "x \<in> topspace (nsphere n)" shows "(x k)\<^sup>2 \<le> 1" | |
| 764 | proof (cases "k \<le> n") | |
| 765 | case True | |
| 766 |   have "(\<Sum>i \<in> {..n} - {k}. (x i)\<^sup>2) = (\<Sum>i\<le>n. (x i)\<^sup>2) - (x k)\<^sup>2"
 | |
| 767 | using \<open>k \<le> n\<close> by (simp add: sum_diff) | |
| 768 | then show ?thesis | |
| 769 | using assms | |
| 770 | apply (simp add: nsphere) | |
| 771 | by (metis diff_ge_0_iff_ge sum_nonneg zero_le_power2) | |
| 772 | next | |
| 773 | case False | |
| 774 | then show ?thesis | |
| 775 | using assms by (simp add: nsphere) | |
| 776 | qed | |
| 777 | ||
| 778 | lemma topspace_nsphere_1_eq_0: | |
| 779 | fixes x :: "nat \<Rightarrow> real" | |
| 780 | assumes x: "x \<in> topspace (nsphere n)" and xk: "(x k)\<^sup>2 = 1" and "i \<noteq> k" | |
| 781 | shows "x i = 0" | |
| 782 | proof (cases "i \<le> n") | |
| 783 | case True | |
| 784 | have "k \<le> n" | |
| 785 | using x | |
| 786 | by (simp add: nsphere) (metis not_less xk zero_neq_one zero_power2) | |
| 787 |   have "(\<Sum>i \<in> {..n} - {k}. (x i)\<^sup>2) = (\<Sum>i\<le>n. (x i)\<^sup>2) - (x k)\<^sup>2"
 | |
| 788 | using \<open>k \<le> n\<close> by (simp add: sum_diff) | |
| 789 | also have "\<dots> = 0" | |
| 790 | using assms by (simp add: nsphere) | |
| 791 |   finally have "\<forall>i\<in>{..n} - {k}. (x i)\<^sup>2 = 0"
 | |
| 792 | by (simp add: sum_nonneg_eq_0_iff) | |
| 793 | then show ?thesis | |
| 794 | using True \<open>i \<noteq> k\<close> by auto | |
| 795 | next | |
| 796 | case False | |
| 797 | with x show ?thesis | |
| 798 | by (simp add: nsphere) | |
| 799 | qed | |
| 800 | ||
| 801 | ||
| 802 | proposition iso_relative_homology_group_upper_hemisphere: | |
| 803 |    "(hom_induced p (subtopology (nsphere n) {x. x k \<ge> 0}) {x. x k = 0} (nsphere n) {x. x k \<le> 0} id)
 | |
| 804 |   \<in> iso (relative_homology_group p (subtopology (nsphere n) {x. x k \<ge> 0}) {x. x k = 0})
 | |
| 805 |         (relative_homology_group p (nsphere n) {x. x k \<le> 0})" (is "?h \<in> iso ?G ?H")
 | |
| 806 | proof - | |
| 807 |   have "topspace (nsphere n) \<inter> {x. x k < - 1 / 2} \<subseteq> {x \<in> topspace (nsphere n). x k \<in> {y. y \<le> - 1 / 2}}"
 | |
| 808 | by force | |
| 809 |   moreover have "closedin (nsphere n) {x \<in> topspace (nsphere n). x k \<in> {y. y \<le> - 1 / 2}}"
 | |
| 810 | apply (rule closedin_continuous_map_preimage [OF continuous_map_nsphere_projection]) | |
| 811 | using closed_Collect_le [of id "\<lambda>x::real. -1/2"] apply simp | |
| 812 | done | |
| 813 |   ultimately have "nsphere n closure_of {x. x k < -1/2} \<subseteq> {x \<in> topspace (nsphere n). x k \<in> {y. y \<le> -1/2}}"
 | |
| 814 | by (metis (no_types, lifting) closure_of_eq closure_of_mono closure_of_restrict) | |
| 815 |   also have "\<dots> \<subseteq> {x \<in> topspace (nsphere n). x k \<in> {y. y < 0}}"
 | |
| 816 | by force | |
| 817 |   also have "\<dots> \<subseteq> nsphere n interior_of {x. x k \<le> 0}"
 | |
| 818 | proof (rule interior_of_maximal) | |
| 819 |     show "{x \<in> topspace (nsphere n). x k \<in> {y. y < 0}} \<subseteq> {x. x k \<le> 0}"
 | |
| 820 | by force | |
| 821 |     show "openin (nsphere n) {x \<in> topspace (nsphere n). x k \<in> {y. y < 0}}"
 | |
| 822 | apply (rule openin_continuous_map_preimage [OF continuous_map_nsphere_projection]) | |
| 823 | using open_Collect_less [of id "\<lambda>x::real. 0"] apply simp | |
| 824 | done | |
| 825 | qed | |
| 826 |   finally have nn: "nsphere n closure_of {x. x k < -1/2} \<subseteq> nsphere n interior_of {x. x k \<le> 0}" .
 | |
| 827 |   have [simp]: "{x::nat\<Rightarrow>real. x k \<le> 0} - {x. x k < - (1/2)} = {x. -1/2 \<le> x k \<and> x k \<le> 0}"
 | |
| 828 |                "UNIV - {x::nat\<Rightarrow>real. x k < a} = {x. a \<le> x k}" for a
 | |
| 829 | by auto | |
| 830 |   let ?T01 = "top_of_set {0..1::real}"
 | |
| 831 |   let ?X12 = "subtopology (nsphere n) {x. -1/2 \<le> x k}"
 | |
| 832 |   have 1: "hom_induced p ?X12 {x. -1/2 \<le> x k \<and> x k \<le> 0} (nsphere n) {x. x k \<le> 0} id
 | |
| 833 |          \<in> iso (relative_homology_group p ?X12 {x. -1/2 \<le> x k \<and> x k \<le> 0})
 | |
| 834 | ?H" | |
| 835 | using homology_excision_axiom [OF nn subset_UNIV, of p] by simp | |
| 836 | define h where "h \<equiv> \<lambda>(T,x). let y = max (x k) (-T) in | |
| 837 | (\<lambda>i. if i = k then y else sqrt(1 - y ^ 2) / sqrt(1 - x k ^ 2) * x i)" | |
| 838 | have h: "h(T,x) = x" if "0 \<le> T" "T \<le> 1" "(\<Sum>i\<le>n. (x i)\<^sup>2) = 1" and 0: "\<forall>i>n. x i = 0" "-T \<le> x k" for T x | |
| 839 | using that by (force simp: nsphere h_def Let_def max_def intro!: topspace_nsphere_1_eq_0) | |
| 840 | have "continuous_map (prod_topology ?T01 ?X12) euclideanreal (\<lambda>x. h x i)" for i | |
| 841 | proof - | |
| 842 | show ?thesis | |
| 843 | proof (rule continuous_map_eq) | |
| 844 | show "continuous_map (prod_topology ?T01 ?X12) | |
| 845 | euclideanreal (\<lambda>(T, x). if 0 \<le> x k then x i else h (T, x) i)" | |
| 846 | unfolding case_prod_unfold | |
| 847 | proof (rule continuous_map_cases_le) | |
| 848 | show "continuous_map (prod_topology ?T01 ?X12) euclideanreal (\<lambda>x. snd x k)" | |
| 849 | apply (subst continuous_map_of_snd [unfolded o_def]) | |
| 850 | by (simp add: continuous_map_from_subtopology continuous_map_nsphere_projection) | |
| 851 | next | |
| 852 |         show "continuous_map (subtopology (prod_topology ?T01 ?X12) {p \<in> topspace (prod_topology ?T01 ?X12). 0 \<le> snd p k})
 | |
| 853 | euclideanreal (\<lambda>x. snd x i)" | |
| 854 | apply (rule continuous_map_from_subtopology) | |
| 855 | apply (subst continuous_map_of_snd [unfolded o_def]) | |
| 856 | by (simp add: continuous_map_from_subtopology continuous_map_nsphere_projection) | |
| 857 | next | |
| 858 | note fst = continuous_map_into_fulltopology [OF continuous_map_subtopology_fst] | |
| 859 | have snd: "continuous_map (subtopology (prod_topology ?T01 (subtopology (nsphere n) T)) S) euclideanreal (\<lambda>x. snd x k)" for k S T | |
| 860 | apply (simp add: nsphere) | |
| 861 | apply (rule continuous_map_from_subtopology) | |
| 862 | apply (subst continuous_map_of_snd [unfolded o_def]) | |
| 863 | using continuous_map_from_subtopology continuous_map_nsphere_projection nsphere by fastforce | |
| 864 |         show "continuous_map (subtopology (prod_topology ?T01 ?X12) {p \<in> topspace (prod_topology ?T01 ?X12). snd p k \<le> 0})
 | |
| 865 | euclideanreal (\<lambda>x. h (fst x, snd x) i)" | |
| 866 | apply (simp add: h_def case_prod_unfold Let_def) | |
| 867 | apply (intro conjI impI fst snd continuous_intros) | |
| 868 | apply (auto simp: nsphere power2_eq_1_iff) | |
| 869 | done | |
| 870 | qed (auto simp: nsphere h) | |
| 871 | qed (auto simp: nsphere h) | |
| 872 | qed | |
| 873 | moreover | |
| 874 |   have "h ` ({0..1} \<times> (topspace (nsphere n) \<inter> {x. - (1/2) \<le> x k}))
 | |
| 875 |      \<subseteq> {x. (\<Sum>i\<le>n. (x i)\<^sup>2) = 1 \<and> (\<forall>i>n. x i = 0)}"
 | |
| 876 | proof - | |
| 877 | have "(\<Sum>i\<le>n. (h (T,x) i)\<^sup>2) = 1" | |
| 878 | if x: "x \<in> topspace (nsphere n)" and xk: "- (1/2) \<le> x k" and T: "0 \<le> T" "T \<le> 1" for T x | |
| 879 | proof (cases "-T \<le> x k ") | |
| 880 | case True | |
| 881 | then show ?thesis | |
| 882 | using that by (auto simp: nsphere h) | |
| 883 | next | |
| 884 | case False | |
| 885 | with x \<open>0 \<le> T\<close> have "k \<le> n" | |
| 886 | apply (simp add: nsphere) | |
| 887 | by (metis neg_le_0_iff_le not_le) | |
| 888 | have "1 - (x k)\<^sup>2 \<ge> 0" | |
| 889 | using topspace_nsphere_1 x by auto | |
| 890 | with False T \<open>k \<le> n\<close> | |
| 891 |       have "(\<Sum>i\<le>n. (h (T,x) i)\<^sup>2) = T\<^sup>2 + (1 - T\<^sup>2) * (\<Sum>i\<in>{..n} - {k}. (x i)\<^sup>2 / (1 - (x k)\<^sup>2))"
 | |
| 892 | unfolding h_def Let_def max_def | |
| 893 | by (simp add: not_le square_le_1 power_mult_distrib power_divide if_distrib [of "\<lambda>x. x ^ 2"] | |
| 894 | sum.delta_remove sum_distrib_left) | |
| 895 | also have "\<dots> = 1" | |
| 896 | using x False xk \<open>0 \<le> T\<close> | |
| 897 | by (simp add: nsphere sum_diff not_le \<open>k \<le> n\<close> power2_eq_1_iff flip: sum_divide_distrib) | |
| 898 | finally show ?thesis . | |
| 899 | qed | |
| 900 | moreover | |
| 901 | have "h (T,x) i = 0" | |
| 902 | if "x \<in> topspace (nsphere n)" "- (1/2) \<le> x k" and "n < i" "0 \<le> T" "T \<le> 1" | |
| 903 | for T x i | |
| 904 | proof (cases "-T \<le> x k ") | |
| 905 | case False | |
| 906 | then show ?thesis | |
| 907 | using that by (auto simp: nsphere h_def Let_def not_le max_def) | |
| 908 | qed (use that in \<open>auto simp: nsphere h\<close>) | |
| 909 | ultimately show ?thesis | |
| 910 | by auto | |
| 911 | qed | |
| 912 | ultimately | |
| 913 | have cmh: "continuous_map (prod_topology ?T01 ?X12) (nsphere n) h" | |
| 914 | by (subst (2) nsphere) (simp add: continuous_map_in_subtopology continuous_map_componentwise_UNIV) | |
| 915 |   have "hom_induced p (subtopology (nsphere n) {x. 0 \<le> x k})
 | |
| 916 |              (topspace (subtopology (nsphere n) {x. 0 \<le> x k}) \<inter> {x. x k = 0}) ?X12
 | |
| 917 |              (topspace ?X12 \<inter> {x. - 1/2 \<le> x k \<and> x k \<le> 0}) id
 | |
| 918 |             \<in> iso (relative_homology_group p (subtopology (nsphere n) {x. 0 \<le> x k})
 | |
| 919 |                        (topspace (subtopology (nsphere n) {x. 0 \<le> x k}) \<inter> {x. x k = 0}))
 | |
| 920 |                 (relative_homology_group p ?X12 (topspace ?X12 \<inter> {x. - 1/2 \<le> x k \<and> x k \<le> 0}))"
 | |
| 921 | proof (rule deformation_retract_relative_homology_group_isomorphism_id) | |
| 922 |     show "retraction_maps ?X12 (subtopology (nsphere n) {x. 0 \<le> x k}) (h \<circ> (\<lambda>x. (0,x))) id"
 | |
| 923 | unfolding retraction_maps_def | |
| 924 | proof (intro conjI ballI) | |
| 925 |       show "continuous_map ?X12 (subtopology (nsphere n) {x. 0 \<le> x k}) (h \<circ> Pair 0)"
 | |
| 926 | apply (simp add: continuous_map_in_subtopology) | |
| 927 | apply (intro conjI continuous_map_compose [OF _ cmh] continuous_intros) | |
| 928 | apply (auto simp: h_def Let_def) | |
| 929 | done | |
| 930 |       show "continuous_map (subtopology (nsphere n) {x. 0 \<le> x k}) ?X12 id"
 | |
| 931 | by (simp add: continuous_map_in_subtopology) (auto simp: nsphere) | |
| 932 | qed (simp add: nsphere h) | |
| 933 | next | |
| 934 | have h0: "\<And>xa. \<lbrakk>xa \<in> topspace (nsphere n); - (1/2) \<le> xa k; xa k \<le> 0\<rbrakk> \<Longrightarrow> h (0, xa) k = 0" | |
| 935 | by (simp add: h_def Let_def) | |
| 936 |     show "(h \<circ> (\<lambda>x. (0,x))) ` (topspace ?X12 \<inter> {x. - 1 / 2 \<le> x k \<and> x k \<le> 0})
 | |
| 937 |         \<subseteq> topspace (subtopology (nsphere n) {x. 0 \<le> x k}) \<inter> {x. x k = 0}"
 | |
| 938 | apply (auto simp: h0) | |
| 939 | apply (rule subsetD [OF continuous_map_image_subset_topspace [OF cmh]]) | |
| 940 | apply (force simp: nsphere) | |
| 941 | done | |
| 942 | have hin: "\<And>t x. \<lbrakk>x \<in> topspace (nsphere n); - (1/2) \<le> x k; 0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> h (t,x) \<in> topspace (nsphere n)" | |
| 943 | apply (rule subsetD [OF continuous_map_image_subset_topspace [OF cmh]]) | |
| 944 | apply (force simp: nsphere) | |
| 945 | done | |
| 946 | have h1: "\<And>x. \<lbrakk>x \<in> topspace (nsphere n); - (1/2) \<le> x k\<rbrakk> \<Longrightarrow> h (1, x) = x" | |
| 947 | by (simp add: h nsphere) | |
| 948 | have "continuous_map (prod_topology ?T01 ?X12) (nsphere n) h" | |
| 949 | using cmh by force | |
| 950 | then show "homotopic_with | |
| 951 |                  (\<lambda>h. h ` (topspace ?X12 \<inter> {x. - 1 / 2 \<le> x k \<and> x k \<le> 0}) \<subseteq> topspace ?X12 \<inter> {x. - 1 / 2 \<le> x k \<and> x k \<le> 0})
 | |
| 952 | ?X12 ?X12 (h \<circ> (\<lambda>x. (0,x))) id" | |
| 953 | apply (subst homotopic_with, force) | |
| 954 | apply (rule_tac x=h in exI) | |
| 955 | apply (auto simp: hin h1 continuous_map_in_subtopology) | |
| 956 | apply (auto simp: h_def Let_def max_def) | |
| 957 | done | |
| 958 | qed auto | |
| 959 |   then have 2: "hom_induced p (subtopology (nsphere n) {x. 0 \<le> x k}) {x. x k = 0}
 | |
| 960 |              ?X12 {x. - 1/2 \<le> x k \<and> x k \<le> 0} id
 | |
| 961 | \<in> Group.iso | |
| 962 |                 (relative_homology_group p (subtopology (nsphere n) {x. 0 \<le> x k}) {x. x k = 0})
 | |
| 963 |                 (relative_homology_group p ?X12 {x. - 1/2 \<le> x k \<and> x k \<le> 0})"
 | |
| 964 | by (metis hom_induced_restrict relative_homology_group_restrict topspace_subtopology) | |
| 965 | show ?thesis | |
| 966 | using iso_set_trans [OF 2 1] | |
| 967 | by (simp add: subset_iff continuous_map_in_subtopology flip: hom_induced_compose) | |
| 968 | qed | |
| 969 | ||
| 970 | ||
| 971 | corollary iso_upper_hemisphere_reduced_homology_group: | |
| 972 |    "(hom_boundary (1 + p) (subtopology (nsphere (Suc n)) {x. x(Suc n) \<ge> 0}) {x. x(Suc n) = 0})
 | |
| 973 |   \<in> iso (relative_homology_group (1 + p) (subtopology (nsphere (Suc n)) {x. x(Suc n) \<ge> 0}) {x. x(Suc n) = 0})
 | |
| 974 | (reduced_homology_group p (nsphere n))" | |
| 975 | proof - | |
| 976 |   have "{x. 0 \<le> x (Suc n)} \<inter> {x. x (Suc n) = 0} = {x. x (Suc n) = (0::real)}"
 | |
| 977 | by auto | |
| 978 |   then have n: "nsphere n = subtopology (subtopology (nsphere (Suc n)) {x. x(Suc n) \<ge> 0}) {x. x(Suc n) = 0}"
 | |
| 979 | by (simp add: subtopology_nsphere_equator subtopology_subtopology) | |
| 980 |   have ne: "(\<lambda>i. if i = n then 1 else 0) \<in> topspace (subtopology (nsphere (Suc n)) {x. 0 \<le> x (Suc n)}) \<inter> {x. x (Suc n) = 0}"
 | |
| 981 | by (simp add: nsphere if_distrib [of "\<lambda>x. x ^ 2"] cong: if_cong) | |
| 982 | show ?thesis | |
| 983 | unfolding n | |
| 984 | apply (rule iso_relative_homology_of_contractible [where p = "1 + p", simplified]) | |
| 985 | using contractible_space_upper_hemisphere ne apply blast+ | |
| 986 | done | |
| 987 | qed | |
| 988 | ||
| 989 | corollary iso_reduced_homology_group_upper_hemisphere: | |
| 990 | assumes "k \<le> n" | |
| 991 |   shows "hom_induced p (nsphere n) {} (nsphere n) {x. x k \<ge> 0} id
 | |
| 992 |       \<in> iso (reduced_homology_group p (nsphere n)) (relative_homology_group p (nsphere n) {x. x k \<ge> 0})"
 | |
| 993 | proof (rule iso_reduced_homology_by_contractible [OF contractible_space_upper_hemisphere [OF assms]]) | |
| 994 |   have "(\<lambda>i. if i = k then 1 else 0) \<in> topspace (nsphere n) \<inter> {x. 0 \<le> x k}"
 | |
| 995 | using assms by (simp add: nsphere if_distrib [of "\<lambda>x. x ^ 2"] cong: if_cong) | |
| 996 |   then show "topspace (nsphere n) \<inter> {x. 0 \<le> x k} \<noteq> {}"
 | |
| 997 | by blast | |
| 998 | qed | |
| 999 | ||
| 1000 | ||
| 1001 | lemma iso_relative_homology_group_lower_hemisphere: | |
| 1002 |   "hom_induced p (subtopology (nsphere n) {x. x k \<le> 0}) {x. x k = 0} (nsphere n) {x. x k \<ge> 0} id
 | |
| 1003 |   \<in> iso (relative_homology_group p (subtopology (nsphere n) {x. x k \<le> 0}) {x. x k = 0})
 | |
| 1004 |         (relative_homology_group p (nsphere n) {x. x k \<ge> 0})" (is "?k \<in> iso ?G ?H")
 | |
| 1005 | proof - | |
| 1006 | define r where "r \<equiv> \<lambda>x i. if i = k then -x i else (x i::real)" | |
| 1007 | then have [simp]: "r \<circ> r = id" | |
| 1008 | by force | |
| 1009 | have cmr: "continuous_map (subtopology (nsphere n) S) (nsphere n) r" for S | |
| 1010 | using continuous_map_nsphere_reflection [of n k] | |
| 1011 | by (simp add: continuous_map_from_subtopology r_def) | |
| 1012 |   let ?f = "hom_induced p (subtopology (nsphere n) {x. x k \<le> 0}) {x. x k = 0}
 | |
| 1013 |                           (subtopology (nsphere n) {x. x k \<ge> 0}) {x. x k = 0} r"
 | |
| 1014 |   let ?g = "hom_induced p (subtopology (nsphere n) {x. x k \<ge> 0}) {x. x k = 0} (nsphere n) {x. x k \<le> 0} id"
 | |
| 1015 |   let ?h = "hom_induced p (nsphere n) {x. x k \<le> 0} (nsphere n) {x. x k \<ge> 0} r"
 | |
| 1016 | obtain f h where | |
| 1017 |         f: "f \<in> iso ?G (relative_homology_group p (subtopology (nsphere n) {x. x k \<ge> 0}) {x. x k = 0})"
 | |
| 1018 |     and h: "h \<in> iso (relative_homology_group p (nsphere n) {x. x k \<le> 0}) ?H"
 | |
| 1019 | and eq: "h \<circ> ?g \<circ> f = ?k" | |
| 1020 | proof | |
| 1021 | have hmr: "homeomorphic_map (nsphere n) (nsphere n) r" | |
| 1022 | unfolding homeomorphic_map_maps | |
| 1023 | by (metis \<open>r \<circ> r = id\<close> cmr homeomorphic_maps_involution pointfree_idE subtopology_topspace) | |
| 1024 |     then have hmrs: "homeomorphic_map (subtopology (nsphere n) {x. x k \<le> 0}) (subtopology (nsphere n) {x. x k \<ge> 0}) r"
 | |
| 1025 | by (simp add: homeomorphic_map_subtopologies_alt r_def) | |
| 1026 |     have rimeq: "r ` (topspace (subtopology (nsphere n) {x. x k \<le> 0}) \<inter> {x. x k = 0})
 | |
| 1027 |                = topspace (subtopology (nsphere n) {x. 0 \<le> x k}) \<inter> {x. x k = 0}"
 | |
| 1028 | using continuous_map_eq_topcontinuous_at continuous_map_nsphere_reflection topcontinuous_at_atin | |
| 1029 | by (fastforce simp: r_def) | |
| 1030 |     show "?f \<in> iso ?G (relative_homology_group p (subtopology (nsphere n) {x. x k \<ge> 0}) {x. x k = 0})"
 | |
| 1031 | using homeomorphic_map_relative_homology_iso [OF hmrs Int_lower1 rimeq] | |
| 1032 | by (metis hom_induced_restrict relative_homology_group_restrict) | |
| 1033 |     have rimeq: "r ` (topspace (nsphere n) \<inter> {x. x k \<le> 0}) = topspace (nsphere n) \<inter> {x. 0 \<le> x k}"
 | |
| 1034 | by (metis hmrs homeomorphic_imp_surjective_map topspace_subtopology) | |
| 1035 |     show "?h \<in> Group.iso (relative_homology_group p (nsphere n) {x. x k \<le> 0}) ?H"
 | |
| 1036 | using homeomorphic_map_relative_homology_iso [OF hmr Int_lower1 rimeq] by simp | |
| 1037 | have [simp]: "\<And>x. x k = 0 \<Longrightarrow> r x k = 0" | |
| 1038 | by (auto simp: r_def) | |
| 1039 | have "?h \<circ> ?g \<circ> ?f | |
| 1040 |         = hom_induced p (subtopology (nsphere n) {x. 0 \<le> x k}) {x. x k = 0} (nsphere n) {x. 0 \<le> x k} r \<circ>
 | |
| 1041 |           hom_induced p (subtopology (nsphere n) {x. x k \<le> 0}) {x. x k = 0} (subtopology (nsphere n) {x. 0 \<le> x k}) {x. x k = 0} r"
 | |
| 1042 | apply (subst hom_induced_compose [symmetric]) | |
| 1043 | using continuous_map_nsphere_reflection apply (force simp: r_def)+ | |
| 1044 | done | |
| 1045 | also have "\<dots> = ?k" | |
| 1046 | apply (subst hom_induced_compose [symmetric]) | |
| 1047 | apply (simp_all add: image_subset_iff cmr) | |
| 1048 | using hmrs homeomorphic_imp_continuous_map apply blast | |
| 1049 | done | |
| 1050 | finally show "?h \<circ> ?g \<circ> ?f = ?k" . | |
| 1051 | qed | |
| 1052 | with iso_relative_homology_group_upper_hemisphere [of p n k] | |
| 1053 |   have "h \<circ> hom_induced p (subtopology (nsphere n) {f. 0 \<le> f k}) {f. f k = 0} (nsphere n) {f. f k \<le> 0} id \<circ> f
 | |
| 1054 |   \<in> Group.iso ?G (relative_homology_group p (nsphere n) {f. 0 \<le> f k})"
 | |
| 1055 | using f h iso_set_trans by blast | |
| 1056 | then show ?thesis | |
| 1057 | by (simp add: eq) | |
| 1058 | qed | |
| 1059 | ||
| 1060 | ||
| 1061 | lemma iso_lower_hemisphere_reduced_homology_group: | |
| 1062 |    "hom_boundary (1 + p) (subtopology (nsphere (Suc n)) {x. x(Suc n) \<le> 0}) {x. x(Suc n) = 0}
 | |
| 1063 |   \<in> iso (relative_homology_group (1 + p) (subtopology (nsphere (Suc n)) {x. x(Suc n) \<le> 0})
 | |
| 1064 |                         {x. x(Suc n) = 0})
 | |
| 1065 | (reduced_homology_group p (nsphere n))" | |
| 1066 | proof - | |
| 1067 |   have "{x. (\<Sum>i\<le>n. (x i)\<^sup>2) = 1 \<and> (\<forall>i>n. x i = 0)} =
 | |
| 1068 |        ({x. (\<Sum>i\<le>n. (x i)\<^sup>2) + (x (Suc n))\<^sup>2 = 1 \<and> (\<forall>i>Suc n. x i = 0)} \<inter> {x. x (Suc n) \<le> 0} \<inter>
 | |
| 1069 |         {x. x (Suc n) = (0::real)})"
 | |
| 1070 | by (force simp: dest: Suc_lessI) | |
| 1071 |   then have n: "nsphere n = subtopology (subtopology (nsphere (Suc n)) {x. x(Suc n) \<le> 0}) {x. x(Suc n) = 0}"
 | |
| 1072 | by (simp add: nsphere subtopology_subtopology) | |
| 1073 |   have ne: "(\<lambda>i. if i = n then 1 else 0) \<in> topspace (subtopology (nsphere (Suc n)) {x. x (Suc n) \<le> 0}) \<inter> {x. x (Suc n) = 0}"
 | |
| 1074 | by (simp add: nsphere if_distrib [of "\<lambda>x. x ^ 2"] cong: if_cong) | |
| 1075 | show ?thesis | |
| 1076 | unfolding n | |
| 1077 | apply (rule iso_relative_homology_of_contractible [where p = "1 + p", simplified]) | |
| 1078 | using contractible_space_lower_hemisphere ne apply blast+ | |
| 1079 | done | |
| 1080 | qed | |
| 1081 | ||
| 1082 | lemma isomorphism_sym: | |
| 1083 | "\<lbrakk>f \<in> iso G1 G2; \<And>x. x \<in> carrier G1 \<Longrightarrow> r'(f x) = f(r x); | |
| 1084 | \<And>x. x \<in> carrier G1 \<Longrightarrow> r x \<in> carrier G1; group G1; group G2\<rbrakk> | |
| 1085 | \<Longrightarrow> \<exists>f \<in> iso G2 G1. \<forall>x \<in> carrier G2. r(f x) = f(r' x)" | |
| 1086 | apply (clarsimp simp add: group.iso_iff_group_isomorphisms Bex_def) | |
| 1087 | by (metis (full_types) group_isomorphisms_def group_isomorphisms_sym hom_in_carrier) | |
| 1088 | ||
| 1089 | lemma isomorphism_trans: | |
| 1090 | "\<lbrakk>\<exists>f \<in> iso G1 G2. \<forall>x \<in> carrier G1. r2(f x) = f(r1 x); \<exists>f \<in> iso G2 G3. \<forall>x \<in> carrier G2. r3(f x) = f(r2 x)\<rbrakk> | |
| 1091 | \<Longrightarrow> \<exists>f \<in> iso G1 G3. \<forall>x \<in> carrier G1. r3(f x) = f(r1 x)" | |
| 1092 | apply clarify | |
| 1093 | apply (rename_tac g f) | |
| 1094 | apply (rule_tac x="f \<circ> g" in bexI) | |
| 1095 | apply (metis iso_iff comp_apply hom_in_carrier) | |
| 1096 | using iso_set_trans by blast | |
| 1097 | ||
| 1098 | lemma reduced_homology_group_nsphere_step: | |
| 1099 | "\<exists>f \<in> iso(reduced_homology_group p (nsphere n)) | |
| 1100 | (reduced_homology_group (1 + p) (nsphere (Suc n))). | |
| 1101 | \<forall>c \<in> carrier(reduced_homology_group p (nsphere n)). | |
| 1102 |              hom_induced (1 + p) (nsphere(Suc n)) {} (nsphere(Suc n)) {}
 | |
| 1103 | (\<lambda>x i. if i = 0 then -x i else x i) (f c) | |
| 1104 |            = f (hom_induced p (nsphere n) {} (nsphere n) {} (\<lambda>x i. if i = 0 then -x i else x i) c)"
 | |
| 1105 | proof - | |
| 1106 | define r where "r \<equiv> \<lambda>x::nat\<Rightarrow>real. \<lambda>i. if i = 0 then -x i else x i" | |
| 1107 | have cmr: "continuous_map (nsphere n) (nsphere n) r" for n | |
| 1108 | unfolding r_def by (rule continuous_map_nsphere_reflection) | |
| 1109 |   have rsub: "r ` {x. 0 \<le> x (Suc n)} \<subseteq> {x. 0 \<le> x (Suc n)}" "r ` {x. x (Suc n) \<le> 0} \<subseteq> {x. x (Suc n) \<le> 0}" "r ` {x. x (Suc n) = 0} \<subseteq> {x. x (Suc n) = 0}"
 | |
| 1110 | by (force simp: r_def)+ | |
| 1111 |   let ?sub = "subtopology (nsphere (Suc n)) {x. x (Suc n) \<ge> 0}"
 | |
| 1112 |   let ?G2 = "relative_homology_group (1 + p) ?sub {x. x (Suc n) = 0}"
 | |
| 1113 |   let ?r2 = "hom_induced (1 + p) ?sub {x. x (Suc n) = 0} ?sub {x. x (Suc n) = 0} r"
 | |
| 1114 |   let ?j = "\<lambda>p n. hom_induced p (nsphere n) {} (nsphere n) {} r"
 | |
| 1115 | show ?thesis | |
| 1116 | unfolding r_def [symmetric] | |
| 1117 | proof (rule isomorphism_trans) | |
| 1118 |     let ?f = "hom_boundary (1 + p) ?sub {x. x (Suc n) = 0}"
 | |
| 1119 | show "\<exists>f\<in>Group.iso (reduced_homology_group p (nsphere n)) ?G2. | |
| 1120 | \<forall>c\<in>carrier (reduced_homology_group p (nsphere n)). ?r2 (f c) = f (?j p n c)" | |
| 1121 | proof (rule isomorphism_sym) | |
| 1122 | show "?f \<in> Group.iso ?G2 (reduced_homology_group p (nsphere n))" | |
| 1123 | using iso_upper_hemisphere_reduced_homology_group | |
| 1124 | by (metis add.commute) | |
| 1125 | next | |
| 1126 | fix c | |
| 1127 | assume "c \<in> carrier ?G2" | |
| 1128 | have cmrs: "continuous_map ?sub ?sub r" | |
| 1129 | by (metis (mono_tags, lifting) IntE cmr continuous_map_from_subtopology continuous_map_in_subtopology image_subset_iff rsub(1) topspace_subtopology) | |
| 1130 |       have "hom_induced p (nsphere n) {} (nsphere n) {} r \<circ> hom_boundary (1 + p) ?sub {x. x (Suc n) = 0}
 | |
| 1131 |           = hom_boundary (1 + p) ?sub {x. x (Suc n) = 0} \<circ>
 | |
| 1132 |             hom_induced (1 + p) ?sub {x. x (Suc n) = 0} ?sub {x. x (Suc n) = 0} r"
 | |
| 1133 | using naturality_hom_induced [OF cmrs rsub(3), symmetric, of "1+p", simplified] | |
| 1134 | by (simp add: subtopology_subtopology subtopology_nsphere_equator flip: Collect_conj_eq cong: rev_conj_cong) | |
| 1135 |       then show "?j p n (?f c) = ?f (hom_induced (1 + p) ?sub {x. x (Suc n) = 0} ?sub {x. x (Suc n) = 0} r c)"
 | |
| 1136 | by (metis comp_def) | |
| 1137 | next | |
| 1138 | fix c | |
| 1139 | assume "c \<in> carrier ?G2" | |
| 1140 |       show "hom_induced (1 + p) ?sub {x. x (Suc n) = 0} ?sub {x. x (Suc n) = 0} r c \<in> carrier ?G2"
 | |
| 1141 | using hom_induced_carrier by blast | |
| 1142 | qed auto | |
| 1143 | next | |
| 1144 |     let ?H2 = "relative_homology_group (1 + p) (nsphere (Suc n)) {x. x (Suc n) \<le> 0}"
 | |
| 1145 |     let ?s2 = "hom_induced (1 + p) (nsphere (Suc n)) {x. x (Suc n) \<le> 0} (nsphere (Suc n)) {x. x (Suc n) \<le> 0} r"
 | |
| 1146 | show "\<exists>f\<in>Group.iso ?G2 (reduced_homology_group (1 + p) (nsphere (Suc n))). \<forall>c\<in>carrier ?G2. ?j (1 + p) (Suc n) (f c) | |
| 1147 | = f (?r2 c)" | |
| 1148 | proof (rule isomorphism_trans) | |
| 1149 | show "\<exists>f\<in>Group.iso ?G2 ?H2. | |
| 1150 | \<forall>c\<in>carrier ?G2. | |
| 1151 |                     ?s2 (f c) = f (hom_induced (1 + p) ?sub {x. x (Suc n) = 0} ?sub {x. x (Suc n) = 0} r c)"
 | |
| 1152 | proof (intro ballI bexI) | |
| 1153 | fix c | |
| 1154 |         assume "c \<in> carrier (relative_homology_group (1 + p) ?sub {x. x (Suc n) = 0})"
 | |
| 1155 |         show "?s2 (hom_induced (1 + p) ?sub {x. x (Suc n) = 0} (nsphere (Suc n)) {x. x (Suc n) \<le> 0} id c)
 | |
| 1156 |             = hom_induced (1 + p) ?sub {x. x (Suc n) = 0} (nsphere (Suc n)) {x. x (Suc n) \<le> 0} id (?r2 c)"
 | |
| 1157 | apply (simp add: rsub hom_induced_compose' Collect_mono_iff cmr) | |
| 1158 | apply (subst hom_induced_compose') | |
| 1159 | apply (simp_all add: continuous_map_in_subtopology continuous_map_from_subtopology [OF cmr] rsub) | |
| 1160 | apply (auto simp: r_def) | |
| 1161 | done | |
| 1162 | qed (simp add: iso_relative_homology_group_upper_hemisphere) | |
| 1163 | next | |
| 1164 |       let ?h = "hom_induced (1 + p) (nsphere(Suc n)) {} (nsphere (Suc n)) {x. x(Suc n) \<le> 0} id"
 | |
| 1165 | show "\<exists>f\<in>Group.iso ?H2 (reduced_homology_group (1 + p) (nsphere (Suc n))). | |
| 1166 | \<forall>c\<in>carrier ?H2. ?j (1 + p) (Suc n) (f c) = f (?s2 c)" | |
| 1167 | proof (rule isomorphism_sym) | |
| 1168 | show "?h \<in> Group.iso (reduced_homology_group (1 + p) (nsphere (Suc n))) | |
| 1169 |                (relative_homology_group (1 + p) (nsphere (Suc n)) {x. x (Suc n) \<le> 0})"
 | |
| 1170 | using iso_reduced_homology_group_lower_hemisphere by blast | |
| 1171 | next | |
| 1172 | fix c | |
| 1173 | assume "c \<in> carrier (reduced_homology_group (1 + p) (nsphere (Suc n)))" | |
| 1174 | show "?s2 (?h c) = ?h (?j (1 + p) (Suc n) c)" | |
| 1175 | by (simp add: hom_induced_compose' cmr rsub) | |
| 1176 | next | |
| 1177 | fix c | |
| 1178 | assume "c \<in> carrier (reduced_homology_group (1 + p) (nsphere (Suc n)))" | |
| 1179 |         then show "hom_induced (1 + p) (nsphere (Suc n)) {} (nsphere (Suc n)) {} r c
 | |
| 1180 | \<in> carrier (reduced_homology_group (1 + p) (nsphere (Suc n)))" | |
| 1181 | by (simp add: hom_induced_reduced) | |
| 1182 | qed auto | |
| 1183 | qed | |
| 1184 | qed | |
| 1185 | qed | |
| 1186 | ||
| 1187 | ||
| 1188 | lemma reduced_homology_group_nsphere_aux: | |
| 1189 | "if p = int n then reduced_homology_group n (nsphere n) \<cong> integer_group | |
| 1190 | else trivial_group(reduced_homology_group p (nsphere n))" | |
| 1191 | proof (induction n arbitrary: p) | |
| 1192 | case 0 | |
| 1193 | let ?a = "\<lambda>i::nat. if i = 0 then 1 else (0::real)" | |
| 1194 | let ?b = "\<lambda>i::nat. if i = 0 then -1 else (0::real)" | |
| 1195 |   have st: "subtopology (powertop_real UNIV) {?a, ?b} = nsphere 0"
 | |
| 1196 | proof - | |
| 1197 |     have "{?a, ?b} = {x. (x 0)\<^sup>2 = 1 \<and> (\<forall>i>0. x i = 0)}"
 | |
| 1198 | using power2_eq_iff by fastforce | |
| 1199 | then show ?thesis | |
| 1200 | by (simp add: nsphere) | |
| 1201 | qed | |
| 1202 |   have *: "reduced_homology_group p (subtopology (powertop_real UNIV) {?a, ?b}) \<cong>
 | |
| 1203 |         homology_group p (subtopology (powertop_real UNIV) {?a})"
 | |
| 1204 | apply (rule reduced_homology_group_pair) | |
| 1205 | apply (simp_all add: fun_eq_iff) | |
| 1206 | apply (simp add: open_fun_def separation_t1 t1_space_def) | |
| 1207 | done | |
| 1208 | have "reduced_homology_group 0 (nsphere 0) \<cong> integer_group" if "p=0" | |
| 1209 | proof - | |
| 1210 |     have "reduced_homology_group 0 (nsphere 0) \<cong> homology_group 0 (top_of_set {?a})" if "p=0"
 | |
| 1211 | by (metis * euclidean_product_topology st that) | |
| 1212 | also have "\<dots> \<cong> integer_group" | |
| 1213 | by (simp add: homology_coefficients) | |
| 1214 | finally show ?thesis | |
| 1215 | using that by blast | |
| 1216 | qed | |
| 1217 | moreover have "trivial_group (reduced_homology_group p (nsphere 0))" if "p\<noteq>0" | |
| 1218 |     using * that homology_dimension_axiom [of "subtopology (powertop_real UNIV) {?a}" ?a p]
 | |
| 1219 | using isomorphic_group_triviality st by force | |
| 1220 | ultimately show ?case | |
| 1221 | by auto | |
| 1222 | next | |
| 1223 | case (Suc n) | |
| 1224 | have eq: "reduced_homology_group (int n) (nsphere n) \<cong> integer_group" if "p-1 = n" | |
| 1225 | by (simp add: Suc.IH) | |
| 1226 | have neq: "trivial_group (reduced_homology_group (p-1) (nsphere n))" if "p-1 \<noteq> n" | |
| 1227 | by (simp add: Suc.IH that) | |
| 1228 | have iso: "reduced_homology_group p (nsphere (Suc n)) \<cong> reduced_homology_group (p-1) (nsphere n)" | |
| 1229 | using reduced_homology_group_nsphere_step [of "p-1" n] group.iso_sym [OF _ is_isoI] group_reduced_homology_group | |
| 1230 | by fastforce | |
| 1231 | then show ?case | |
| 1232 | using eq iso_trans iso isomorphic_group_triviality neq | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
72632diff
changeset | 1233 | by (metis (no_types, opaque_lifting) add.commute add_left_cancel diff_add_cancel group_reduced_homology_group of_nat_Suc) | 
| 70095 | 1234 | qed | 
| 1235 | ||
| 1236 | ||
| 1237 | lemma reduced_homology_group_nsphere: | |
| 1238 | "reduced_homology_group n (nsphere n) \<cong> integer_group" | |
| 1239 | "p \<noteq> n \<Longrightarrow> trivial_group(reduced_homology_group p (nsphere n))" | |
| 1240 | using reduced_homology_group_nsphere_aux by auto | |
| 1241 | ||
| 1242 | lemma cyclic_reduced_homology_group_nsphere: | |
| 1243 | "cyclic_group(reduced_homology_group p (nsphere n))" | |
| 1244 | by (metis reduced_homology_group_nsphere trivial_imp_cyclic_group cyclic_integer_group | |
| 1245 | group_integer_group group_reduced_homology_group isomorphic_group_cyclicity) | |
| 1246 | ||
| 1247 | lemma trivial_reduced_homology_group_nsphere: | |
| 1248 | "trivial_group(reduced_homology_group p (nsphere n)) \<longleftrightarrow> (p \<noteq> n)" | |
| 1249 | using group_integer_group isomorphic_group_triviality nontrivial_integer_group reduced_homology_group_nsphere(1) reduced_homology_group_nsphere(2) trivial_group_def by blast | |
| 1250 | ||
| 1251 | lemma non_contractible_space_nsphere: "\<not> (contractible_space(nsphere n))" | |
| 1252 | proof (clarsimp simp add: contractible_eq_homotopy_equivalent_singleton_subtopology) | |
| 1253 | fix a :: "nat \<Rightarrow> real" | |
| 1254 | assume a: "a \<in> topspace (nsphere n)" | |
| 1255 |     and he: "nsphere n homotopy_equivalent_space subtopology (nsphere n) {a}"
 | |
| 1256 |   have "trivial_group (reduced_homology_group (int n) (subtopology (nsphere n) {a}))"
 | |
| 1257 | by (simp add: a homology_dimension_reduced [where a=a]) | |
| 1258 | then show "False" | |
| 1259 | using isomorphic_group_triviality [OF homotopy_equivalent_space_imp_isomorphic_reduced_homology_groups [OF he, of n]] | |
| 1260 | by (simp add: trivial_reduced_homology_group_nsphere) | |
| 1261 | qed | |
| 1262 | ||
| 1263 | ||
| 1264 | subsection\<open>Brouwer degree of a Map\<close> | |
| 1265 | ||
| 1266 | definition Brouwer_degree2 :: "nat \<Rightarrow> ((nat \<Rightarrow> real) \<Rightarrow> nat \<Rightarrow> real) \<Rightarrow> int" | |
| 1267 | where | |
| 1268 | "Brouwer_degree2 p f \<equiv> | |
| 1269 | @d::int. \<forall>x \<in> carrier(reduced_homology_group p (nsphere p)). | |
| 1270 |                 hom_induced p (nsphere p) {} (nsphere p) {} f x = pow (reduced_homology_group p (nsphere p)) x d"
 | |
| 1271 | ||
| 1272 | lemma Brouwer_degree2_eq: | |
| 1273 | "(\<And>x. x \<in> topspace(nsphere p) \<Longrightarrow> f x = g x) \<Longrightarrow> Brouwer_degree2 p f = Brouwer_degree2 p g" | |
| 1274 | unfolding Brouwer_degree2_def Ball_def | |
| 1275 | apply (intro Eps_cong all_cong) | |
| 1276 | by (metis (mono_tags, lifting) hom_induced_eq) | |
| 1277 | ||
| 1278 | lemma Brouwer_degree2: | |
| 1279 | assumes "x \<in> carrier(reduced_homology_group p (nsphere p))" | |
| 1280 |   shows "hom_induced p (nsphere p) {} (nsphere p) {} f x
 | |
| 1281 | = pow (reduced_homology_group p (nsphere p)) x (Brouwer_degree2 p f)" | |
| 1282 | (is "?h x = pow ?G x _") | |
| 1283 | proof (cases "continuous_map(nsphere p) (nsphere p) f") | |
| 1284 | case True | |
| 1285 | interpret group ?G | |
| 1286 | by simp | |
| 1287 | interpret group_hom ?G ?G ?h | |
| 1288 | using hom_induced_reduced_hom group_hom_axioms_def group_hom_def is_group by blast | |
| 1289 | obtain a where a: "a \<in> carrier ?G" | |
| 1290 |     and aeq: "subgroup_generated ?G {a} = ?G"
 | |
| 1291 | using cyclic_reduced_homology_group_nsphere [of p p] by (auto simp: cyclic_group_def) | |
| 1292 |   then have carra: "carrier (subgroup_generated ?G {a}) = range (\<lambda>n::int. pow ?G a n)"
 | |
| 1293 | using carrier_subgroup_generated_by_singleton by blast | |
| 1294 |   moreover have "?h a \<in> carrier (subgroup_generated ?G {a})"
 | |
| 1295 | by (simp add: a aeq hom_induced_reduced) | |
| 1296 | ultimately obtain d::int where d: "?h a = pow ?G a d" | |
| 1297 | by auto | |
| 1298 |   have *: "hom_induced (int p) (nsphere p) {} (nsphere p) {} f x = x [^]\<^bsub>?G\<^esub> d"
 | |
| 1299 | if x: "x \<in> carrier ?G" for x | |
| 1300 | proof - | |
| 1301 | obtain n::int where xeq: "x = pow ?G a n" | |
| 1302 | using carra x aeq by moura | |
| 1303 | show ?thesis | |
| 1304 | by (simp add: xeq a d hom_int_pow int_pow_pow mult.commute) | |
| 1305 | qed | |
| 1306 | show ?thesis | |
| 1307 | unfolding Brouwer_degree2_def | |
| 1308 | apply (rule someI2 [where a=d]) | |
| 1309 | using assms * apply blast+ | |
| 1310 | done | |
| 1311 | next | |
| 1312 | case False | |
| 1313 | show ?thesis | |
| 1314 | unfolding Brouwer_degree2_def | |
| 1315 | by (rule someI2 [where a=0]) (simp_all add: hom_induced_default False one_reduced_homology_group assms) | |
| 1316 | qed | |
| 1317 | ||
| 1318 | ||
| 1319 | ||
| 1320 | lemma Brouwer_degree2_iff: | |
| 1321 | assumes f: "continuous_map (nsphere p) (nsphere p) f" | |
| 1322 | and x: "x \<in> carrier(reduced_homology_group p (nsphere p))" | |
| 1323 |   shows "(hom_induced (int p) (nsphere p) {} (nsphere p) {} f x =
 | |
| 1324 | x [^]\<^bsub>reduced_homology_group (int p) (nsphere p)\<^esub> d) | |
| 1325 | \<longleftrightarrow> (x = \<one>\<^bsub>reduced_homology_group (int p) (nsphere p)\<^esub> \<or> Brouwer_degree2 p f = d)" | |
| 1326 | (is "(?h x = x [^]\<^bsub>?G\<^esub> d) \<longleftrightarrow> _") | |
| 1327 | proof - | |
| 1328 | interpret group "?G" | |
| 1329 | by simp | |
| 1330 | obtain a where a: "a \<in> carrier ?G" | |
| 1331 |     and aeq: "subgroup_generated ?G {a} = ?G"
 | |
| 1332 | using cyclic_reduced_homology_group_nsphere [of p p] by (auto simp: cyclic_group_def) | |
| 1333 | then obtain i::int where i: "x = (a [^]\<^bsub>?G\<^esub> i)" | |
| 1334 | using carrier_subgroup_generated_by_singleton x by fastforce | |
| 1335 | then have "a [^]\<^bsub>?G\<^esub> i \<in> carrier ?G" | |
| 1336 | using x by blast | |
| 1337 | have [simp]: "ord a = 0" | |
| 1338 | by (simp add: a aeq iso_finite [OF reduced_homology_group_nsphere(1)] flip: infinite_cyclic_subgroup_order) | |
| 1339 | show ?thesis | |
| 1340 | by (auto simp: Brouwer_degree2 int_pow_eq_id x i a int_pow_pow int_pow_eq) | |
| 1341 | qed | |
| 1342 | ||
| 1343 | ||
| 1344 | lemma Brouwer_degree2_unique: | |
| 1345 | assumes f: "continuous_map (nsphere p) (nsphere p) f" | |
| 1346 | and hi: "\<And>x. x \<in> carrier(reduced_homology_group p (nsphere p)) | |
| 1347 |                \<Longrightarrow> hom_induced p (nsphere p) {} (nsphere p) {} f x = pow (reduced_homology_group p (nsphere p)) x d"
 | |
| 1348 | (is "\<And>x. x \<in> carrier ?G \<Longrightarrow> ?h x = _") | |
| 1349 | shows "Brouwer_degree2 p f = d" | |
| 1350 | proof - | |
| 1351 | obtain a where a: "a \<in> carrier ?G" | |
| 1352 |     and aeq: "subgroup_generated ?G {a} = ?G"
 | |
| 1353 | using cyclic_reduced_homology_group_nsphere [of p p] by (auto simp: cyclic_group_def) | |
| 1354 | show ?thesis | |
| 1355 | using hi [OF a] | |
| 1356 | apply (simp add: Brouwer_degree2 a) | |
| 1357 | by (metis Brouwer_degree2_iff a aeq f group.trivial_group_subgroup_generated group_reduced_homology_group subsetI trivial_reduced_homology_group_nsphere) | |
| 1358 | qed | |
| 1359 | ||
| 1360 | lemma Brouwer_degree2_unique_generator: | |
| 1361 | assumes f: "continuous_map (nsphere p) (nsphere p) f" | |
| 1362 |     and eq: "subgroup_generated (reduced_homology_group p (nsphere p)) {a}
 | |
| 1363 | = reduced_homology_group p (nsphere p)" | |
| 1364 |     and hi: "hom_induced p (nsphere p) {} (nsphere p) {} f a = pow (reduced_homology_group p (nsphere p)) a d"
 | |
| 1365 | (is "?h a = pow ?G a _") | |
| 1366 | shows "Brouwer_degree2 p f = d" | |
| 1367 | proof (cases "a \<in> carrier ?G") | |
| 1368 | case True | |
| 1369 | then show ?thesis | |
| 1370 | by (metis Brouwer_degree2_iff hi eq f group.trivial_group_subgroup_generated group_reduced_homology_group | |
| 1371 | subset_singleton_iff trivial_reduced_homology_group_nsphere) | |
| 1372 | next | |
| 1373 | case False | |
| 1374 | then show ?thesis | |
| 1375 | using trivial_reduced_homology_group_nsphere [of p p] | |
| 1376 | by (metis group.trivial_group_subgroup_generated_eq disjoint_insert(1) eq group_reduced_homology_group inf_bot_right subset_singleton_iff) | |
| 1377 | qed | |
| 1378 | ||
| 1379 | lemma Brouwer_degree2_homotopic: | |
| 1380 | assumes "homotopic_with (\<lambda>x. True) (nsphere p) (nsphere p) f g" | |
| 1381 | shows "Brouwer_degree2 p f = Brouwer_degree2 p g" | |
| 1382 | proof - | |
| 1383 | have "continuous_map (nsphere p) (nsphere p) f" | |
| 1384 | using homotopic_with_imp_continuous_maps [OF assms] by auto | |
| 1385 | show ?thesis | |
| 1386 | using Brouwer_degree2_def assms homology_homotopy_empty by fastforce | |
| 1387 | qed | |
| 1388 | ||
| 1389 | lemma Brouwer_degree2_id [simp]: "Brouwer_degree2 p id = 1" | |
| 1390 | proof (rule Brouwer_degree2_unique) | |
| 1391 | fix x | |
| 1392 | assume x: "x \<in> carrier (reduced_homology_group (int p) (nsphere p))" | |
| 1393 | then have "x \<in> carrier (homology_group (int p) (nsphere p))" | |
| 1394 | using carrier_reduced_homology_group_subset by blast | |
| 1395 |   then show "hom_induced (int p) (nsphere p) {} (nsphere p) {} id x =
 | |
| 1396 | x [^]\<^bsub>reduced_homology_group (int p) (nsphere p)\<^esub> (1::int)" | |
| 1397 | by (simp add: hom_induced_id group.int_pow_1 x) | |
| 1398 | qed auto | |
| 1399 | ||
| 1400 | lemma Brouwer_degree2_compose: | |
| 1401 | assumes f: "continuous_map (nsphere p) (nsphere p) f" and g: "continuous_map (nsphere p) (nsphere p) g" | |
| 1402 | shows "Brouwer_degree2 p (g \<circ> f) = Brouwer_degree2 p g * Brouwer_degree2 p f" | |
| 1403 | proof (rule Brouwer_degree2_unique) | |
| 1404 | show "continuous_map (nsphere p) (nsphere p) (g \<circ> f)" | |
| 1405 | by (meson continuous_map_compose f g) | |
| 1406 | next | |
| 1407 | fix x | |
| 1408 | assume x: "x \<in> carrier (reduced_homology_group (int p) (nsphere p))" | |
| 1409 |   have "hom_induced (int p) (nsphere p) {} (nsphere p) {} (g \<circ> f) =
 | |
| 1410 |              hom_induced (int p) (nsphere p) {} (nsphere p) {} g \<circ>
 | |
| 1411 |              hom_induced (int p) (nsphere p) {} (nsphere p) {} f"
 | |
| 1412 | by (blast intro: hom_induced_compose [OF f _ g]) | |
| 1413 |   with x show "hom_induced (int p) (nsphere p) {} (nsphere p) {} (g \<circ> f) x =
 | |
| 1414 | x [^]\<^bsub>reduced_homology_group (int p) (nsphere p)\<^esub> (Brouwer_degree2 p g * Brouwer_degree2 p f)" | |
| 1415 | by (simp add: mult.commute hom_induced_reduced flip: Brouwer_degree2 group.int_pow_pow) | |
| 1416 | qed | |
| 1417 | ||
| 1418 | lemma Brouwer_degree2_homotopy_equivalence: | |
| 1419 | assumes f: "continuous_map (nsphere p) (nsphere p) f" and g: "continuous_map (nsphere p) (nsphere p) g" | |
| 1420 | and hom: "homotopic_with (\<lambda>x. True) (nsphere p) (nsphere p) (f \<circ> g) id" | |
| 1421 | obtains "\<bar>Brouwer_degree2 p f\<bar> = 1" "\<bar>Brouwer_degree2 p g\<bar> = 1" "Brouwer_degree2 p g = Brouwer_degree2 p f" | |
| 1422 | using Brouwer_degree2_homotopic [OF hom] Brouwer_degree2_compose f g zmult_eq_1_iff by auto | |
| 1423 | ||
| 1424 | lemma Brouwer_degree2_homeomorphic_maps: | |
| 1425 | assumes "homeomorphic_maps (nsphere p) (nsphere p) f g" | |
| 1426 | obtains "\<bar>Brouwer_degree2 p f\<bar> = 1" "\<bar>Brouwer_degree2 p g\<bar> = 1" "Brouwer_degree2 p g = Brouwer_degree2 p f" | |
| 1427 | using assms | |
| 1428 | by (auto simp: homeomorphic_maps_def homotopic_with_equal continuous_map_compose intro: Brouwer_degree2_homotopy_equivalence) | |
| 1429 | ||
| 1430 | ||
| 1431 | lemma Brouwer_degree2_retraction_map: | |
| 1432 | assumes "retraction_map (nsphere p) (nsphere p) f" | |
| 1433 | shows "\<bar>Brouwer_degree2 p f\<bar> = 1" | |
| 1434 | proof - | |
| 1435 | obtain g where g: "retraction_maps (nsphere p) (nsphere p) f g" | |
| 1436 | using assms by (auto simp: retraction_map_def) | |
| 1437 | show ?thesis | |
| 1438 | proof (rule Brouwer_degree2_homotopy_equivalence) | |
| 1439 | show "homotopic_with (\<lambda>x. True) (nsphere p) (nsphere p) (f \<circ> g) id" | |
| 1440 | using g apply (auto simp: retraction_maps_def) | |
| 1441 | by (simp add: homotopic_with_equal continuous_map_compose) | |
| 1442 | show "continuous_map (nsphere p) (nsphere p) f" "continuous_map (nsphere p) (nsphere p) g" | |
| 1443 | using g retraction_maps_def by blast+ | |
| 1444 | qed | |
| 1445 | qed | |
| 1446 | ||
| 1447 | lemma Brouwer_degree2_section_map: | |
| 1448 | assumes "section_map (nsphere p) (nsphere p) f" | |
| 1449 | shows "\<bar>Brouwer_degree2 p f\<bar> = 1" | |
| 1450 | proof - | |
| 1451 | obtain g where g: "retraction_maps (nsphere p) (nsphere p) g f" | |
| 1452 | using assms by (auto simp: section_map_def) | |
| 1453 | show ?thesis | |
| 1454 | proof (rule Brouwer_degree2_homotopy_equivalence) | |
| 1455 | show "homotopic_with (\<lambda>x. True) (nsphere p) (nsphere p) (g \<circ> f) id" | |
| 1456 | using g apply (auto simp: retraction_maps_def) | |
| 1457 | by (simp add: homotopic_with_equal continuous_map_compose) | |
| 1458 | show "continuous_map (nsphere p) (nsphere p) g" "continuous_map (nsphere p) (nsphere p) f" | |
| 1459 | using g retraction_maps_def by blast+ | |
| 1460 | qed | |
| 1461 | qed | |
| 1462 | ||
| 1463 | lemma Brouwer_degree2_homeomorphic_map: | |
| 1464 | "homeomorphic_map (nsphere p) (nsphere p) f \<Longrightarrow> \<bar>Brouwer_degree2 p f\<bar> = 1" | |
| 1465 | using Brouwer_degree2_retraction_map section_and_retraction_eq_homeomorphic_map by blast | |
| 1466 | ||
| 1467 | ||
| 1468 | lemma Brouwer_degree2_nullhomotopic: | |
| 1469 | assumes "homotopic_with (\<lambda>x. True) (nsphere p) (nsphere p) f (\<lambda>x. a)" | |
| 1470 | shows "Brouwer_degree2 p f = 0" | |
| 1471 | proof - | |
| 1472 | have contf: "continuous_map (nsphere p) (nsphere p) f" | |
| 1473 | and contc: "continuous_map (nsphere p) (nsphere p) (\<lambda>x. a)" | |
| 1474 | using homotopic_with_imp_continuous_maps [OF assms] by metis+ | |
| 1475 | have "Brouwer_degree2 p f = Brouwer_degree2 p (\<lambda>x. a)" | |
| 1476 | using Brouwer_degree2_homotopic [OF assms] . | |
| 1477 | moreover | |
| 1478 | let ?G = "reduced_homology_group (int p) (nsphere p)" | |
| 1479 | interpret group ?G | |
| 1480 | by simp | |
| 1481 | have "Brouwer_degree2 p (\<lambda>x. a) = 0" | |
| 1482 | proof (rule Brouwer_degree2_unique [OF contc]) | |
| 1483 | fix c | |
| 1484 | assume c: "c \<in> carrier ?G" | |
| 1485 |     have "continuous_map (nsphere p) (subtopology (nsphere p) {a}) (\<lambda>f. a)"
 | |
| 1486 | using contc continuous_map_in_subtopology by blast | |
| 1487 |     then have he: "hom_induced p (nsphere p) {} (nsphere p) {} (\<lambda>x. a)
 | |
| 1488 |                  = hom_induced p (subtopology (nsphere p) {a}) {} (nsphere p) {} id \<circ>
 | |
| 1489 |                    hom_induced p (nsphere p) {} (subtopology (nsphere p) {a}) {} (\<lambda>x. a)"
 | |
| 1490 | by (metis continuous_map_id_subt hom_induced_compose id_comp image_empty order_refl) | |
| 1491 |     have 1: "hom_induced p (nsphere p) {} (subtopology (nsphere p) {a}) {} (\<lambda>x. a) c =
 | |
| 1492 |              \<one>\<^bsub>reduced_homology_group (int p) (subtopology (nsphere p) {a})\<^esub>"
 | |
| 1493 |       using c trivial_reduced_homology_group_contractible_space [of "subtopology (nsphere p) {a}" p]
 | |
| 1494 | by (simp add: hom_induced_reduced contractible_space_subtopology_singleton trivial_group_subset group.trivial_group_subset subset_iff) | |
| 1495 |     show "hom_induced (int p) (nsphere p) {} (nsphere p) {} (\<lambda>x. a) c =
 | |
| 1496 | c [^]\<^bsub>?G\<^esub> (0::int)" | |
| 1497 | apply (simp add: he 1) | |
| 1498 | using hom_induced_reduced_hom group_hom.hom_one group_hom_axioms_def group_hom_def group_reduced_homology_group by blast | |
| 1499 | qed | |
| 1500 | ultimately show ?thesis | |
| 1501 | by metis | |
| 1502 | qed | |
| 1503 | ||
| 1504 | ||
| 1505 | lemma Brouwer_degree2_const: "Brouwer_degree2 p (\<lambda>x. a) = 0" | |
| 1506 | proof (cases "continuous_map(nsphere p) (nsphere p) (\<lambda>x. a)") | |
| 1507 | case True | |
| 1508 | then show ?thesis | |
| 1509 | by (auto intro: Brouwer_degree2_nullhomotopic [where a=a]) | |
| 1510 | next | |
| 1511 | case False | |
| 1512 | let ?G = "reduced_homology_group (int p) (nsphere p)" | |
| 1513 | let ?H = "homology_group (int p) (nsphere p)" | |
| 1514 | interpret group ?G | |
| 1515 | by simp | |
| 1516 | have eq1: "\<one>\<^bsub>?H\<^esub> = \<one>\<^bsub>?G\<^esub>" | |
| 1517 | by (simp add: one_reduced_homology_group) | |
| 1518 |   have *: "\<forall>x\<in>carrier ?G. hom_induced (int p) (nsphere p) {} (nsphere p) {} (\<lambda>x. a) x = \<one>\<^bsub>?H\<^esub>"
 | |
| 1519 | by (metis False hom_induced_default one_relative_homology_group) | |
| 1520 |   obtain c where c: "c \<in> carrier ?G" and ceq: "subgroup_generated ?G {c} = ?G"
 | |
| 1521 | using cyclic_reduced_homology_group_nsphere [of p p] by (force simp: cyclic_group_def) | |
| 1522 | have [simp]: "ord c = 0" | |
| 1523 | by (simp add: c ceq iso_finite [OF reduced_homology_group_nsphere(1)] flip: infinite_cyclic_subgroup_order) | |
| 1524 | show ?thesis | |
| 1525 | unfolding Brouwer_degree2_def | |
| 1526 | proof (rule some_equality) | |
| 1527 | fix d :: "int" | |
| 1528 |     assume "\<forall>x\<in>carrier ?G. hom_induced (int p) (nsphere p) {} (nsphere p) {} (\<lambda>x. a) x = x [^]\<^bsub>?G\<^esub> d"
 | |
| 1529 | then have "c [^]\<^bsub>?G\<^esub> d = \<one>\<^bsub>?H\<^esub>" | |
| 1530 | using "*" c by blast | |
| 1531 | then have "int (ord c) dvd d" | |
| 1532 | using c eq1 int_pow_eq_id by auto | |
| 1533 | then show "d = 0" | |
| 1534 | by (simp add: * del: one_relative_homology_group) | |
| 1535 | qed (use "*" eq1 in force) | |
| 1536 | qed | |
| 1537 | ||
| 1538 | ||
| 1539 | corollary Brouwer_degree2_nonsurjective: | |
| 1540 | "\<lbrakk>continuous_map(nsphere p) (nsphere p) f; f ` topspace (nsphere p) \<noteq> topspace (nsphere p)\<rbrakk> | |
| 1541 | \<Longrightarrow> Brouwer_degree2 p f = 0" | |
| 1542 | by (meson Brouwer_degree2_nullhomotopic nullhomotopic_nonsurjective_sphere_map) | |
| 1543 | ||
| 1544 | ||
| 1545 | proposition Brouwer_degree2_reflection: | |
| 1546 | "Brouwer_degree2 p (\<lambda>x i. if i = 0 then -x i else x i) = -1" (is "Brouwer_degree2 _ ?r = -1") | |
| 1547 | proof (induction p) | |
| 1548 | case 0 | |
| 1549 | let ?G = "homology_group 0 (nsphere 0)" | |
| 1550 |   let ?D = "homology_group 0 (discrete_topology {()})"
 | |
| 1551 | interpret group ?G | |
| 1552 | by simp | |
| 1553 | define r where "r \<equiv> \<lambda>x::nat\<Rightarrow>real. \<lambda>i. if i = 0 then -x i else x i" | |
| 1554 | then have [simp]: "r \<circ> r = id" | |
| 1555 | by force | |
| 1556 | have cmr: "continuous_map (nsphere 0) (nsphere 0) r" | |
| 1557 | by (simp add: r_def continuous_map_nsphere_reflection) | |
| 1558 |   have *: "hom_induced 0 (nsphere 0) {} (nsphere 0) {} r c = inv\<^bsub>?G\<^esub> c"
 | |
| 1559 | if "c \<in> carrier(reduced_homology_group 0 (nsphere 0))" for c | |
| 1560 | proof - | |
| 1561 | have c: "c \<in> carrier ?G" | |
| 1562 |       and ceq: "hom_induced 0 (nsphere 0) {} (discrete_topology {()}) {} (\<lambda>x. ()) c = \<one>\<^bsub>?D\<^esub>"
 | |
| 1563 | using that by (auto simp: carrier_reduced_homology_group kernel_def) | |
| 1564 | define pp::"nat\<Rightarrow>real" where "pp \<equiv> \<lambda>i. if i = 0 then 1 else 0" | |
| 1565 | define nn::"nat\<Rightarrow>real" where "nn \<equiv> \<lambda>i. if i = 0 then -1 else 0" | |
| 1566 |     have topn0: "topspace(nsphere 0) = {pp,nn}"
 | |
| 1567 | by (auto simp: nsphere pp_def nn_def fun_eq_iff power2_eq_1_iff split: if_split_asm) | |
| 1568 | have "t1_space (nsphere 0)" | |
| 1569 | unfolding nsphere | |
| 1570 | apply (rule t1_space_subtopology) | |
| 1571 | by (metis (full_types) open_fun_def t1_space t1_space_def) | |
| 1572 |     then have dtn0: "discrete_topology {pp,nn} = nsphere 0"
 | |
| 1573 | using finite_t1_space_imp_discrete_topology [OF topn0] by auto | |
| 1574 | have "pp \<noteq> nn" | |
| 1575 | by (auto simp: pp_def nn_def fun_eq_iff) | |
| 1576 | have [simp]: "r pp = nn" "r nn = pp" | |
| 1577 | by (auto simp: r_def pp_def nn_def fun_eq_iff) | |
| 1578 |     have iso: "(\<lambda>(a,b). hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} id a
 | |
| 1579 |                   \<otimes>\<^bsub>?G\<^esub> hom_induced 0 (subtopology (nsphere 0) {nn}) {} (nsphere 0) {} id b)
 | |
| 1580 |             \<in> iso (homology_group 0 (subtopology (nsphere 0) {pp}) \<times>\<times> homology_group 0 (subtopology (nsphere 0) {nn}))
 | |
| 1581 | ?G" (is "?f \<in> iso (?P \<times>\<times> ?N) ?G") | |
| 1582 | apply (rule homology_additivity_explicit) | |
| 1583 | using dtn0 \<open>pp \<noteq> nn\<close> by (auto simp: discrete_topology_unique) | |
| 1584 | then have fim: "?f ` carrier(?P \<times>\<times> ?N) = carrier ?G" | |
| 1585 | by (simp add: iso_def bij_betw_def) | |
| 1586 | obtain d d' where d: "d \<in> carrier ?P" and d': "d' \<in> carrier ?N" and eqc: "?f(d,d') = c" | |
| 1587 | using c by (force simp flip: fim) | |
| 1588 |     let ?h = "\<lambda>xx. hom_induced 0 (subtopology (nsphere 0) {xx}) {} (discrete_topology {()}) {} (\<lambda>x. ())"
 | |
| 1589 |     have "retraction_map (subtopology (nsphere 0) {pp}) (subtopology (nsphere 0) {nn}) r"
 | |
| 1590 | apply (simp add: retraction_map_def retraction_maps_def continuous_map_in_subtopology continuous_map_from_subtopology cmr image_subset_iff) | |
| 1591 | apply (rule_tac x=r in exI) | |
| 1592 | apply (force simp: retraction_map_def retraction_maps_def continuous_map_in_subtopology continuous_map_from_subtopology cmr) | |
| 1593 | done | |
| 1594 |     then have "carrier ?N = (hom_induced 0 (subtopology (nsphere 0) {pp}) {} (subtopology (nsphere 0) {nn}) {} r) ` carrier ?P"
 | |
| 1595 | by (rule surj_hom_induced_retraction_map) | |
| 1596 |     then obtain e where e: "e \<in> carrier ?P" and eqd': "hom_induced 0 (subtopology (nsphere 0) {pp}) {} (subtopology (nsphere 0) {nn}) {} r e = d'"
 | |
| 1597 | using d' by auto | |
| 1598 |     have "section_map (subtopology (nsphere 0) {pp}) (discrete_topology {()}) (\<lambda>x. ())"
 | |
| 1599 | by (force simp: section_map_def retraction_maps_def topn0) | |
| 1600 | then have "?h pp \<in> mon ?P ?D" | |
| 1601 | by (rule mon_hom_induced_section_map) | |
| 1602 | then have one: "x = one ?P" | |
| 1603 | if "?h pp x = \<one>\<^bsub>?D\<^esub>" "x \<in> carrier ?P" for x | |
| 1604 | using that by (simp add: mon_iff_hom_one) | |
| 1605 | interpret hpd: group_hom ?P ?D "?h pp" | |
| 1606 | using hom_induced_empty_hom by (simp add: hom_induced_empty_hom group_hom_axioms_def group_hom_def) | |
| 1607 |     interpret hgd: group_hom ?G ?D "hom_induced 0 (nsphere 0) {} (discrete_topology {()}) {} (\<lambda>x. ())"
 | |
| 1608 | using hom_induced_empty_hom by (simp add: hom_induced_empty_hom group_hom_axioms_def group_hom_def) | |
| 1609 |     interpret hpg: group_hom ?P ?G "hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} r"
 | |
| 1610 | using hom_induced_empty_hom by (simp add: hom_induced_empty_hom group_hom_axioms_def group_hom_def) | |
| 1611 |     interpret hgg: group_hom ?G ?G "hom_induced 0 (nsphere 0) {} (nsphere 0) {} r"
 | |
| 1612 | using hom_induced_empty_hom by (simp add: hom_induced_empty_hom group_hom_axioms_def group_hom_def) | |
| 1613 | have "?h pp d = | |
| 1614 |           (hom_induced 0 (nsphere 0) {} (discrete_topology {()}) {} (\<lambda>x. ())
 | |
| 1615 |            \<circ> hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} id) d"
 | |
| 1616 | by (simp flip: hom_induced_compose_empty) | |
| 1617 | moreover | |
| 1618 |     have "?h pp = ?h nn \<circ> hom_induced 0 (subtopology (nsphere 0) {pp}) {} (subtopology (nsphere 0) {nn}) {} r"
 | |
| 1619 | by (simp add: cmr continuous_map_from_subtopology continuous_map_in_subtopology image_subset_iff flip: hom_induced_compose_empty) | |
| 1620 | then have "?h pp e = | |
| 1621 |                (hom_induced 0 (nsphere 0) {} (discrete_topology {()}) {} (\<lambda>x. ())
 | |
| 1622 |                 \<circ> hom_induced 0 (subtopology (nsphere 0) {nn}) {} (nsphere 0) {} id) d'"
 | |
| 1623 | by (simp flip: hom_induced_compose_empty eqd') | |
| 1624 |     ultimately have "?h pp (d \<otimes>\<^bsub>?P\<^esub> e) = hom_induced 0 (nsphere 0) {} (discrete_topology {()}) {} (\<lambda>x. ()) (?f(d,d'))"
 | |
| 1625 | by (simp add: d e hom_induced_carrier) | |
| 1626 | then have "?h pp (d \<otimes>\<^bsub>?P\<^esub> e) = \<one>\<^bsub>?D\<^esub>" | |
| 1627 | using ceq eqc by simp | |
| 1628 | then have inv_p: "inv\<^bsub>?P\<^esub> d = e" | |
| 1629 | by (metis (no_types, lifting) Group.group_def d e group.inv_equality group.r_inv group_relative_homology_group one monoid.m_closed) | |
| 1630 |     have cmr_pn: "continuous_map (subtopology (nsphere 0) {pp}) (subtopology (nsphere 0) {nn}) r"
 | |
| 1631 | by (simp add: cmr continuous_map_from_subtopology continuous_map_in_subtopology image_subset_iff) | |
| 1632 |     then have "hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} (id \<circ> r) =
 | |
| 1633 |                hom_induced 0 (subtopology (nsphere 0) {nn}) {} (nsphere 0) {} id \<circ>
 | |
| 1634 |                hom_induced 0 (subtopology (nsphere 0) {pp}) {} (subtopology (nsphere 0) {nn}) {} r"
 | |
| 1635 | using hom_induced_compose_empty continuous_map_id_subt by blast | |
| 1636 |     then have "inv\<^bsub>?G\<^esub> hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} r d =
 | |
| 1637 |                   hom_induced 0 (subtopology (nsphere 0) {nn}) {} (nsphere 0) {} id d'"
 | |
| 1638 | apply (simp add: flip: inv_p eqd') | |
| 1639 | using d hpg.hom_inv by auto | |
| 1640 |     then have c: "c = (hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} id d)
 | |
| 1641 |                        \<otimes>\<^bsub>?G\<^esub> inv\<^bsub>?G\<^esub> (hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} r d)"
 | |
| 1642 | by (simp flip: eqc) | |
| 1643 |     have "hom_induced 0 (nsphere 0) {} (nsphere 0) {} r \<circ>
 | |
| 1644 |           hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} id =
 | |
| 1645 |           hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} r"
 | |
| 1646 | by (metis cmr comp_id continuous_map_id_subt hom_induced_compose_empty) | |
| 1647 | moreover | |
| 1648 |     have "hom_induced 0 (nsphere 0) {} (nsphere 0) {} r \<circ>
 | |
| 1649 |           hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} r =
 | |
| 1650 |           hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} id"
 | |
| 1651 | by (metis \<open>r \<circ> r = id\<close> cmr continuous_map_from_subtopology hom_induced_compose_empty) | |
| 1652 | ultimately show ?thesis | |
| 1653 | by (metis inv_p c comp_def d e hgg.hom_inv hgg.hom_mult hom_induced_carrier hpd.G.inv_inv hpg.hom_inv inv_mult_group) | |
| 1654 | qed | |
| 1655 | show ?case | |
| 1656 | unfolding r_def [symmetric] | |
| 1657 | using Brouwer_degree2_unique [OF cmr] | |
| 1658 | by (auto simp: * group.int_pow_neg group.int_pow_1 reduced_homology_group_def intro!: Brouwer_degree2_unique [OF cmr]) | |
| 1659 | next | |
| 1660 | case (Suc p) | |
| 1661 | let ?G = "reduced_homology_group (int p) (nsphere p)" | |
| 1662 | let ?G1 = "reduced_homology_group (1 + int p) (nsphere (Suc p))" | |
| 1663 | obtain f g where fg: "group_isomorphisms ?G ?G1 f g" | |
| 1664 | and *: "\<forall>c\<in>carrier ?G. | |
| 1665 |            hom_induced (1 + int p) (nsphere (Suc p)) {} (nsphere (Suc p)) {} ?r (f c) =
 | |
| 1666 |            f (hom_induced p (nsphere p) {} (nsphere p) {} ?r c)"
 | |
| 1667 | using reduced_homology_group_nsphere_step | |
| 1668 | by (meson group.iso_iff_group_isomorphisms group_reduced_homology_group) | |
| 1669 | then have eq: "carrier ?G1 = f ` carrier ?G" | |
| 1670 | by (fastforce simp add: iso_iff dest: group_isomorphisms_imp_iso) | |
| 1671 | interpret group_hom ?G ?G1 f | |
| 1672 | by (meson fg group_hom_axioms_def group_hom_def group_isomorphisms_def group_reduced_homology_group) | |
| 1673 | have homf: "f \<in> hom ?G ?G1" | |
| 1674 | using fg group_isomorphisms_def by blast | |
| 1675 |   have "hom_induced (1 + int p) (nsphere (Suc p)) {} (nsphere (Suc p)) {} ?r (f y) = f y [^]\<^bsub>?G1\<^esub> (-1::int)"
 | |
| 1676 | if "y \<in> carrier ?G" for y | |
| 1677 | by (simp add: that * Brouwer_degree2 Suc hom_int_pow) | |
| 1678 | then show ?case | |
| 1679 | by (fastforce simp: eq intro: Brouwer_degree2_unique [OF continuous_map_nsphere_reflection]) | |
| 1680 | qed | |
| 1681 | ||
| 1682 | end |