author | nipkow |
Thu, 17 Jul 2025 21:06:22 +0100 | |
changeset 82885 | 5d2a599f88af |
parent 82323 | b022c013b04b |
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:
70097
diff
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 - |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
99 |
have *: "continuous_map X (discrete_topology {()}) (\<lambda>x. ())" "(\<lambda>x. ()) \<in> S \<rightarrow> {()}" |
70095 | 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: |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
438 |
"\<lbrakk>retraction_maps X Y r s; r \<in> U \<rightarrow> V; s \<in> V \<rightarrow> U; homotopic_with (\<lambda>h. h ` U \<subseteq> U) X X (s \<circ> r) id\<rbrakk> |
70095 | 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) |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
443 |
apply (auto simp: image_subset_iff_funcset Pi_iff continuous_map_compose homotopic_with_equal) |
70095 | 444 |
done |
445 |
||
446 |
||
447 |
lemma deformation_retract_relative_homology_group_isomorphisms: |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
448 |
"\<lbrakk>retraction_maps X Y r id; V \<subseteq> U; r \<in> U \<rightarrow> V; homotopic_with (\<lambda>h. h ` U \<subseteq> U) X X r id\<rbrakk> |
70095 | 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)" |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
451 |
by (simp add: deformation_retraction_relative_homology_group_isomorphisms |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
452 |
in_mono) |
70095 | 453 |
|
454 |
lemma deformation_retract_relative_homology_group_isomorphism: |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
455 |
"\<lbrakk>retraction_maps X Y r id; V \<subseteq> U; r \<in> U \<rightarrow> V; homotopic_with (\<lambda>h. h ` U \<subseteq> U) X X r id\<rbrakk> |
70095 | 456 |
\<Longrightarrow> (hom_induced p X U Y V r) \<in> iso (relative_homology_group p X U) (relative_homology_group p Y V)" |
457 |
by (metis deformation_retract_relative_homology_group_isomorphisms group_isomorphisms_imp_iso) |
|
458 |
||
459 |
lemma deformation_retract_relative_homology_group_isomorphism_id: |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
460 |
"\<lbrakk>retraction_maps X Y r id; V \<subseteq> U; r \<in> U \<rightarrow> V; homotopic_with (\<lambda>h. h ` U \<subseteq> U) X X r id\<rbrakk> |
70095 | 461 |
\<Longrightarrow> (hom_induced p Y V X U id) \<in> iso (relative_homology_group p Y V) (relative_homology_group p X U)" |
462 |
by (metis deformation_retract_relative_homology_group_isomorphisms group_isomorphisms_imp_iso group_isomorphisms_sym) |
|
463 |
||
464 |
lemma deformation_retraction_imp_isomorphic_relative_homology_groups: |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
465 |
"\<lbrakk>retraction_maps X Y r s; r \<in> U \<rightarrow> V; s ` V \<subseteq> U; homotopic_with (\<lambda>h. h ` U \<subseteq> U) X X (s \<circ> r) id\<rbrakk> |
70095 | 466 |
\<Longrightarrow> relative_homology_group p X U \<cong> relative_homology_group p Y V" |
467 |
by (blast intro: is_isoI group_isomorphisms_imp_iso deformation_retraction_relative_homology_group_isomorphisms) |
|
468 |
||
469 |
lemma deformation_retraction_imp_isomorphic_homology_groups: |
|
470 |
"\<lbrakk>retraction_maps X Y r s; homotopic_with (\<lambda>h. True) X X (s \<circ> r) id\<rbrakk> |
|
471 |
\<Longrightarrow> homology_group p X \<cong> homology_group p Y" |
|
472 |
by (simp add: deformation_retraction_imp_homotopy_equivalent_space homotopy_equivalent_space_imp_isomorphic_homology_groups) |
|
473 |
||
474 |
lemma deformation_retract_imp_isomorphic_relative_homology_groups: |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
475 |
"\<lbrakk>retraction_maps X X' r id; V \<subseteq> U; r \<in> U \<rightarrow> V; homotopic_with (\<lambda>h. h ` U \<subseteq> U) X X r id\<rbrakk> |
70095 | 476 |
\<Longrightarrow> relative_homology_group p X U \<cong> relative_homology_group p X' V" |
477 |
by (simp add: deformation_retraction_imp_isomorphic_relative_homology_groups) |
|
478 |
||
479 |
lemma deformation_retract_imp_isomorphic_homology_groups: |
|
480 |
"\<lbrakk>retraction_maps X X' r id; homotopic_with (\<lambda>h. True) X X r id\<rbrakk> |
|
481 |
\<Longrightarrow> homology_group p X \<cong> homology_group p X'" |
|
482 |
by (simp add: deformation_retraction_imp_isomorphic_homology_groups) |
|
483 |
||
484 |
||
485 |
lemma epi_hom_induced_inclusion: |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
486 |
assumes "homotopic_with (\<lambda>x. True) X X id f" and "f \<in> topspace X \<rightarrow> S" |
70095 | 487 |
shows "(hom_induced p (subtopology X S) {} X {} id) |
488 |
\<in> epi (homology_group p (subtopology X S)) (homology_group p X)" |
|
489 |
proof (rule epi_right_invertible) |
|
490 |
show "hom_induced p (subtopology X S) {} X {} id |
|
491 |
\<in> hom (homology_group p (subtopology X S)) (homology_group p X)" |
|
492 |
by (simp add: hom_induced_empty_hom) |
|
493 |
show "hom_induced p X {} (subtopology X S) {} f |
|
494 |
\<in> carrier (homology_group p X) \<rightarrow> carrier (homology_group p (subtopology X S))" |
|
495 |
by (simp add: hom_induced_carrier) |
|
496 |
fix x |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
497 |
assume x: "x \<in> carrier (homology_group p X)" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
498 |
show "hom_induced p (subtopology X S) {} X {} id (hom_induced p X {} (subtopology X S) {} f x) = x" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
499 |
proof (subst hom_induced_compose') |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
500 |
show "continuous_map X (subtopology X S) f" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
501 |
by (meson assms continuous_map_into_subtopology |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
502 |
homotopic_with_imp_continuous_maps) |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
503 |
show "hom_induced p X {} X {} (id \<circ> f) x = x" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
504 |
by (metis assms(1) hom_induced_id homology_homotopy_empty id_comp x) |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
505 |
qed (use assms in auto) |
70095 | 506 |
qed |
507 |
||
508 |
||
509 |
lemma trivial_homomorphism_hom_induced_relativization: |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
510 |
assumes "homotopic_with (\<lambda>x. True) X X id f" and "f \<in> topspace X \<rightarrow> S" |
70095 | 511 |
shows "trivial_homomorphism (homology_group p X) (relative_homology_group p X S) |
512 |
(hom_induced p X {} X S id)" |
|
513 |
proof - |
|
514 |
have "(hom_induced p (subtopology X S) {} X {} id) |
|
515 |
\<in> epi (homology_group p (subtopology X S)) (homology_group p X)" |
|
516 |
by (metis assms epi_hom_induced_inclusion) |
|
517 |
then show ?thesis |
|
518 |
using homology_exactness_axiom_3 [of p X S] homology_exactness_axiom_1 [of p X S] |
|
519 |
by (simp add: epi_def group.trivial_homomorphism_image group_hom.trivial_hom_iff) |
|
520 |
qed |
|
521 |
||
522 |
||
523 |
lemma mon_hom_boundary_inclusion: |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
524 |
assumes "homotopic_with (\<lambda>x. True) X X id f" and "f \<in> topspace X \<rightarrow> S" |
70095 | 525 |
shows "(hom_boundary p X S) \<in> mon |
526 |
(relative_homology_group p X S) (homology_group (p - 1) (subtopology X S))" |
|
527 |
proof - |
|
528 |
have "(hom_induced p (subtopology X S) {} X {} id) |
|
529 |
\<in> epi (homology_group p (subtopology X S)) (homology_group p X)" |
|
530 |
by (metis assms epi_hom_induced_inclusion) |
|
531 |
then show ?thesis |
|
532 |
using homology_exactness_axiom_3 [of p X S] homology_exactness_axiom_1 [of p X S] |
|
533 |
apply (simp add: mon_def epi_def hom_boundary_hom) |
|
73932
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
72632
diff
changeset
|
534 |
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 | 535 |
qed |
536 |
||
537 |
lemma short_exact_sequence_hom_induced_relativization: |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
538 |
assumes "homotopic_with (\<lambda>x. True) X X id f" and "f \<in> topspace X \<rightarrow> S" |
70095 | 539 |
shows "short_exact_sequence (homology_group (p-1) X) (homology_group (p-1) (subtopology X S)) (relative_homology_group p X S) |
540 |
(hom_induced (p-1) (subtopology X S) {} X {} id) (hom_boundary p X S)" |
|
541 |
unfolding short_exact_sequence_iff |
|
542 |
by (intro conjI homology_exactness_axiom_2 epi_hom_induced_inclusion [OF assms] mon_hom_boundary_inclusion [OF assms]) |
|
543 |
||
544 |
||
545 |
lemma group_isomorphisms_homology_group_prod_deformation: |
|
546 |
fixes p::int |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
547 |
assumes "homotopic_with (\<lambda>x. True) X X id f" and "f \<in> topspace X \<rightarrow> S" |
70095 | 548 |
obtains H K where |
549 |
"subgroup H (homology_group p (subtopology X S))" |
|
550 |
"subgroup K (homology_group p (subtopology X S))" |
|
551 |
"(\<lambda>(x, y). x \<otimes>\<^bsub>homology_group p (subtopology X S)\<^esub> y) |
|
552 |
\<in> Group.iso (subgroup_generated (homology_group p (subtopology X S)) H \<times>\<times> |
|
553 |
subgroup_generated (homology_group p (subtopology X S)) K) |
|
554 |
(homology_group p (subtopology X S))" |
|
555 |
"hom_boundary (p + 1) X S |
|
556 |
\<in> Group.iso (relative_homology_group (p + 1) X S) |
|
557 |
(subgroup_generated (homology_group p (subtopology X S)) H)" |
|
558 |
"hom_induced p (subtopology X S) {} X {} id |
|
559 |
\<in> Group.iso |
|
560 |
(subgroup_generated (homology_group p (subtopology X S)) K) |
|
561 |
(homology_group p X)" |
|
562 |
proof - |
|
563 |
let ?rhs = "relative_homology_group (p + 1) X S" |
|
564 |
let ?pXS = "homology_group p (subtopology X S)" |
|
565 |
let ?pX = "homology_group p X" |
|
566 |
let ?hb = "hom_boundary (p + 1) X S" |
|
567 |
let ?hi = "hom_induced p (subtopology X S) {} X {} id" |
|
568 |
have x: "short_exact_sequence (?pX) ?pXS ?rhs ?hi ?hb" |
|
569 |
using short_exact_sequence_hom_induced_relativization [OF assms, of "p + 1"] by simp |
|
570 |
have contf: "continuous_map X (subtopology X S) f" |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
571 |
by (metis assms continuous_map_into_subtopology homotopic_with_imp_continuous_maps) |
70095 | 572 |
obtain H K where HK: "H \<lhd> ?pXS" "subgroup K ?pXS" "H \<inter> K \<subseteq> {one ?pXS}" "set_mult ?pXS H K = carrier ?pXS" |
573 |
and iso: "?hb \<in> iso ?rhs (subgroup_generated ?pXS H)" "?hi \<in> iso (subgroup_generated ?pXS K) ?pX" |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
574 |
proof (rule splitting_lemma_right [OF x, where g' = "hom_induced p X {} (subtopology X S) {} f"]) |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
575 |
show "hom_induced p X {} (subtopology X S) {} f \<in> hom (homology_group p X) (homology_group p (subtopology X S))" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
576 |
using hom_induced_empty_hom by blast |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
577 |
next |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
578 |
fix z |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
579 |
assume "z \<in> carrier (homology_group p X)" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
580 |
then show "hom_induced p (subtopology X S) {} X {} id (hom_induced p X {} (subtopology X S) {} f z) = z" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
581 |
using assms(1) contf hom_induced_id homology_homotopy_empty |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
582 |
by (fastforce simp add: hom_induced_compose') |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
583 |
qed blast |
70095 | 584 |
show ?thesis |
585 |
proof |
|
586 |
show "subgroup H ?pXS" |
|
587 |
using HK(1) normal_imp_subgroup by blast |
|
588 |
then show "(\<lambda>(x, y). x \<otimes>\<^bsub>?pXS\<^esub> y) |
|
589 |
\<in> Group.iso (subgroup_generated (?pXS) H \<times>\<times> subgroup_generated (?pXS) K) (?pXS)" |
|
590 |
by (meson HK abelian_relative_homology_group group_disjoint_sum.iso_group_mul group_disjoint_sum_def group_relative_homology_group) |
|
591 |
show "subgroup K ?pXS" |
|
592 |
by (rule HK) |
|
593 |
show "hom_boundary (p + 1) X S \<in> Group.iso ?rhs (subgroup_generated (?pXS) H)" |
|
594 |
using iso int_ops(4) by presburger |
|
595 |
show "hom_induced p (subtopology X S) {} X {} id \<in> Group.iso (subgroup_generated (?pXS) K) (?pX)" |
|
596 |
by (simp add: iso(2)) |
|
597 |
qed |
|
598 |
qed |
|
599 |
||
600 |
lemma iso_homology_group_prod_deformation: |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
601 |
assumes "homotopic_with (\<lambda>x. True) X X id f" and "f \<in> topspace X \<rightarrow> S" |
70095 | 602 |
shows "homology_group p (subtopology X S) |
603 |
\<cong> DirProd (homology_group p X) (relative_homology_group(p + 1) X S)" |
|
604 |
(is "?G \<cong> DirProd ?H ?R") |
|
605 |
proof - |
|
606 |
obtain H K where HK: |
|
607 |
"(\<lambda>(x, y). x \<otimes>\<^bsub>?G\<^esub> y) |
|
608 |
\<in> Group.iso (subgroup_generated (?G) H \<times>\<times> subgroup_generated (?G) K) (?G)" |
|
609 |
"hom_boundary (p + 1) X S \<in> Group.iso (?R) (subgroup_generated (?G) H)" |
|
610 |
"hom_induced p (subtopology X S) {} X {} id \<in> Group.iso (subgroup_generated (?G) K) (?H)" |
|
611 |
by (blast intro: group_isomorphisms_homology_group_prod_deformation [OF assms]) |
|
612 |
have "?G \<cong> DirProd (subgroup_generated (?G) H) (subgroup_generated (?G) K)" |
|
613 |
by (meson DirProd_group HK(1) group.group_subgroup_generated group.iso_sym group_relative_homology_group is_isoI) |
|
614 |
also have "\<dots> \<cong> DirProd ?R ?H" |
|
615 |
by (meson HK group.DirProd_iso_trans group.group_subgroup_generated group.iso_sym group_relative_homology_group is_isoI) |
|
616 |
also have "\<dots> \<cong> DirProd ?H ?R" |
|
617 |
by (simp add: DirProd_commute_iso) |
|
618 |
finally show ?thesis . |
|
619 |
qed |
|
620 |
||
621 |
||
622 |
||
623 |
lemma iso_homology_contractible_space_subtopology1: |
|
624 |
assumes "contractible_space X" "S \<subseteq> topspace X" "S \<noteq> {}" |
|
625 |
shows "homology_group 0 (subtopology X S) \<cong> DirProd integer_group (relative_homology_group(1) X S)" |
|
626 |
proof - |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
627 |
obtain f where "homotopic_with (\<lambda>x. True) X X id f" and "f \<in> topspace X \<rightarrow> S" |
70095 | 628 |
using assms contractible_space_alt by fastforce |
629 |
then have "homology_group 0 (subtopology X S) \<cong> homology_group 0 X \<times>\<times> relative_homology_group 1 X S" |
|
630 |
using iso_homology_group_prod_deformation [of X _ S 0] by auto |
|
631 |
also have "\<dots> \<cong> integer_group \<times>\<times> relative_homology_group 1 X S" |
|
632 |
using assms contractible_imp_path_connected_space group.DirProd_iso_trans group_relative_homology_group iso_refl isomorphic_integer_zeroth_homology_group by blast |
|
633 |
finally show ?thesis . |
|
634 |
qed |
|
635 |
||
636 |
lemma iso_homology_contractible_space_subtopology2: |
|
637 |
"\<lbrakk>contractible_space X; S \<subseteq> topspace X; p \<noteq> 0; S \<noteq> {}\<rbrakk> |
|
638 |
\<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:
72632
diff
changeset
|
639 |
by (metis (no_types, opaque_lifting) add.commute isomorphic_group_reduced_homology_of_contractible topspace_subtopology topspace_subtopology_subset un_reduced_homology_group) |
70095 | 640 |
|
641 |
lemma trivial_relative_homology_group_contractible_spaces: |
|
642 |
"\<lbrakk>contractible_space X; contractible_space(subtopology X S); topspace X \<inter> S \<noteq> {}\<rbrakk> |
|
643 |
\<Longrightarrow> trivial_group(relative_homology_group p X S)" |
|
644 |
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 |
|
645 |
||
646 |
lemma trivial_relative_homology_group_alt: |
|
647 |
assumes contf: "continuous_map X (subtopology X S) f" and hom: "homotopic_with (\<lambda>k. k ` S \<subseteq> S) X X f id" |
|
648 |
shows "trivial_group (relative_homology_group p X S)" |
|
649 |
proof (rule trivial_relative_homology_group_gen [OF contf]) |
|
650 |
show "homotopic_with (\<lambda>h. True) (subtopology X S) (subtopology X S) f id" |
|
651 |
using hom unfolding homotopic_with_def |
|
652 |
apply (rule ex_forward) |
|
653 |
apply (auto simp: prod_topology_subtopology continuous_map_in_subtopology continuous_map_from_subtopology image_subset_iff topspace_subtopology) |
|
654 |
done |
|
655 |
show "homotopic_with (\<lambda>k. True) X X f id" |
|
656 |
using assms by (force simp: homotopic_with_def) |
|
657 |
qed |
|
658 |
||
659 |
||
660 |
lemma iso_hom_induced_relativization_contractible: |
|
661 |
assumes "contractible_space(subtopology X S)" "contractible_space(subtopology X T)" "T \<subseteq> S" "topspace X \<inter> T \<noteq> {}" |
|
662 |
shows "(hom_induced p X T X S id) \<in> iso (relative_homology_group p X T) (relative_homology_group p X S)" |
|
663 |
proof (rule very_short_exact_sequence) |
|
664 |
show "exact_seq |
|
665 |
([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], |
|
666 |
[hom_relboundary p X S T, hom_induced p X T X S id, hom_induced p (subtopology X S) T X T id])" |
|
667 |
using homology_exactness_triple_1 [OF \<open>T \<subseteq> S\<close>] homology_exactness_triple_3 [OF \<open>T \<subseteq> S\<close>] |
|
668 |
by fastforce |
|
669 |
show "trivial_group (relative_homology_group p (subtopology X S) T)" "trivial_group (relative_homology_group(p - 1) (subtopology X S) T)" |
|
670 |
using assms |
|
671 |
by (force simp: inf.absorb_iff2 subtopology_subtopology topspace_subtopology intro!: trivial_relative_homology_group_contractible_spaces)+ |
|
672 |
qed |
|
673 |
||
674 |
corollary isomorphic_relative_homology_groups_relativization_contractible: |
|
675 |
assumes "contractible_space(subtopology X S)" "contractible_space(subtopology X T)" "T \<subseteq> S" "topspace X \<inter> T \<noteq> {}" |
|
676 |
shows "relative_homology_group p X T \<cong> relative_homology_group p X S" |
|
677 |
by (rule is_isoI) (rule iso_hom_induced_relativization_contractible [OF assms]) |
|
678 |
||
679 |
lemma iso_hom_induced_inclusion_contractible: |
|
680 |
assumes "contractible_space X" "contractible_space(subtopology X S)" "T \<subseteq> S" "topspace X \<inter> S \<noteq> {}" |
|
681 |
shows "(hom_induced p (subtopology X S) T X T id) |
|
682 |
\<in> iso (relative_homology_group p (subtopology X S) T) (relative_homology_group p X T)" |
|
683 |
proof (rule very_short_exact_sequence) |
|
684 |
show "exact_seq |
|
685 |
([relative_homology_group p X S, relative_homology_group p X T, |
|
686 |
relative_homology_group p (subtopology X S) T, relative_homology_group (p+1) X S], |
|
687 |
[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])" |
|
688 |
using homology_exactness_triple_2 [OF \<open>T \<subseteq> S\<close>] homology_exactness_triple_3 [OF \<open>T \<subseteq> S\<close>] |
|
689 |
by (metis add_diff_cancel_left' diff_add_cancel exact_seq_cons_iff) |
|
690 |
show "trivial_group (relative_homology_group (p+1) X S)" "trivial_group (relative_homology_group p X S)" |
|
691 |
using assms |
|
692 |
by (auto simp: subtopology_subtopology topspace_subtopology intro!: trivial_relative_homology_group_contractible_spaces) |
|
693 |
qed |
|
694 |
||
695 |
corollary isomorphic_relative_homology_groups_inclusion_contractible: |
|
696 |
assumes "contractible_space X" "contractible_space(subtopology X S)" "T \<subseteq> S" "topspace X \<inter> S \<noteq> {}" |
|
697 |
shows "relative_homology_group p (subtopology X S) T \<cong> relative_homology_group p X T" |
|
698 |
by (rule is_isoI) (rule iso_hom_induced_inclusion_contractible [OF assms]) |
|
699 |
||
700 |
lemma iso_hom_relboundary_contractible: |
|
701 |
assumes "contractible_space X" "contractible_space(subtopology X T)" "T \<subseteq> S" "topspace X \<inter> T \<noteq> {}" |
|
702 |
shows "hom_relboundary p X S T |
|
703 |
\<in> iso (relative_homology_group p X S) (relative_homology_group (p - 1) (subtopology X S) T)" |
|
704 |
proof (rule very_short_exact_sequence) |
|
705 |
show "exact_seq |
|
706 |
([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], |
|
707 |
[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])" |
|
708 |
using homology_exactness_triple_1 [OF \<open>T \<subseteq> S\<close>] homology_exactness_triple_2 [OF \<open>T \<subseteq> S\<close>] by simp |
|
709 |
show "trivial_group (relative_homology_group p X T)" "trivial_group (relative_homology_group (p - 1) X T)" |
|
710 |
using assms |
|
711 |
by (auto simp: subtopology_subtopology topspace_subtopology intro!: trivial_relative_homology_group_contractible_spaces) |
|
712 |
qed |
|
713 |
||
714 |
corollary isomorphic_relative_homology_groups_relboundary_contractible: |
|
715 |
assumes "contractible_space X" "contractible_space(subtopology X T)" "T \<subseteq> S" "topspace X \<inter> T \<noteq> {}" |
|
716 |
shows "relative_homology_group p X S \<cong> relative_homology_group (p - 1) (subtopology X S) T" |
|
717 |
by (rule is_isoI) (rule iso_hom_relboundary_contractible [OF assms]) |
|
718 |
||
719 |
lemma isomorphic_relative_contractible_space_imp_homology_groups: |
|
720 |
assumes "contractible_space X" "contractible_space Y" "S \<subseteq> topspace X" "T \<subseteq> topspace Y" |
|
721 |
and ST: "S = {} \<longleftrightarrow> T = {}" |
|
722 |
and iso: "\<And>p. relative_homology_group p X S \<cong> relative_homology_group p Y T" |
|
723 |
shows "homology_group p (subtopology X S) \<cong> homology_group p (subtopology Y T)" |
|
724 |
proof (cases "T = {}") |
|
725 |
case True |
|
726 |
have "homology_group p (subtopology X {}) \<cong> homology_group p (subtopology Y {})" |
|
727 |
by (simp add: homeomorphic_empty_space_eq homeomorphic_space_imp_isomorphic_homology_groups) |
|
728 |
then show ?thesis |
|
729 |
using ST True by blast |
|
730 |
next |
|
731 |
case False |
|
732 |
show ?thesis |
|
733 |
proof (cases "p = 0") |
|
734 |
case True |
|
735 |
have "homology_group p (subtopology X S) \<cong> integer_group \<times>\<times> relative_homology_group 1 X S" |
|
736 |
using assms True \<open>T \<noteq> {}\<close> |
|
737 |
by (simp add: iso_homology_contractible_space_subtopology1) |
|
738 |
also have "\<dots> \<cong> integer_group \<times>\<times> relative_homology_group 1 Y T" |
|
739 |
by (simp add: assms group.DirProd_iso_trans iso_refl) |
|
740 |
also have "\<dots> \<cong> homology_group p (subtopology Y T)" |
|
741 |
by (simp add: True \<open>T \<noteq> {}\<close> assms group.iso_sym iso_homology_contractible_space_subtopology1) |
|
742 |
finally show ?thesis . |
|
743 |
next |
|
744 |
case False |
|
745 |
have "homology_group p (subtopology X S) \<cong> relative_homology_group (p+1) X S" |
|
746 |
using assms False \<open>T \<noteq> {}\<close> |
|
747 |
by (simp add: iso_homology_contractible_space_subtopology2) |
|
748 |
also have "\<dots> \<cong> relative_homology_group (p+1) Y T" |
|
749 |
by (simp add: assms) |
|
750 |
also have "\<dots> \<cong> homology_group p (subtopology Y T)" |
|
751 |
by (simp add: False \<open>T \<noteq> {}\<close> assms group.iso_sym iso_homology_contractible_space_subtopology2) |
|
752 |
finally show ?thesis . |
|
753 |
qed |
|
754 |
qed |
|
755 |
||
756 |
||
757 |
subsection\<open>Homology groups of spheres\<close> |
|
758 |
||
759 |
lemma iso_reduced_homology_group_lower_hemisphere: |
|
760 |
assumes "k \<le> n" |
|
761 |
shows "hom_induced p (nsphere n) {} (nsphere n) {x. x k \<le> 0} id |
|
762 |
\<in> iso (reduced_homology_group p (nsphere n)) (relative_homology_group p (nsphere n) {x. x k \<le> 0})" |
|
763 |
proof (rule iso_reduced_homology_by_contractible) |
|
764 |
show "contractible_space (subtopology (nsphere n) {x. x k \<le> 0})" |
|
765 |
by (simp add: assms contractible_space_lower_hemisphere) |
|
766 |
have "(\<lambda>i. if i = k then -1 else 0) \<in> topspace (nsphere n) \<inter> {x. x k \<le> 0}" |
|
767 |
using assms by (simp add: nsphere if_distrib [of "\<lambda>x. x ^ 2"] cong: if_cong) |
|
768 |
then show "topspace (nsphere n) \<inter> {x. x k \<le> 0} \<noteq> {}" |
|
769 |
by blast |
|
770 |
qed |
|
771 |
||
772 |
||
773 |
lemma topspace_nsphere_1: |
|
774 |
assumes "x \<in> topspace (nsphere n)" shows "(x k)\<^sup>2 \<le> 1" |
|
775 |
proof (cases "k \<le> n") |
|
776 |
case True |
|
777 |
have "(\<Sum>i \<in> {..n} - {k}. (x i)\<^sup>2) = (\<Sum>i\<le>n. (x i)\<^sup>2) - (x k)\<^sup>2" |
|
778 |
using \<open>k \<le> n\<close> by (simp add: sum_diff) |
|
779 |
then show ?thesis |
|
780 |
using assms |
|
781 |
apply (simp add: nsphere) |
|
782 |
by (metis diff_ge_0_iff_ge sum_nonneg zero_le_power2) |
|
783 |
next |
|
784 |
case False |
|
785 |
then show ?thesis |
|
786 |
using assms by (simp add: nsphere) |
|
787 |
qed |
|
788 |
||
789 |
lemma topspace_nsphere_1_eq_0: |
|
790 |
fixes x :: "nat \<Rightarrow> real" |
|
791 |
assumes x: "x \<in> topspace (nsphere n)" and xk: "(x k)\<^sup>2 = 1" and "i \<noteq> k" |
|
792 |
shows "x i = 0" |
|
793 |
proof (cases "i \<le> n") |
|
794 |
case True |
|
795 |
have "k \<le> n" |
|
796 |
using x |
|
797 |
by (simp add: nsphere) (metis not_less xk zero_neq_one zero_power2) |
|
798 |
have "(\<Sum>i \<in> {..n} - {k}. (x i)\<^sup>2) = (\<Sum>i\<le>n. (x i)\<^sup>2) - (x k)\<^sup>2" |
|
799 |
using \<open>k \<le> n\<close> by (simp add: sum_diff) |
|
800 |
also have "\<dots> = 0" |
|
801 |
using assms by (simp add: nsphere) |
|
802 |
finally have "\<forall>i\<in>{..n} - {k}. (x i)\<^sup>2 = 0" |
|
803 |
by (simp add: sum_nonneg_eq_0_iff) |
|
804 |
then show ?thesis |
|
805 |
using True \<open>i \<noteq> k\<close> by auto |
|
806 |
next |
|
807 |
case False |
|
808 |
with x show ?thesis |
|
809 |
by (simp add: nsphere) |
|
810 |
qed |
|
811 |
||
812 |
||
813 |
proposition iso_relative_homology_group_upper_hemisphere: |
|
814 |
"(hom_induced p (subtopology (nsphere n) {x. x k \<ge> 0}) {x. x k = 0} (nsphere n) {x. x k \<le> 0} id) |
|
815 |
\<in> iso (relative_homology_group p (subtopology (nsphere n) {x. x k \<ge> 0}) {x. x k = 0}) |
|
816 |
(relative_homology_group p (nsphere n) {x. x k \<le> 0})" (is "?h \<in> iso ?G ?H") |
|
817 |
proof - |
|
818 |
have "topspace (nsphere n) \<inter> {x. x k < - 1 / 2} \<subseteq> {x \<in> topspace (nsphere n). x k \<in> {y. y \<le> - 1 / 2}}" |
|
819 |
by force |
|
820 |
moreover have "closedin (nsphere n) {x \<in> topspace (nsphere n). x k \<in> {y. y \<le> - 1 / 2}}" |
|
821 |
apply (rule closedin_continuous_map_preimage [OF continuous_map_nsphere_projection]) |
|
822 |
using closed_Collect_le [of id "\<lambda>x::real. -1/2"] apply simp |
|
823 |
done |
|
824 |
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}}" |
|
825 |
by (metis (no_types, lifting) closure_of_eq closure_of_mono closure_of_restrict) |
|
826 |
also have "\<dots> \<subseteq> {x \<in> topspace (nsphere n). x k \<in> {y. y < 0}}" |
|
827 |
by force |
|
828 |
also have "\<dots> \<subseteq> nsphere n interior_of {x. x k \<le> 0}" |
|
829 |
proof (rule interior_of_maximal) |
|
830 |
show "{x \<in> topspace (nsphere n). x k \<in> {y. y < 0}} \<subseteq> {x. x k \<le> 0}" |
|
831 |
by force |
|
832 |
show "openin (nsphere n) {x \<in> topspace (nsphere n). x k \<in> {y. y < 0}}" |
|
833 |
apply (rule openin_continuous_map_preimage [OF continuous_map_nsphere_projection]) |
|
834 |
using open_Collect_less [of id "\<lambda>x::real. 0"] apply simp |
|
835 |
done |
|
836 |
qed |
|
837 |
finally have nn: "nsphere n closure_of {x. x k < -1/2} \<subseteq> nsphere n interior_of {x. x k \<le> 0}" . |
|
838 |
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}" |
|
839 |
"UNIV - {x::nat\<Rightarrow>real. x k < a} = {x. a \<le> x k}" for a |
|
840 |
by auto |
|
841 |
let ?T01 = "top_of_set {0..1::real}" |
|
842 |
let ?X12 = "subtopology (nsphere n) {x. -1/2 \<le> x k}" |
|
843 |
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 |
|
844 |
\<in> iso (relative_homology_group p ?X12 {x. -1/2 \<le> x k \<and> x k \<le> 0}) |
|
845 |
?H" |
|
846 |
using homology_excision_axiom [OF nn subset_UNIV, of p] by simp |
|
847 |
define h where "h \<equiv> \<lambda>(T,x). let y = max (x k) (-T) in |
|
848 |
(\<lambda>i. if i = k then y else sqrt(1 - y ^ 2) / sqrt(1 - x k ^ 2) * x i)" |
|
849 |
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 |
|
850 |
using that by (force simp: nsphere h_def Let_def max_def intro!: topspace_nsphere_1_eq_0) |
|
851 |
have "continuous_map (prod_topology ?T01 ?X12) euclideanreal (\<lambda>x. h x i)" for i |
|
852 |
proof - |
|
853 |
show ?thesis |
|
854 |
proof (rule continuous_map_eq) |
|
855 |
show "continuous_map (prod_topology ?T01 ?X12) |
|
856 |
euclideanreal (\<lambda>(T, x). if 0 \<le> x k then x i else h (T, x) i)" |
|
857 |
unfolding case_prod_unfold |
|
858 |
proof (rule continuous_map_cases_le) |
|
859 |
show "continuous_map (prod_topology ?T01 ?X12) euclideanreal (\<lambda>x. snd x k)" |
|
860 |
apply (subst continuous_map_of_snd [unfolded o_def]) |
|
861 |
by (simp add: continuous_map_from_subtopology continuous_map_nsphere_projection) |
|
862 |
next |
|
863 |
show "continuous_map (subtopology (prod_topology ?T01 ?X12) {p \<in> topspace (prod_topology ?T01 ?X12). 0 \<le> snd p k}) |
|
864 |
euclideanreal (\<lambda>x. snd x i)" |
|
865 |
apply (rule continuous_map_from_subtopology) |
|
866 |
apply (subst continuous_map_of_snd [unfolded o_def]) |
|
867 |
by (simp add: continuous_map_from_subtopology continuous_map_nsphere_projection) |
|
868 |
next |
|
869 |
note fst = continuous_map_into_fulltopology [OF continuous_map_subtopology_fst] |
|
870 |
have snd: "continuous_map (subtopology (prod_topology ?T01 (subtopology (nsphere n) T)) S) euclideanreal (\<lambda>x. snd x k)" for k S T |
|
871 |
apply (simp add: nsphere) |
|
872 |
apply (rule continuous_map_from_subtopology) |
|
873 |
apply (subst continuous_map_of_snd [unfolded o_def]) |
|
874 |
using continuous_map_from_subtopology continuous_map_nsphere_projection nsphere by fastforce |
|
875 |
show "continuous_map (subtopology (prod_topology ?T01 ?X12) {p \<in> topspace (prod_topology ?T01 ?X12). snd p k \<le> 0}) |
|
876 |
euclideanreal (\<lambda>x. h (fst x, snd x) i)" |
|
877 |
apply (simp add: h_def case_prod_unfold Let_def) |
|
878 |
apply (intro conjI impI fst snd continuous_intros) |
|
879 |
apply (auto simp: nsphere power2_eq_1_iff) |
|
880 |
done |
|
881 |
qed (auto simp: nsphere h) |
|
882 |
qed (auto simp: nsphere h) |
|
883 |
qed |
|
884 |
moreover |
|
885 |
have "h ` ({0..1} \<times> (topspace (nsphere n) \<inter> {x. - (1/2) \<le> x k})) |
|
886 |
\<subseteq> {x. (\<Sum>i\<le>n. (x i)\<^sup>2) = 1 \<and> (\<forall>i>n. x i = 0)}" |
|
887 |
proof - |
|
888 |
have "(\<Sum>i\<le>n. (h (T,x) i)\<^sup>2) = 1" |
|
889 |
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 |
|
890 |
proof (cases "-T \<le> x k ") |
|
891 |
case True |
|
892 |
then show ?thesis |
|
893 |
using that by (auto simp: nsphere h) |
|
894 |
next |
|
895 |
case False |
|
896 |
with x \<open>0 \<le> T\<close> have "k \<le> n" |
|
897 |
apply (simp add: nsphere) |
|
898 |
by (metis neg_le_0_iff_le not_le) |
|
899 |
have "1 - (x k)\<^sup>2 \<ge> 0" |
|
900 |
using topspace_nsphere_1 x by auto |
|
901 |
with False T \<open>k \<le> n\<close> |
|
902 |
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))" |
|
903 |
unfolding h_def Let_def max_def |
|
904 |
by (simp add: not_le square_le_1 power_mult_distrib power_divide if_distrib [of "\<lambda>x. x ^ 2"] |
|
905 |
sum.delta_remove sum_distrib_left) |
|
906 |
also have "\<dots> = 1" |
|
907 |
using x False xk \<open>0 \<le> T\<close> |
|
908 |
by (simp add: nsphere sum_diff not_le \<open>k \<le> n\<close> power2_eq_1_iff flip: sum_divide_distrib) |
|
909 |
finally show ?thesis . |
|
910 |
qed |
|
911 |
moreover |
|
912 |
have "h (T,x) i = 0" |
|
913 |
if "x \<in> topspace (nsphere n)" "- (1/2) \<le> x k" and "n < i" "0 \<le> T" "T \<le> 1" |
|
914 |
for T x i |
|
915 |
proof (cases "-T \<le> x k ") |
|
916 |
case False |
|
917 |
then show ?thesis |
|
918 |
using that by (auto simp: nsphere h_def Let_def not_le max_def) |
|
919 |
qed (use that in \<open>auto simp: nsphere h\<close>) |
|
920 |
ultimately show ?thesis |
|
921 |
by auto |
|
922 |
qed |
|
923 |
ultimately |
|
924 |
have cmh: "continuous_map (prod_topology ?T01 ?X12) (nsphere n) h" |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
925 |
proof (subst (2) nsphere) |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
926 |
qed (fastforce simp add: continuous_map_in_subtopology continuous_map_componentwise_UNIV) |
70095 | 927 |
have "hom_induced p (subtopology (nsphere n) {x. 0 \<le> x k}) |
928 |
(topspace (subtopology (nsphere n) {x. 0 \<le> x k}) \<inter> {x. x k = 0}) ?X12 |
|
929 |
(topspace ?X12 \<inter> {x. - 1/2 \<le> x k \<and> x k \<le> 0}) id |
|
930 |
\<in> iso (relative_homology_group p (subtopology (nsphere n) {x. 0 \<le> x k}) |
|
931 |
(topspace (subtopology (nsphere n) {x. 0 \<le> x k}) \<inter> {x. x k = 0})) |
|
932 |
(relative_homology_group p ?X12 (topspace ?X12 \<inter> {x. - 1/2 \<le> x k \<and> x k \<le> 0}))" |
|
933 |
proof (rule deformation_retract_relative_homology_group_isomorphism_id) |
|
934 |
show "retraction_maps ?X12 (subtopology (nsphere n) {x. 0 \<le> x k}) (h \<circ> (\<lambda>x. (0,x))) id" |
|
935 |
unfolding retraction_maps_def |
|
936 |
proof (intro conjI ballI) |
|
937 |
show "continuous_map ?X12 (subtopology (nsphere n) {x. 0 \<le> x k}) (h \<circ> Pair 0)" |
|
938 |
apply (simp add: continuous_map_in_subtopology) |
|
939 |
apply (intro conjI continuous_map_compose [OF _ cmh] continuous_intros) |
|
940 |
apply (auto simp: h_def Let_def) |
|
941 |
done |
|
942 |
show "continuous_map (subtopology (nsphere n) {x. 0 \<le> x k}) ?X12 id" |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
943 |
by (simp add: continuous_map_in_subtopology) |
70095 | 944 |
qed (simp add: nsphere h) |
945 |
next |
|
946 |
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" |
|
947 |
by (simp add: h_def Let_def) |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
948 |
show "(h \<circ> (\<lambda>x. (0,x))) \<in> (topspace ?X12 \<inter> {x. - 1 / 2 \<le> x k \<and> x k \<le> 0}) |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
949 |
\<rightarrow> topspace (subtopology (nsphere n) {x. 0 \<le> x k}) \<inter> {x. x k = 0}" |
70095 | 950 |
apply (auto simp: h0) |
951 |
apply (rule subsetD [OF continuous_map_image_subset_topspace [OF cmh]]) |
|
952 |
apply (force simp: nsphere) |
|
953 |
done |
|
954 |
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)" |
|
955 |
apply (rule subsetD [OF continuous_map_image_subset_topspace [OF cmh]]) |
|
956 |
apply (force simp: nsphere) |
|
957 |
done |
|
958 |
have h1: "\<And>x. \<lbrakk>x \<in> topspace (nsphere n); - (1/2) \<le> x k\<rbrakk> \<Longrightarrow> h (1, x) = x" |
|
959 |
by (simp add: h nsphere) |
|
960 |
have "continuous_map (prod_topology ?T01 ?X12) (nsphere n) h" |
|
961 |
using cmh by force |
|
962 |
then show "homotopic_with |
|
963 |
(\<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}) |
|
964 |
?X12 ?X12 (h \<circ> (\<lambda>x. (0,x))) id" |
|
965 |
apply (subst homotopic_with, force) |
|
966 |
apply (rule_tac x=h in exI) |
|
967 |
apply (auto simp: hin h1 continuous_map_in_subtopology) |
|
968 |
apply (auto simp: h_def Let_def max_def) |
|
969 |
done |
|
970 |
qed auto |
|
971 |
then have 2: "hom_induced p (subtopology (nsphere n) {x. 0 \<le> x k}) {x. x k = 0} |
|
972 |
?X12 {x. - 1/2 \<le> x k \<and> x k \<le> 0} id |
|
973 |
\<in> Group.iso |
|
974 |
(relative_homology_group p (subtopology (nsphere n) {x. 0 \<le> x k}) {x. x k = 0}) |
|
975 |
(relative_homology_group p ?X12 {x. - 1/2 \<le> x k \<and> x k \<le> 0})" |
|
976 |
by (metis hom_induced_restrict relative_homology_group_restrict topspace_subtopology) |
|
977 |
show ?thesis |
|
978 |
using iso_set_trans [OF 2 1] |
|
979 |
by (simp add: subset_iff continuous_map_in_subtopology flip: hom_induced_compose) |
|
980 |
qed |
|
981 |
||
982 |
||
983 |
corollary iso_upper_hemisphere_reduced_homology_group: |
|
984 |
"(hom_boundary (1 + p) (subtopology (nsphere (Suc n)) {x. x(Suc n) \<ge> 0}) {x. x(Suc n) = 0}) |
|
985 |
\<in> iso (relative_homology_group (1 + p) (subtopology (nsphere (Suc n)) {x. x(Suc n) \<ge> 0}) {x. x(Suc n) = 0}) |
|
986 |
(reduced_homology_group p (nsphere n))" |
|
987 |
proof - |
|
988 |
have "{x. 0 \<le> x (Suc n)} \<inter> {x. x (Suc n) = 0} = {x. x (Suc n) = (0::real)}" |
|
989 |
by auto |
|
990 |
then have n: "nsphere n = subtopology (subtopology (nsphere (Suc n)) {x. x(Suc n) \<ge> 0}) {x. x(Suc n) = 0}" |
|
991 |
by (simp add: subtopology_nsphere_equator subtopology_subtopology) |
|
992 |
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}" |
|
993 |
by (simp add: nsphere if_distrib [of "\<lambda>x. x ^ 2"] cong: if_cong) |
|
994 |
show ?thesis |
|
995 |
unfolding n |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
996 |
using iso_relative_homology_of_contractible [where p = "1 + p", simplified] |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
997 |
by (metis contractible_space_upper_hemisphere dual_order.refl empty_iff ne) |
70095 | 998 |
qed |
999 |
||
1000 |
corollary iso_reduced_homology_group_upper_hemisphere: |
|
1001 |
assumes "k \<le> n" |
|
1002 |
shows "hom_induced p (nsphere n) {} (nsphere n) {x. x k \<ge> 0} id |
|
1003 |
\<in> iso (reduced_homology_group p (nsphere n)) (relative_homology_group p (nsphere n) {x. x k \<ge> 0})" |
|
1004 |
proof (rule iso_reduced_homology_by_contractible [OF contractible_space_upper_hemisphere [OF assms]]) |
|
1005 |
have "(\<lambda>i. if i = k then 1 else 0) \<in> topspace (nsphere n) \<inter> {x. 0 \<le> x k}" |
|
1006 |
using assms by (simp add: nsphere if_distrib [of "\<lambda>x. x ^ 2"] cong: if_cong) |
|
1007 |
then show "topspace (nsphere n) \<inter> {x. 0 \<le> x k} \<noteq> {}" |
|
1008 |
by blast |
|
1009 |
qed |
|
1010 |
||
1011 |
||
1012 |
lemma iso_relative_homology_group_lower_hemisphere: |
|
1013 |
"hom_induced p (subtopology (nsphere n) {x. x k \<le> 0}) {x. x k = 0} (nsphere n) {x. x k \<ge> 0} id |
|
1014 |
\<in> iso (relative_homology_group p (subtopology (nsphere n) {x. x k \<le> 0}) {x. x k = 0}) |
|
1015 |
(relative_homology_group p (nsphere n) {x. x k \<ge> 0})" (is "?k \<in> iso ?G ?H") |
|
1016 |
proof - |
|
1017 |
define r where "r \<equiv> \<lambda>x i. if i = k then -x i else (x i::real)" |
|
1018 |
then have [simp]: "r \<circ> r = id" |
|
1019 |
by force |
|
1020 |
have cmr: "continuous_map (subtopology (nsphere n) S) (nsphere n) r" for S |
|
1021 |
using continuous_map_nsphere_reflection [of n k] |
|
1022 |
by (simp add: continuous_map_from_subtopology r_def) |
|
1023 |
let ?f = "hom_induced p (subtopology (nsphere n) {x. x k \<le> 0}) {x. x k = 0} |
|
1024 |
(subtopology (nsphere n) {x. x k \<ge> 0}) {x. x k = 0} r" |
|
1025 |
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" |
|
1026 |
let ?h = "hom_induced p (nsphere n) {x. x k \<le> 0} (nsphere n) {x. x k \<ge> 0} r" |
|
1027 |
obtain f h where |
|
1028 |
f: "f \<in> iso ?G (relative_homology_group p (subtopology (nsphere n) {x. x k \<ge> 0}) {x. x k = 0})" |
|
1029 |
and h: "h \<in> iso (relative_homology_group p (nsphere n) {x. x k \<le> 0}) ?H" |
|
1030 |
and eq: "h \<circ> ?g \<circ> f = ?k" |
|
1031 |
proof |
|
1032 |
have hmr: "homeomorphic_map (nsphere n) (nsphere n) r" |
|
1033 |
unfolding homeomorphic_map_maps |
|
1034 |
by (metis \<open>r \<circ> r = id\<close> cmr homeomorphic_maps_involution pointfree_idE subtopology_topspace) |
|
1035 |
then have hmrs: "homeomorphic_map (subtopology (nsphere n) {x. x k \<le> 0}) (subtopology (nsphere n) {x. x k \<ge> 0}) r" |
|
1036 |
by (simp add: homeomorphic_map_subtopologies_alt r_def) |
|
1037 |
have rimeq: "r ` (topspace (subtopology (nsphere n) {x. x k \<le> 0}) \<inter> {x. x k = 0}) |
|
1038 |
= topspace (subtopology (nsphere n) {x. 0 \<le> x k}) \<inter> {x. x k = 0}" |
|
1039 |
using continuous_map_eq_topcontinuous_at continuous_map_nsphere_reflection topcontinuous_at_atin |
|
78322
74c75da4cb01
Some fixes, and SOME TIME LIMITS
paulson <lp15@cam.ac.uk>
parents:
73932
diff
changeset
|
1040 |
by (fastforce simp: r_def Pi_iff) |
70095 | 1041 |
show "?f \<in> iso ?G (relative_homology_group p (subtopology (nsphere n) {x. x k \<ge> 0}) {x. x k = 0})" |
1042 |
using homeomorphic_map_relative_homology_iso [OF hmrs Int_lower1 rimeq] |
|
1043 |
by (metis hom_induced_restrict relative_homology_group_restrict) |
|
1044 |
have rimeq: "r ` (topspace (nsphere n) \<inter> {x. x k \<le> 0}) = topspace (nsphere n) \<inter> {x. 0 \<le> x k}" |
|
1045 |
by (metis hmrs homeomorphic_imp_surjective_map topspace_subtopology) |
|
1046 |
show "?h \<in> Group.iso (relative_homology_group p (nsphere n) {x. x k \<le> 0}) ?H" |
|
1047 |
using homeomorphic_map_relative_homology_iso [OF hmr Int_lower1 rimeq] by simp |
|
1048 |
have [simp]: "\<And>x. x k = 0 \<Longrightarrow> r x k = 0" |
|
1049 |
by (auto simp: r_def) |
|
1050 |
have "?h \<circ> ?g \<circ> ?f |
|
1051 |
= hom_induced p (subtopology (nsphere n) {x. 0 \<le> x k}) {x. x k = 0} (nsphere n) {x. 0 \<le> x k} r \<circ> |
|
1052 |
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" |
|
1053 |
apply (subst hom_induced_compose [symmetric]) |
|
1054 |
using continuous_map_nsphere_reflection apply (force simp: r_def)+ |
|
1055 |
done |
|
1056 |
also have "\<dots> = ?k" |
|
1057 |
apply (subst hom_induced_compose [symmetric]) |
|
1058 |
apply (simp_all add: image_subset_iff cmr) |
|
1059 |
using hmrs homeomorphic_imp_continuous_map apply blast |
|
1060 |
done |
|
1061 |
finally show "?h \<circ> ?g \<circ> ?f = ?k" . |
|
1062 |
qed |
|
1063 |
with iso_relative_homology_group_upper_hemisphere [of p n k] |
|
1064 |
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 |
|
1065 |
\<in> Group.iso ?G (relative_homology_group p (nsphere n) {f. 0 \<le> f k})" |
|
1066 |
using f h iso_set_trans by blast |
|
1067 |
then show ?thesis |
|
1068 |
by (simp add: eq) |
|
1069 |
qed |
|
1070 |
||
1071 |
||
1072 |
lemma iso_lower_hemisphere_reduced_homology_group: |
|
1073 |
"hom_boundary (1 + p) (subtopology (nsphere (Suc n)) {x. x(Suc n) \<le> 0}) {x. x(Suc n) = 0} |
|
1074 |
\<in> iso (relative_homology_group (1 + p) (subtopology (nsphere (Suc n)) {x. x(Suc n) \<le> 0}) |
|
1075 |
{x. x(Suc n) = 0}) |
|
1076 |
(reduced_homology_group p (nsphere n))" |
|
1077 |
proof - |
|
1078 |
have "{x. (\<Sum>i\<le>n. (x i)\<^sup>2) = 1 \<and> (\<forall>i>n. x i = 0)} = |
|
1079 |
({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> |
|
1080 |
{x. x (Suc n) = (0::real)})" |
|
1081 |
by (force simp: dest: Suc_lessI) |
|
1082 |
then have n: "nsphere n = subtopology (subtopology (nsphere (Suc n)) {x. x(Suc n) \<le> 0}) {x. x(Suc n) = 0}" |
|
1083 |
by (simp add: nsphere subtopology_subtopology) |
|
1084 |
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}" |
|
1085 |
by (simp add: nsphere if_distrib [of "\<lambda>x. x ^ 2"] cong: if_cong) |
|
1086 |
show ?thesis |
|
1087 |
unfolding n |
|
1088 |
apply (rule iso_relative_homology_of_contractible [where p = "1 + p", simplified]) |
|
1089 |
using contractible_space_lower_hemisphere ne apply blast+ |
|
1090 |
done |
|
1091 |
qed |
|
1092 |
||
1093 |
lemma isomorphism_sym: |
|
1094 |
"\<lbrakk>f \<in> iso G1 G2; \<And>x. x \<in> carrier G1 \<Longrightarrow> r'(f x) = f(r x); |
|
1095 |
\<And>x. x \<in> carrier G1 \<Longrightarrow> r x \<in> carrier G1; group G1; group G2\<rbrakk> |
|
1096 |
\<Longrightarrow> \<exists>f \<in> iso G2 G1. \<forall>x \<in> carrier G2. r(f x) = f(r' x)" |
|
1097 |
apply (clarsimp simp add: group.iso_iff_group_isomorphisms Bex_def) |
|
1098 |
by (metis (full_types) group_isomorphisms_def group_isomorphisms_sym hom_in_carrier) |
|
1099 |
||
1100 |
lemma isomorphism_trans: |
|
1101 |
"\<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> |
|
1102 |
\<Longrightarrow> \<exists>f \<in> iso G1 G3. \<forall>x \<in> carrier G1. r3(f x) = f(r1 x)" |
|
1103 |
apply clarify |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1104 |
by (smt (verit, ccfv_threshold) Group.iso_iff hom_in_carrier iso_set_trans |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1105 |
o_apply) |
70095 | 1106 |
|
1107 |
lemma reduced_homology_group_nsphere_step: |
|
1108 |
"\<exists>f \<in> iso(reduced_homology_group p (nsphere n)) |
|
1109 |
(reduced_homology_group (1 + p) (nsphere (Suc n))). |
|
1110 |
\<forall>c \<in> carrier(reduced_homology_group p (nsphere n)). |
|
1111 |
hom_induced (1 + p) (nsphere(Suc n)) {} (nsphere(Suc n)) {} |
|
1112 |
(\<lambda>x i. if i = 0 then -x i else x i) (f c) |
|
1113 |
= f (hom_induced p (nsphere n) {} (nsphere n) {} (\<lambda>x i. if i = 0 then -x i else x i) c)" |
|
1114 |
proof - |
|
1115 |
define r where "r \<equiv> \<lambda>x::nat\<Rightarrow>real. \<lambda>i. if i = 0 then -x i else x i" |
|
1116 |
have cmr: "continuous_map (nsphere n) (nsphere n) r" for n |
|
1117 |
unfolding r_def by (rule continuous_map_nsphere_reflection) |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1118 |
have rsub: "r \<in> {x. 0 \<le> x (Suc n)} \<rightarrow> {x. 0 \<le> x (Suc n)}" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1119 |
"r \<in> {x. x (Suc n) \<le> 0} \<rightarrow> {x. x (Suc n) \<le> 0}" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1120 |
"r \<in> {x. x (Suc n) = 0} \<rightarrow> {x. x (Suc n) = 0}" |
70095 | 1121 |
by (force simp: r_def)+ |
1122 |
let ?sub = "subtopology (nsphere (Suc n)) {x. x (Suc n) \<ge> 0}" |
|
1123 |
let ?G2 = "relative_homology_group (1 + p) ?sub {x. x (Suc n) = 0}" |
|
1124 |
let ?r2 = "hom_induced (1 + p) ?sub {x. x (Suc n) = 0} ?sub {x. x (Suc n) = 0} r" |
|
1125 |
let ?j = "\<lambda>p n. hom_induced p (nsphere n) {} (nsphere n) {} r" |
|
1126 |
show ?thesis |
|
1127 |
unfolding r_def [symmetric] |
|
1128 |
proof (rule isomorphism_trans) |
|
1129 |
let ?f = "hom_boundary (1 + p) ?sub {x. x (Suc n) = 0}" |
|
1130 |
show "\<exists>f\<in>Group.iso (reduced_homology_group p (nsphere n)) ?G2. |
|
1131 |
\<forall>c\<in>carrier (reduced_homology_group p (nsphere n)). ?r2 (f c) = f (?j p n c)" |
|
1132 |
proof (rule isomorphism_sym) |
|
1133 |
show "?f \<in> Group.iso ?G2 (reduced_homology_group p (nsphere n))" |
|
1134 |
using iso_upper_hemisphere_reduced_homology_group |
|
1135 |
by (metis add.commute) |
|
1136 |
next |
|
1137 |
fix c |
|
1138 |
assume "c \<in> carrier ?G2" |
|
1139 |
have cmrs: "continuous_map ?sub ?sub r" |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1140 |
by (metis (no_types, lifting) IntE Pi_iff cmr continuous_map_from_subtopology |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1141 |
continuous_map_into_subtopology rsub(1) topspace_subtopology) |
70095 | 1142 |
have "hom_induced p (nsphere n) {} (nsphere n) {} r \<circ> hom_boundary (1 + p) ?sub {x. x (Suc n) = 0} |
1143 |
= hom_boundary (1 + p) ?sub {x. x (Suc n) = 0} \<circ> |
|
1144 |
hom_induced (1 + p) ?sub {x. x (Suc n) = 0} ?sub {x. x (Suc n) = 0} r" |
|
1145 |
using naturality_hom_induced [OF cmrs rsub(3), symmetric, of "1+p", simplified] |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1146 |
by (simp add: Pi_iff subtopology_subtopology subtopology_nsphere_equator flip: Collect_conj_eq cong: rev_conj_cong) |
70095 | 1147 |
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)" |
1148 |
by (metis comp_def) |
|
1149 |
next |
|
1150 |
fix c |
|
1151 |
assume "c \<in> carrier ?G2" |
|
1152 |
show "hom_induced (1 + p) ?sub {x. x (Suc n) = 0} ?sub {x. x (Suc n) = 0} r c \<in> carrier ?G2" |
|
1153 |
using hom_induced_carrier by blast |
|
1154 |
qed auto |
|
1155 |
next |
|
1156 |
let ?H2 = "relative_homology_group (1 + p) (nsphere (Suc n)) {x. x (Suc n) \<le> 0}" |
|
1157 |
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" |
|
1158 |
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) |
|
1159 |
= f (?r2 c)" |
|
1160 |
proof (rule isomorphism_trans) |
|
1161 |
show "\<exists>f\<in>Group.iso ?G2 ?H2. |
|
1162 |
\<forall>c\<in>carrier ?G2. |
|
1163 |
?s2 (f c) = f (hom_induced (1 + p) ?sub {x. x (Suc n) = 0} ?sub {x. x (Suc n) = 0} r c)" |
|
1164 |
proof (intro ballI bexI) |
|
1165 |
fix c |
|
1166 |
assume "c \<in> carrier (relative_homology_group (1 + p) ?sub {x. x (Suc n) = 0})" |
|
1167 |
show "?s2 (hom_induced (1 + p) ?sub {x. x (Suc n) = 0} (nsphere (Suc n)) {x. x (Suc n) \<le> 0} id c) |
|
1168 |
= hom_induced (1 + p) ?sub {x. x (Suc n) = 0} (nsphere (Suc n)) {x. x (Suc n) \<le> 0} id (?r2 c)" |
|
1169 |
apply (simp add: rsub hom_induced_compose' Collect_mono_iff cmr) |
|
1170 |
apply (subst hom_induced_compose') |
|
1171 |
apply (simp_all add: continuous_map_in_subtopology continuous_map_from_subtopology [OF cmr] rsub) |
|
1172 |
apply (auto simp: r_def) |
|
1173 |
done |
|
1174 |
qed (simp add: iso_relative_homology_group_upper_hemisphere) |
|
1175 |
next |
|
1176 |
let ?h = "hom_induced (1 + p) (nsphere(Suc n)) {} (nsphere (Suc n)) {x. x(Suc n) \<le> 0} id" |
|
1177 |
show "\<exists>f\<in>Group.iso ?H2 (reduced_homology_group (1 + p) (nsphere (Suc n))). |
|
1178 |
\<forall>c\<in>carrier ?H2. ?j (1 + p) (Suc n) (f c) = f (?s2 c)" |
|
1179 |
proof (rule isomorphism_sym) |
|
1180 |
show "?h \<in> Group.iso (reduced_homology_group (1 + p) (nsphere (Suc n))) |
|
1181 |
(relative_homology_group (1 + p) (nsphere (Suc n)) {x. x (Suc n) \<le> 0})" |
|
1182 |
using iso_reduced_homology_group_lower_hemisphere by blast |
|
1183 |
next |
|
1184 |
fix c |
|
1185 |
assume "c \<in> carrier (reduced_homology_group (1 + p) (nsphere (Suc n)))" |
|
1186 |
show "?s2 (?h c) = ?h (?j (1 + p) (Suc n) c)" |
|
1187 |
by (simp add: hom_induced_compose' cmr rsub) |
|
1188 |
next |
|
1189 |
fix c |
|
1190 |
assume "c \<in> carrier (reduced_homology_group (1 + p) (nsphere (Suc n)))" |
|
1191 |
then show "hom_induced (1 + p) (nsphere (Suc n)) {} (nsphere (Suc n)) {} r c |
|
1192 |
\<in> carrier (reduced_homology_group (1 + p) (nsphere (Suc n)))" |
|
1193 |
by (simp add: hom_induced_reduced) |
|
1194 |
qed auto |
|
1195 |
qed |
|
1196 |
qed |
|
1197 |
qed |
|
1198 |
||
1199 |
||
1200 |
lemma reduced_homology_group_nsphere_aux: |
|
1201 |
"if p = int n then reduced_homology_group n (nsphere n) \<cong> integer_group |
|
1202 |
else trivial_group(reduced_homology_group p (nsphere n))" |
|
1203 |
proof (induction n arbitrary: p) |
|
1204 |
case 0 |
|
1205 |
let ?a = "\<lambda>i::nat. if i = 0 then 1 else (0::real)" |
|
1206 |
let ?b = "\<lambda>i::nat. if i = 0 then -1 else (0::real)" |
|
1207 |
have st: "subtopology (powertop_real UNIV) {?a, ?b} = nsphere 0" |
|
1208 |
proof - |
|
1209 |
have "{?a, ?b} = {x. (x 0)\<^sup>2 = 1 \<and> (\<forall>i>0. x i = 0)}" |
|
1210 |
using power2_eq_iff by fastforce |
|
1211 |
then show ?thesis |
|
1212 |
by (simp add: nsphere) |
|
1213 |
qed |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1214 |
have "t1_space (powertop_real UNIV)" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1215 |
using t1_space_euclidean t1_space_product_topology by blast |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1216 |
then have *: "reduced_homology_group p (subtopology (powertop_real UNIV) {?a, ?b}) \<cong> |
70095 | 1217 |
homology_group p (subtopology (powertop_real UNIV) {?a})" |
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1218 |
by (intro reduced_homology_group_pair) (auto simp: fun_eq_iff) |
70095 | 1219 |
have "reduced_homology_group 0 (nsphere 0) \<cong> integer_group" if "p=0" |
1220 |
proof - |
|
1221 |
have "reduced_homology_group 0 (nsphere 0) \<cong> homology_group 0 (top_of_set {?a})" if "p=0" |
|
1222 |
by (metis * euclidean_product_topology st that) |
|
1223 |
also have "\<dots> \<cong> integer_group" |
|
1224 |
by (simp add: homology_coefficients) |
|
1225 |
finally show ?thesis |
|
1226 |
using that by blast |
|
1227 |
qed |
|
1228 |
moreover have "trivial_group (reduced_homology_group p (nsphere 0))" if "p\<noteq>0" |
|
1229 |
using * that homology_dimension_axiom [of "subtopology (powertop_real UNIV) {?a}" ?a p] |
|
1230 |
using isomorphic_group_triviality st by force |
|
1231 |
ultimately show ?case |
|
1232 |
by auto |
|
1233 |
next |
|
1234 |
case (Suc n) |
|
1235 |
have eq: "reduced_homology_group (int n) (nsphere n) \<cong> integer_group" if "p-1 = n" |
|
1236 |
by (simp add: Suc.IH) |
|
1237 |
have neq: "trivial_group (reduced_homology_group (p-1) (nsphere n))" if "p-1 \<noteq> n" |
|
1238 |
by (simp add: Suc.IH that) |
|
1239 |
have iso: "reduced_homology_group p (nsphere (Suc n)) \<cong> reduced_homology_group (p-1) (nsphere n)" |
|
1240 |
using reduced_homology_group_nsphere_step [of "p-1" n] group.iso_sym [OF _ is_isoI] group_reduced_homology_group |
|
1241 |
by fastforce |
|
1242 |
then show ?case |
|
1243 |
using eq iso_trans iso isomorphic_group_triviality neq |
|
73932
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
72632
diff
changeset
|
1244 |
by (metis (no_types, opaque_lifting) add.commute add_left_cancel diff_add_cancel group_reduced_homology_group of_nat_Suc) |
70095 | 1245 |
qed |
1246 |
||
1247 |
||
1248 |
lemma reduced_homology_group_nsphere: |
|
1249 |
"reduced_homology_group n (nsphere n) \<cong> integer_group" |
|
1250 |
"p \<noteq> n \<Longrightarrow> trivial_group(reduced_homology_group p (nsphere n))" |
|
1251 |
using reduced_homology_group_nsphere_aux by auto |
|
1252 |
||
1253 |
lemma cyclic_reduced_homology_group_nsphere: |
|
1254 |
"cyclic_group(reduced_homology_group p (nsphere n))" |
|
1255 |
by (metis reduced_homology_group_nsphere trivial_imp_cyclic_group cyclic_integer_group |
|
1256 |
group_integer_group group_reduced_homology_group isomorphic_group_cyclicity) |
|
1257 |
||
1258 |
lemma trivial_reduced_homology_group_nsphere: |
|
1259 |
"trivial_group(reduced_homology_group p (nsphere n)) \<longleftrightarrow> (p \<noteq> n)" |
|
1260 |
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 |
|
1261 |
||
1262 |
lemma non_contractible_space_nsphere: "\<not> (contractible_space(nsphere n))" |
|
1263 |
proof (clarsimp simp add: contractible_eq_homotopy_equivalent_singleton_subtopology) |
|
1264 |
fix a :: "nat \<Rightarrow> real" |
|
1265 |
assume a: "a \<in> topspace (nsphere n)" |
|
1266 |
and he: "nsphere n homotopy_equivalent_space subtopology (nsphere n) {a}" |
|
1267 |
have "trivial_group (reduced_homology_group (int n) (subtopology (nsphere n) {a}))" |
|
1268 |
by (simp add: a homology_dimension_reduced [where a=a]) |
|
1269 |
then show "False" |
|
1270 |
using isomorphic_group_triviality [OF homotopy_equivalent_space_imp_isomorphic_reduced_homology_groups [OF he, of n]] |
|
1271 |
by (simp add: trivial_reduced_homology_group_nsphere) |
|
1272 |
qed |
|
1273 |
||
1274 |
||
1275 |
subsection\<open>Brouwer degree of a Map\<close> |
|
1276 |
||
1277 |
definition Brouwer_degree2 :: "nat \<Rightarrow> ((nat \<Rightarrow> real) \<Rightarrow> nat \<Rightarrow> real) \<Rightarrow> int" |
|
1278 |
where |
|
1279 |
"Brouwer_degree2 p f \<equiv> |
|
1280 |
@d::int. \<forall>x \<in> carrier(reduced_homology_group p (nsphere p)). |
|
1281 |
hom_induced p (nsphere p) {} (nsphere p) {} f x = pow (reduced_homology_group p (nsphere p)) x d" |
|
1282 |
||
1283 |
lemma Brouwer_degree2_eq: |
|
1284 |
"(\<And>x. x \<in> topspace(nsphere p) \<Longrightarrow> f x = g x) \<Longrightarrow> Brouwer_degree2 p f = Brouwer_degree2 p g" |
|
1285 |
unfolding Brouwer_degree2_def Ball_def |
|
1286 |
apply (intro Eps_cong all_cong) |
|
1287 |
by (metis (mono_tags, lifting) hom_induced_eq) |
|
1288 |
||
1289 |
lemma Brouwer_degree2: |
|
1290 |
assumes "x \<in> carrier(reduced_homology_group p (nsphere p))" |
|
1291 |
shows "hom_induced p (nsphere p) {} (nsphere p) {} f x |
|
1292 |
= pow (reduced_homology_group p (nsphere p)) x (Brouwer_degree2 p f)" |
|
1293 |
(is "?h x = pow ?G x _") |
|
1294 |
proof (cases "continuous_map(nsphere p) (nsphere p) f") |
|
1295 |
case True |
|
1296 |
interpret group ?G |
|
1297 |
by simp |
|
1298 |
interpret group_hom ?G ?G ?h |
|
1299 |
using hom_induced_reduced_hom group_hom_axioms_def group_hom_def is_group by blast |
|
1300 |
obtain a where a: "a \<in> carrier ?G" |
|
1301 |
and aeq: "subgroup_generated ?G {a} = ?G" |
|
1302 |
using cyclic_reduced_homology_group_nsphere [of p p] by (auto simp: cyclic_group_def) |
|
1303 |
then have carra: "carrier (subgroup_generated ?G {a}) = range (\<lambda>n::int. pow ?G a n)" |
|
1304 |
using carrier_subgroup_generated_by_singleton by blast |
|
1305 |
moreover have "?h a \<in> carrier (subgroup_generated ?G {a})" |
|
1306 |
by (simp add: a aeq hom_induced_reduced) |
|
1307 |
ultimately obtain d::int where d: "?h a = pow ?G a d" |
|
1308 |
by auto |
|
1309 |
have *: "hom_induced (int p) (nsphere p) {} (nsphere p) {} f x = x [^]\<^bsub>?G\<^esub> d" |
|
1310 |
if x: "x \<in> carrier ?G" for x |
|
1311 |
proof - |
|
1312 |
obtain n::int where xeq: "x = pow ?G a n" |
|
79712
658f17274845
new less ad hoc implementation of the 'moura' tactic for skolemization
blanchet
parents:
78322
diff
changeset
|
1313 |
using carra x aeq by auto |
70095 | 1314 |
show ?thesis |
1315 |
by (simp add: xeq a d hom_int_pow int_pow_pow mult.commute) |
|
1316 |
qed |
|
1317 |
show ?thesis |
|
1318 |
unfolding Brouwer_degree2_def |
|
1319 |
apply (rule someI2 [where a=d]) |
|
1320 |
using assms * apply blast+ |
|
1321 |
done |
|
1322 |
next |
|
1323 |
case False |
|
1324 |
show ?thesis |
|
1325 |
unfolding Brouwer_degree2_def |
|
1326 |
by (rule someI2 [where a=0]) (simp_all add: hom_induced_default False one_reduced_homology_group assms) |
|
1327 |
qed |
|
1328 |
||
1329 |
||
1330 |
||
1331 |
lemma Brouwer_degree2_iff: |
|
1332 |
assumes f: "continuous_map (nsphere p) (nsphere p) f" |
|
1333 |
and x: "x \<in> carrier(reduced_homology_group p (nsphere p))" |
|
1334 |
shows "(hom_induced (int p) (nsphere p) {} (nsphere p) {} f x = |
|
1335 |
x [^]\<^bsub>reduced_homology_group (int p) (nsphere p)\<^esub> d) |
|
1336 |
\<longleftrightarrow> (x = \<one>\<^bsub>reduced_homology_group (int p) (nsphere p)\<^esub> \<or> Brouwer_degree2 p f = d)" |
|
1337 |
(is "(?h x = x [^]\<^bsub>?G\<^esub> d) \<longleftrightarrow> _") |
|
1338 |
proof - |
|
1339 |
interpret group "?G" |
|
1340 |
by simp |
|
1341 |
obtain a where a: "a \<in> carrier ?G" |
|
1342 |
and aeq: "subgroup_generated ?G {a} = ?G" |
|
1343 |
using cyclic_reduced_homology_group_nsphere [of p p] by (auto simp: cyclic_group_def) |
|
1344 |
then obtain i::int where i: "x = (a [^]\<^bsub>?G\<^esub> i)" |
|
1345 |
using carrier_subgroup_generated_by_singleton x by fastforce |
|
1346 |
then have "a [^]\<^bsub>?G\<^esub> i \<in> carrier ?G" |
|
1347 |
using x by blast |
|
1348 |
have [simp]: "ord a = 0" |
|
1349 |
by (simp add: a aeq iso_finite [OF reduced_homology_group_nsphere(1)] flip: infinite_cyclic_subgroup_order) |
|
1350 |
show ?thesis |
|
1351 |
by (auto simp: Brouwer_degree2 int_pow_eq_id x i a int_pow_pow int_pow_eq) |
|
1352 |
qed |
|
1353 |
||
1354 |
||
1355 |
lemma Brouwer_degree2_unique: |
|
1356 |
assumes f: "continuous_map (nsphere p) (nsphere p) f" |
|
1357 |
and hi: "\<And>x. x \<in> carrier(reduced_homology_group p (nsphere p)) |
|
1358 |
\<Longrightarrow> hom_induced p (nsphere p) {} (nsphere p) {} f x = pow (reduced_homology_group p (nsphere p)) x d" |
|
1359 |
(is "\<And>x. x \<in> carrier ?G \<Longrightarrow> ?h x = _") |
|
1360 |
shows "Brouwer_degree2 p f = d" |
|
1361 |
proof - |
|
1362 |
obtain a where a: "a \<in> carrier ?G" |
|
1363 |
and aeq: "subgroup_generated ?G {a} = ?G" |
|
1364 |
using cyclic_reduced_homology_group_nsphere [of p p] by (auto simp: cyclic_group_def) |
|
1365 |
show ?thesis |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1366 |
using hi [OF a] unfolding Brouwer_degree2 a |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1367 |
by (metis Brouwer_degree2_iff a aeq f group.trivial_group_subgroup_generated |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1368 |
group_reduced_homology_group subsetI trivial_reduced_homology_group_nsphere) |
70095 | 1369 |
qed |
1370 |
||
1371 |
lemma Brouwer_degree2_unique_generator: |
|
1372 |
assumes f: "continuous_map (nsphere p) (nsphere p) f" |
|
1373 |
and eq: "subgroup_generated (reduced_homology_group p (nsphere p)) {a} |
|
1374 |
= reduced_homology_group p (nsphere p)" |
|
1375 |
and hi: "hom_induced p (nsphere p) {} (nsphere p) {} f a = pow (reduced_homology_group p (nsphere p)) a d" |
|
1376 |
(is "?h a = pow ?G a _") |
|
1377 |
shows "Brouwer_degree2 p f = d" |
|
1378 |
proof (cases "a \<in> carrier ?G") |
|
1379 |
case True |
|
1380 |
then show ?thesis |
|
1381 |
by (metis Brouwer_degree2_iff hi eq f group.trivial_group_subgroup_generated group_reduced_homology_group |
|
1382 |
subset_singleton_iff trivial_reduced_homology_group_nsphere) |
|
1383 |
next |
|
1384 |
case False |
|
1385 |
then show ?thesis |
|
1386 |
using trivial_reduced_homology_group_nsphere [of p p] |
|
1387 |
by (metis group.trivial_group_subgroup_generated_eq disjoint_insert(1) eq group_reduced_homology_group inf_bot_right subset_singleton_iff) |
|
1388 |
qed |
|
1389 |
||
1390 |
lemma Brouwer_degree2_homotopic: |
|
1391 |
assumes "homotopic_with (\<lambda>x. True) (nsphere p) (nsphere p) f g" |
|
1392 |
shows "Brouwer_degree2 p f = Brouwer_degree2 p g" |
|
1393 |
proof - |
|
1394 |
have "continuous_map (nsphere p) (nsphere p) f" |
|
1395 |
using homotopic_with_imp_continuous_maps [OF assms] by auto |
|
1396 |
show ?thesis |
|
1397 |
using Brouwer_degree2_def assms homology_homotopy_empty by fastforce |
|
1398 |
qed |
|
1399 |
||
1400 |
lemma Brouwer_degree2_id [simp]: "Brouwer_degree2 p id = 1" |
|
1401 |
proof (rule Brouwer_degree2_unique) |
|
1402 |
fix x |
|
1403 |
assume x: "x \<in> carrier (reduced_homology_group (int p) (nsphere p))" |
|
1404 |
then have "x \<in> carrier (homology_group (int p) (nsphere p))" |
|
1405 |
using carrier_reduced_homology_group_subset by blast |
|
1406 |
then show "hom_induced (int p) (nsphere p) {} (nsphere p) {} id x = |
|
1407 |
x [^]\<^bsub>reduced_homology_group (int p) (nsphere p)\<^esub> (1::int)" |
|
1408 |
by (simp add: hom_induced_id group.int_pow_1 x) |
|
1409 |
qed auto |
|
1410 |
||
1411 |
lemma Brouwer_degree2_compose: |
|
1412 |
assumes f: "continuous_map (nsphere p) (nsphere p) f" and g: "continuous_map (nsphere p) (nsphere p) g" |
|
1413 |
shows "Brouwer_degree2 p (g \<circ> f) = Brouwer_degree2 p g * Brouwer_degree2 p f" |
|
1414 |
proof (rule Brouwer_degree2_unique) |
|
1415 |
show "continuous_map (nsphere p) (nsphere p) (g \<circ> f)" |
|
1416 |
by (meson continuous_map_compose f g) |
|
1417 |
next |
|
1418 |
fix x |
|
1419 |
assume x: "x \<in> carrier (reduced_homology_group (int p) (nsphere p))" |
|
1420 |
have "hom_induced (int p) (nsphere p) {} (nsphere p) {} (g \<circ> f) = |
|
1421 |
hom_induced (int p) (nsphere p) {} (nsphere p) {} g \<circ> |
|
1422 |
hom_induced (int p) (nsphere p) {} (nsphere p) {} f" |
|
1423 |
by (blast intro: hom_induced_compose [OF f _ g]) |
|
1424 |
with x show "hom_induced (int p) (nsphere p) {} (nsphere p) {} (g \<circ> f) x = |
|
1425 |
x [^]\<^bsub>reduced_homology_group (int p) (nsphere p)\<^esub> (Brouwer_degree2 p g * Brouwer_degree2 p f)" |
|
1426 |
by (simp add: mult.commute hom_induced_reduced flip: Brouwer_degree2 group.int_pow_pow) |
|
1427 |
qed |
|
1428 |
||
1429 |
lemma Brouwer_degree2_homotopy_equivalence: |
|
1430 |
assumes f: "continuous_map (nsphere p) (nsphere p) f" and g: "continuous_map (nsphere p) (nsphere p) g" |
|
1431 |
and hom: "homotopic_with (\<lambda>x. True) (nsphere p) (nsphere p) (f \<circ> g) id" |
|
1432 |
obtains "\<bar>Brouwer_degree2 p f\<bar> = 1" "\<bar>Brouwer_degree2 p g\<bar> = 1" "Brouwer_degree2 p g = Brouwer_degree2 p f" |
|
1433 |
using Brouwer_degree2_homotopic [OF hom] Brouwer_degree2_compose f g zmult_eq_1_iff by auto |
|
1434 |
||
1435 |
lemma Brouwer_degree2_homeomorphic_maps: |
|
1436 |
assumes "homeomorphic_maps (nsphere p) (nsphere p) f g" |
|
1437 |
obtains "\<bar>Brouwer_degree2 p f\<bar> = 1" "\<bar>Brouwer_degree2 p g\<bar> = 1" "Brouwer_degree2 p g = Brouwer_degree2 p f" |
|
1438 |
using assms |
|
1439 |
by (auto simp: homeomorphic_maps_def homotopic_with_equal continuous_map_compose intro: Brouwer_degree2_homotopy_equivalence) |
|
1440 |
||
1441 |
||
1442 |
lemma Brouwer_degree2_retraction_map: |
|
1443 |
assumes "retraction_map (nsphere p) (nsphere p) f" |
|
1444 |
shows "\<bar>Brouwer_degree2 p f\<bar> = 1" |
|
1445 |
proof - |
|
1446 |
obtain g where g: "retraction_maps (nsphere p) (nsphere p) f g" |
|
1447 |
using assms by (auto simp: retraction_map_def) |
|
1448 |
show ?thesis |
|
1449 |
proof (rule Brouwer_degree2_homotopy_equivalence) |
|
1450 |
show "homotopic_with (\<lambda>x. True) (nsphere p) (nsphere p) (f \<circ> g) id" |
|
1451 |
using g apply (auto simp: retraction_maps_def) |
|
1452 |
by (simp add: homotopic_with_equal continuous_map_compose) |
|
1453 |
show "continuous_map (nsphere p) (nsphere p) f" "continuous_map (nsphere p) (nsphere p) g" |
|
1454 |
using g retraction_maps_def by blast+ |
|
1455 |
qed |
|
1456 |
qed |
|
1457 |
||
1458 |
lemma Brouwer_degree2_section_map: |
|
1459 |
assumes "section_map (nsphere p) (nsphere p) f" |
|
1460 |
shows "\<bar>Brouwer_degree2 p f\<bar> = 1" |
|
1461 |
proof - |
|
1462 |
obtain g where g: "retraction_maps (nsphere p) (nsphere p) g f" |
|
1463 |
using assms by (auto simp: section_map_def) |
|
1464 |
show ?thesis |
|
1465 |
proof (rule Brouwer_degree2_homotopy_equivalence) |
|
1466 |
show "homotopic_with (\<lambda>x. True) (nsphere p) (nsphere p) (g \<circ> f) id" |
|
1467 |
using g apply (auto simp: retraction_maps_def) |
|
1468 |
by (simp add: homotopic_with_equal continuous_map_compose) |
|
1469 |
show "continuous_map (nsphere p) (nsphere p) g" "continuous_map (nsphere p) (nsphere p) f" |
|
1470 |
using g retraction_maps_def by blast+ |
|
1471 |
qed |
|
1472 |
qed |
|
1473 |
||
1474 |
lemma Brouwer_degree2_homeomorphic_map: |
|
1475 |
"homeomorphic_map (nsphere p) (nsphere p) f \<Longrightarrow> \<bar>Brouwer_degree2 p f\<bar> = 1" |
|
1476 |
using Brouwer_degree2_retraction_map section_and_retraction_eq_homeomorphic_map by blast |
|
1477 |
||
1478 |
||
1479 |
lemma Brouwer_degree2_nullhomotopic: |
|
1480 |
assumes "homotopic_with (\<lambda>x. True) (nsphere p) (nsphere p) f (\<lambda>x. a)" |
|
1481 |
shows "Brouwer_degree2 p f = 0" |
|
1482 |
proof - |
|
1483 |
have contf: "continuous_map (nsphere p) (nsphere p) f" |
|
1484 |
and contc: "continuous_map (nsphere p) (nsphere p) (\<lambda>x. a)" |
|
1485 |
using homotopic_with_imp_continuous_maps [OF assms] by metis+ |
|
1486 |
have "Brouwer_degree2 p f = Brouwer_degree2 p (\<lambda>x. a)" |
|
1487 |
using Brouwer_degree2_homotopic [OF assms] . |
|
1488 |
moreover |
|
1489 |
let ?G = "reduced_homology_group (int p) (nsphere p)" |
|
1490 |
interpret group ?G |
|
1491 |
by simp |
|
1492 |
have "Brouwer_degree2 p (\<lambda>x. a) = 0" |
|
1493 |
proof (rule Brouwer_degree2_unique [OF contc]) |
|
1494 |
fix c |
|
1495 |
assume c: "c \<in> carrier ?G" |
|
1496 |
have "continuous_map (nsphere p) (subtopology (nsphere p) {a}) (\<lambda>f. a)" |
|
1497 |
using contc continuous_map_in_subtopology by blast |
|
1498 |
then have he: "hom_induced p (nsphere p) {} (nsphere p) {} (\<lambda>x. a) |
|
1499 |
= hom_induced p (subtopology (nsphere p) {a}) {} (nsphere p) {} id \<circ> |
|
1500 |
hom_induced p (nsphere p) {} (subtopology (nsphere p) {a}) {} (\<lambda>x. a)" |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1501 |
by (metis continuous_map_id_subt fun.map_id hom_induced_compose_empty) |
70095 | 1502 |
have 1: "hom_induced p (nsphere p) {} (subtopology (nsphere p) {a}) {} (\<lambda>x. a) c = |
1503 |
\<one>\<^bsub>reduced_homology_group (int p) (subtopology (nsphere p) {a})\<^esub>" |
|
1504 |
using c trivial_reduced_homology_group_contractible_space [of "subtopology (nsphere p) {a}" p] |
|
1505 |
by (simp add: hom_induced_reduced contractible_space_subtopology_singleton trivial_group_subset group.trivial_group_subset subset_iff) |
|
1506 |
show "hom_induced (int p) (nsphere p) {} (nsphere p) {} (\<lambda>x. a) c = |
|
1507 |
c [^]\<^bsub>?G\<^esub> (0::int)" |
|
1508 |
apply (simp add: he 1) |
|
1509 |
using hom_induced_reduced_hom group_hom.hom_one group_hom_axioms_def group_hom_def group_reduced_homology_group by blast |
|
1510 |
qed |
|
1511 |
ultimately show ?thesis |
|
1512 |
by metis |
|
1513 |
qed |
|
1514 |
||
1515 |
||
1516 |
lemma Brouwer_degree2_const: "Brouwer_degree2 p (\<lambda>x. a) = 0" |
|
1517 |
proof (cases "continuous_map(nsphere p) (nsphere p) (\<lambda>x. a)") |
|
1518 |
case True |
|
1519 |
then show ?thesis |
|
1520 |
by (auto intro: Brouwer_degree2_nullhomotopic [where a=a]) |
|
1521 |
next |
|
1522 |
case False |
|
1523 |
let ?G = "reduced_homology_group (int p) (nsphere p)" |
|
1524 |
let ?H = "homology_group (int p) (nsphere p)" |
|
1525 |
interpret group ?G |
|
1526 |
by simp |
|
1527 |
have eq1: "\<one>\<^bsub>?H\<^esub> = \<one>\<^bsub>?G\<^esub>" |
|
1528 |
by (simp add: one_reduced_homology_group) |
|
1529 |
have *: "\<forall>x\<in>carrier ?G. hom_induced (int p) (nsphere p) {} (nsphere p) {} (\<lambda>x. a) x = \<one>\<^bsub>?H\<^esub>" |
|
1530 |
by (metis False hom_induced_default one_relative_homology_group) |
|
1531 |
obtain c where c: "c \<in> carrier ?G" and ceq: "subgroup_generated ?G {c} = ?G" |
|
1532 |
using cyclic_reduced_homology_group_nsphere [of p p] by (force simp: cyclic_group_def) |
|
1533 |
have [simp]: "ord c = 0" |
|
1534 |
by (simp add: c ceq iso_finite [OF reduced_homology_group_nsphere(1)] flip: infinite_cyclic_subgroup_order) |
|
1535 |
show ?thesis |
|
1536 |
unfolding Brouwer_degree2_def |
|
1537 |
proof (rule some_equality) |
|
1538 |
fix d :: "int" |
|
1539 |
assume "\<forall>x\<in>carrier ?G. hom_induced (int p) (nsphere p) {} (nsphere p) {} (\<lambda>x. a) x = x [^]\<^bsub>?G\<^esub> d" |
|
1540 |
then have "c [^]\<^bsub>?G\<^esub> d = \<one>\<^bsub>?H\<^esub>" |
|
1541 |
using "*" c by blast |
|
1542 |
then have "int (ord c) dvd d" |
|
1543 |
using c eq1 int_pow_eq_id by auto |
|
1544 |
then show "d = 0" |
|
1545 |
by (simp add: * del: one_relative_homology_group) |
|
1546 |
qed (use "*" eq1 in force) |
|
1547 |
qed |
|
1548 |
||
1549 |
||
1550 |
corollary Brouwer_degree2_nonsurjective: |
|
1551 |
"\<lbrakk>continuous_map(nsphere p) (nsphere p) f; f ` topspace (nsphere p) \<noteq> topspace (nsphere p)\<rbrakk> |
|
1552 |
\<Longrightarrow> Brouwer_degree2 p f = 0" |
|
1553 |
by (meson Brouwer_degree2_nullhomotopic nullhomotopic_nonsurjective_sphere_map) |
|
1554 |
||
1555 |
||
1556 |
proposition Brouwer_degree2_reflection: |
|
1557 |
"Brouwer_degree2 p (\<lambda>x i. if i = 0 then -x i else x i) = -1" (is "Brouwer_degree2 _ ?r = -1") |
|
1558 |
proof (induction p) |
|
1559 |
case 0 |
|
1560 |
let ?G = "homology_group 0 (nsphere 0)" |
|
1561 |
let ?D = "homology_group 0 (discrete_topology {()})" |
|
1562 |
interpret group ?G |
|
1563 |
by simp |
|
1564 |
define r where "r \<equiv> \<lambda>x::nat\<Rightarrow>real. \<lambda>i. if i = 0 then -x i else x i" |
|
1565 |
then have [simp]: "r \<circ> r = id" |
|
1566 |
by force |
|
1567 |
have cmr: "continuous_map (nsphere 0) (nsphere 0) r" |
|
1568 |
by (simp add: r_def continuous_map_nsphere_reflection) |
|
1569 |
have *: "hom_induced 0 (nsphere 0) {} (nsphere 0) {} r c = inv\<^bsub>?G\<^esub> c" |
|
1570 |
if "c \<in> carrier(reduced_homology_group 0 (nsphere 0))" for c |
|
1571 |
proof - |
|
1572 |
have c: "c \<in> carrier ?G" |
|
1573 |
and ceq: "hom_induced 0 (nsphere 0) {} (discrete_topology {()}) {} (\<lambda>x. ()) c = \<one>\<^bsub>?D\<^esub>" |
|
1574 |
using that by (auto simp: carrier_reduced_homology_group kernel_def) |
|
1575 |
define pp::"nat\<Rightarrow>real" where "pp \<equiv> \<lambda>i. if i = 0 then 1 else 0" |
|
1576 |
define nn::"nat\<Rightarrow>real" where "nn \<equiv> \<lambda>i. if i = 0 then -1 else 0" |
|
1577 |
have topn0: "topspace(nsphere 0) = {pp,nn}" |
|
1578 |
by (auto simp: nsphere pp_def nn_def fun_eq_iff power2_eq_1_iff split: if_split_asm) |
|
1579 |
have "t1_space (nsphere 0)" |
|
1580 |
unfolding nsphere |
|
1581 |
apply (rule t1_space_subtopology) |
|
1582 |
by (metis (full_types) open_fun_def t1_space t1_space_def) |
|
1583 |
then have dtn0: "discrete_topology {pp,nn} = nsphere 0" |
|
1584 |
using finite_t1_space_imp_discrete_topology [OF topn0] by auto |
|
1585 |
have "pp \<noteq> nn" |
|
1586 |
by (auto simp: pp_def nn_def fun_eq_iff) |
|
1587 |
have [simp]: "r pp = nn" "r nn = pp" |
|
1588 |
by (auto simp: r_def pp_def nn_def fun_eq_iff) |
|
1589 |
have iso: "(\<lambda>(a,b). hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} id a |
|
1590 |
\<otimes>\<^bsub>?G\<^esub> hom_induced 0 (subtopology (nsphere 0) {nn}) {} (nsphere 0) {} id b) |
|
1591 |
\<in> iso (homology_group 0 (subtopology (nsphere 0) {pp}) \<times>\<times> homology_group 0 (subtopology (nsphere 0) {nn})) |
|
1592 |
?G" (is "?f \<in> iso (?P \<times>\<times> ?N) ?G") |
|
1593 |
apply (rule homology_additivity_explicit) |
|
1594 |
using dtn0 \<open>pp \<noteq> nn\<close> by (auto simp: discrete_topology_unique) |
|
1595 |
then have fim: "?f ` carrier(?P \<times>\<times> ?N) = carrier ?G" |
|
1596 |
by (simp add: iso_def bij_betw_def) |
|
1597 |
obtain d d' where d: "d \<in> carrier ?P" and d': "d' \<in> carrier ?N" and eqc: "?f(d,d') = c" |
|
1598 |
using c by (force simp flip: fim) |
|
1599 |
let ?h = "\<lambda>xx. hom_induced 0 (subtopology (nsphere 0) {xx}) {} (discrete_topology {()}) {} (\<lambda>x. ())" |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1600 |
have "continuous_map (subtopology (nsphere 0) {nn}) (nsphere 0) r" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1601 |
using cmr continuous_map_from_subtopology by blast |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1602 |
then have "retraction_map (subtopology (nsphere 0) {pp}) (subtopology (nsphere 0) {nn}) r" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1603 |
apply (simp add: retraction_map_def retraction_maps_def continuous_map_in_subtopology) |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1604 |
using \<open>r nn = pp\<close> \<open>r pp = nn\<close> cmr continuous_map_from_subtopology |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
79712
diff
changeset
|
1605 |
by blast |
70095 | 1606 |
then have "carrier ?N = (hom_induced 0 (subtopology (nsphere 0) {pp}) {} (subtopology (nsphere 0) {nn}) {} r) ` carrier ?P" |
1607 |
by (rule surj_hom_induced_retraction_map) |
|
1608 |
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'" |
|
1609 |
using d' by auto |
|
1610 |
have "section_map (subtopology (nsphere 0) {pp}) (discrete_topology {()}) (\<lambda>x. ())" |
|
1611 |
by (force simp: section_map_def retraction_maps_def topn0) |
|
1612 |
then have "?h pp \<in> mon ?P ?D" |
|
1613 |
by (rule mon_hom_induced_section_map) |
|
1614 |
then have one: "x = one ?P" |
|
1615 |
if "?h pp x = \<one>\<^bsub>?D\<^esub>" "x \<in> carrier ?P" for x |
|
1616 |
using that by (simp add: mon_iff_hom_one) |
|
1617 |
interpret hpd: group_hom ?P ?D "?h pp" |
|
1618 |
using hom_induced_empty_hom by (simp add: hom_induced_empty_hom group_hom_axioms_def group_hom_def) |
|
1619 |
interpret hgd: group_hom ?G ?D "hom_induced 0 (nsphere 0) {} (discrete_topology {()}) {} (\<lambda>x. ())" |
|
1620 |
using hom_induced_empty_hom by (simp add: hom_induced_empty_hom group_hom_axioms_def group_hom_def) |
|
1621 |
interpret hpg: group_hom ?P ?G "hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} r" |
|
1622 |
using hom_induced_empty_hom by (simp add: hom_induced_empty_hom group_hom_axioms_def group_hom_def) |
|
1623 |
interpret hgg: group_hom ?G ?G "hom_induced 0 (nsphere 0) {} (nsphere 0) {} r" |
|
1624 |
using hom_induced_empty_hom by (simp add: hom_induced_empty_hom group_hom_axioms_def group_hom_def) |
|
1625 |
have "?h pp d = |
|
1626 |
(hom_induced 0 (nsphere 0) {} (discrete_topology {()}) {} (\<lambda>x. ()) |
|
1627 |
\<circ> hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} id) d" |
|
1628 |
by (simp flip: hom_induced_compose_empty) |
|
1629 |
moreover |
|
1630 |
have "?h pp = ?h nn \<circ> hom_induced 0 (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 flip: hom_induced_compose_empty) |
|
1632 |
then have "?h pp e = |
|
1633 |
(hom_induced 0 (nsphere 0) {} (discrete_topology {()}) {} (\<lambda>x. ()) |
|
1634 |
\<circ> hom_induced 0 (subtopology (nsphere 0) {nn}) {} (nsphere 0) {} id) d'" |
|
1635 |
by (simp flip: hom_induced_compose_empty eqd') |
|
1636 |
ultimately have "?h pp (d \<otimes>\<^bsub>?P\<^esub> e) = hom_induced 0 (nsphere 0) {} (discrete_topology {()}) {} (\<lambda>x. ()) (?f(d,d'))" |
|
1637 |
by (simp add: d e hom_induced_carrier) |
|
1638 |
then have "?h pp (d \<otimes>\<^bsub>?P\<^esub> e) = \<one>\<^bsub>?D\<^esub>" |
|
1639 |
using ceq eqc by simp |
|
1640 |
then have inv_p: "inv\<^bsub>?P\<^esub> d = e" |
|
1641 |
by (metis (no_types, lifting) Group.group_def d e group.inv_equality group.r_inv group_relative_homology_group one monoid.m_closed) |
|
1642 |
have cmr_pn: "continuous_map (subtopology (nsphere 0) {pp}) (subtopology (nsphere 0) {nn}) r" |
|
1643 |
by (simp add: cmr continuous_map_from_subtopology continuous_map_in_subtopology image_subset_iff) |
|
1644 |
then have "hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} (id \<circ> r) = |
|
1645 |
hom_induced 0 (subtopology (nsphere 0) {nn}) {} (nsphere 0) {} id \<circ> |
|
1646 |
hom_induced 0 (subtopology (nsphere 0) {pp}) {} (subtopology (nsphere 0) {nn}) {} r" |
|
1647 |
using hom_induced_compose_empty continuous_map_id_subt by blast |
|
1648 |
then have "inv\<^bsub>?G\<^esub> hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} r d = |
|
1649 |
hom_induced 0 (subtopology (nsphere 0) {nn}) {} (nsphere 0) {} id d'" |
|
1650 |
apply (simp add: flip: inv_p eqd') |
|
1651 |
using d hpg.hom_inv by auto |
|
1652 |
then have c: "c = (hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} id d) |
|
1653 |
\<otimes>\<^bsub>?G\<^esub> inv\<^bsub>?G\<^esub> (hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} r d)" |
|
1654 |
by (simp flip: eqc) |
|
1655 |
have "hom_induced 0 (nsphere 0) {} (nsphere 0) {} r \<circ> |
|
1656 |
hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} id = |
|
1657 |
hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} r" |
|
1658 |
by (metis cmr comp_id continuous_map_id_subt hom_induced_compose_empty) |
|
1659 |
moreover |
|
1660 |
have "hom_induced 0 (nsphere 0) {} (nsphere 0) {} r \<circ> |
|
1661 |
hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} r = |
|
1662 |
hom_induced 0 (subtopology (nsphere 0) {pp}) {} (nsphere 0) {} id" |
|
1663 |
by (metis \<open>r \<circ> r = id\<close> cmr continuous_map_from_subtopology hom_induced_compose_empty) |
|
1664 |
ultimately show ?thesis |
|
1665 |
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) |
|
1666 |
qed |
|
1667 |
show ?case |
|
1668 |
unfolding r_def [symmetric] |
|
1669 |
using Brouwer_degree2_unique [OF cmr] |
|
1670 |
by (auto simp: * group.int_pow_neg group.int_pow_1 reduced_homology_group_def intro!: Brouwer_degree2_unique [OF cmr]) |
|
1671 |
next |
|
1672 |
case (Suc p) |
|
1673 |
let ?G = "reduced_homology_group (int p) (nsphere p)" |
|
1674 |
let ?G1 = "reduced_homology_group (1 + int p) (nsphere (Suc p))" |
|
1675 |
obtain f g where fg: "group_isomorphisms ?G ?G1 f g" |
|
1676 |
and *: "\<forall>c\<in>carrier ?G. |
|
1677 |
hom_induced (1 + int p) (nsphere (Suc p)) {} (nsphere (Suc p)) {} ?r (f c) = |
|
1678 |
f (hom_induced p (nsphere p) {} (nsphere p) {} ?r c)" |
|
1679 |
using reduced_homology_group_nsphere_step |
|
1680 |
by (meson group.iso_iff_group_isomorphisms group_reduced_homology_group) |
|
1681 |
then have eq: "carrier ?G1 = f ` carrier ?G" |
|
1682 |
by (fastforce simp add: iso_iff dest: group_isomorphisms_imp_iso) |
|
1683 |
interpret group_hom ?G ?G1 f |
|
1684 |
by (meson fg group_hom_axioms_def group_hom_def group_isomorphisms_def group_reduced_homology_group) |
|
1685 |
have homf: "f \<in> hom ?G ?G1" |
|
1686 |
using fg group_isomorphisms_def by blast |
|
1687 |
have "hom_induced (1 + int p) (nsphere (Suc p)) {} (nsphere (Suc p)) {} ?r (f y) = f y [^]\<^bsub>?G1\<^esub> (-1::int)" |
|
1688 |
if "y \<in> carrier ?G" for y |
|
1689 |
by (simp add: that * Brouwer_degree2 Suc hom_int_pow) |
|
1690 |
then show ?case |
|
1691 |
by (fastforce simp: eq intro: Brouwer_degree2_unique [OF continuous_map_nsphere_reflection]) |
|
1692 |
qed |
|
1693 |
||
1694 |
end |