author | paulson <lp15@cam.ac.uk> |
Thu, 17 Apr 2025 22:57:26 +0100 | |
changeset 82522 | 62afd98e3f3e |
parent 82323 | b022c013b04b |
child 82538 | 4b132ea7d575 |
permissions | -rw-r--r-- |
69620 | 1 |
(* Title: HOL/Analysis/Path_Connected.thy |
2 |
Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light |
|
3 |
*) |
|
4 |
||
5 |
section \<open>Homotopy of Maps\<close> |
|
6 |
||
7 |
theory Homotopy |
|
77200
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
76874
diff
changeset
|
8 |
imports Path_Connected Product_Topology Uncountable_Sets |
69620 | 9 |
begin |
10 |
||
70136 | 11 |
definition\<^marker>\<open>tag important\<close> homotopic_with |
69620 | 12 |
where |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
13 |
"homotopic_with P X Y f g \<equiv> |
71745 | 14 |
(\<exists>h. continuous_map (prod_topology (top_of_set {0..1::real}) X) Y h \<and> |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
15 |
(\<forall>x. h(0, x) = f x) \<and> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
16 |
(\<forall>x. h(1, x) = g x) \<and> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
17 |
(\<forall>t \<in> {0..1}. P(\<lambda>x. h(t,x))))" |
69620 | 18 |
|
19 |
text\<open>\<open>p\<close>, \<open>q\<close> are functions \<open>X \<rightarrow> Y\<close>, and the property \<open>P\<close> restricts all intermediate maps. |
|
20 |
We often just want to require that \<open>P\<close> fixes some subset, but to include the case of a loop homotopy, |
|
21 |
it is convenient to have a general property \<open>P\<close>.\<close> |
|
22 |
||
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
23 |
abbreviation homotopic_with_canon :: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
24 |
"[('a::topological_space \<Rightarrow> 'b::topological_space) \<Rightarrow> bool, 'a set, 'b set, 'a \<Rightarrow> 'b, 'a \<Rightarrow> 'b] \<Rightarrow> bool" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
25 |
where |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
26 |
"homotopic_with_canon P S T p q \<equiv> homotopic_with P (top_of_set S) (top_of_set T) p q" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
27 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
28 |
lemma split_01: "{0..1::real} = {0..1/2} \<union> {1/2..1}" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
29 |
by force |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
30 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
31 |
lemma split_01_prod: "{0..1::real} \<times> X = ({0..1/2} \<times> X) \<union> ({1/2..1} \<times> X)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
32 |
by force |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
33 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
34 |
lemma image_Pair_const: "(\<lambda>x. (x, c)) ` A = A \<times> {c}" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
35 |
by auto |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
36 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
37 |
lemma fst_o_paired [simp]: "fst \<circ> (\<lambda>(x,y). (f x y, g x y)) = (\<lambda>(x,y). f x y)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
38 |
by auto |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
39 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
40 |
lemma snd_o_paired [simp]: "snd \<circ> (\<lambda>(x,y). (f x y, g x y)) = (\<lambda>(x,y). g x y)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
41 |
by auto |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
42 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
43 |
lemma continuous_on_o_Pair: "\<lbrakk>continuous_on (T \<times> X) h; t \<in> T\<rbrakk> \<Longrightarrow> continuous_on X (h \<circ> Pair t)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
44 |
by (fast intro: continuous_intros elim!: continuous_on_subset) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
45 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
46 |
lemma continuous_map_o_Pair: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
47 |
assumes h: "continuous_map (prod_topology X Y) Z h" and t: "t \<in> topspace X" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
48 |
shows "continuous_map Y Z (h \<circ> Pair t)" |
71745 | 49 |
by (intro continuous_map_compose [OF _ h] continuous_intros; simp add: t) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
50 |
|
70136 | 51 |
subsection\<^marker>\<open>tag unimportant\<close>\<open>Trivial properties\<close> |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
52 |
|
69620 | 53 |
text \<open>We often want to just localize the ending function equality or whatever.\<close> |
70136 | 54 |
text\<^marker>\<open>tag important\<close> \<open>%whitespace\<close> |
69620 | 55 |
proposition homotopic_with: |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
56 |
assumes "\<And>h k. (\<And>x. x \<in> topspace X \<Longrightarrow> h x = k x) \<Longrightarrow> (P h \<longleftrightarrow> P k)" |
69620 | 57 |
shows "homotopic_with P X Y p q \<longleftrightarrow> |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
58 |
(\<exists>h. continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h \<and> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
59 |
(\<forall>x \<in> topspace X. h(0,x) = p x) \<and> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
60 |
(\<forall>x \<in> topspace X. h(1,x) = q x) \<and> |
69620 | 61 |
(\<forall>t \<in> {0..1}. P(\<lambda>x. h(t, x))))" |
62 |
unfolding homotopic_with_def |
|
63 |
apply (rule iffI, blast, clarify) |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
64 |
apply (rule_tac x="\<lambda>(u,v). if v \<in> topspace X then h(u,v) else if u = 0 then p v else q v" in exI) |
77684 | 65 |
apply simp |
66 |
by (smt (verit, best) SigmaE assms case_prod_conv continuous_map_eq topspace_prod_topology) |
|
69620 | 67 |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
68 |
lemma homotopic_with_mono: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
69 |
assumes hom: "homotopic_with P X Y f g" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
70 |
and Q: "\<And>h. \<lbrakk>continuous_map X Y h; P h\<rbrakk> \<Longrightarrow> Q h" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
71 |
shows "homotopic_with Q X Y f g" |
71745 | 72 |
using hom unfolding homotopic_with_def |
73 |
by (force simp: o_def dest: continuous_map_o_Pair intro: Q) |
|
69620 | 74 |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
75 |
lemma homotopic_with_imp_continuous_maps: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
76 |
assumes "homotopic_with P X Y f g" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
77 |
shows "continuous_map X Y f \<and> continuous_map X Y g" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
78 |
proof - |
71745 | 79 |
obtain h :: "real \<times> 'a \<Rightarrow> 'b" |
80 |
where conth: "continuous_map (prod_topology (top_of_set {0..1}) X) Y h" |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
81 |
and h: "\<forall>x. h (0, x) = f x" "\<forall>x. h (1, x) = g x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
82 |
using assms by (auto simp: homotopic_with_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
83 |
have *: "t \<in> {0..1} \<Longrightarrow> continuous_map X Y (h \<circ> (\<lambda>x. (t,x)))" for t |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
84 |
by (rule continuous_map_compose [OF _ conth]) (simp add: o_def continuous_map_pairwise) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
85 |
show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
86 |
using h *[of 0] *[of 1] by (simp add: continuous_map_eq) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
87 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
88 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
89 |
lemma homotopic_with_imp_continuous: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
90 |
assumes "homotopic_with_canon P X Y f g" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
91 |
shows "continuous_on X f \<and> continuous_on X g" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
92 |
by (meson assms continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
93 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
94 |
lemma homotopic_with_imp_property: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
95 |
assumes "homotopic_with P X Y f g" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
96 |
shows "P f \<and> P g" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
97 |
proof |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
98 |
obtain h where h: "\<And>x. h(0, x) = f x" "\<And>x. h(1, x) = g x" and P: "\<And>t. t \<in> {0..1::real} \<Longrightarrow> P(\<lambda>x. h(t,x))" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
99 |
using assms by (force simp: homotopic_with_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
100 |
show "P f" "P g" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
101 |
using P [of 0] P [of 1] by (force simp: h)+ |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
102 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
103 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
104 |
lemma homotopic_with_equal: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
105 |
assumes "P f" "P g" and contf: "continuous_map X Y f" and fg: "\<And>x. x \<in> topspace X \<Longrightarrow> f x = g x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
106 |
shows "homotopic_with P X Y f g" |
69620 | 107 |
unfolding homotopic_with_def |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
108 |
proof (intro exI conjI allI ballI) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
109 |
let ?h = "\<lambda>(t::real,x). if t = 1 then g x else f x" |
71745 | 110 |
show "continuous_map (prod_topology (top_of_set {0..1}) X) Y ?h" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
111 |
proof (rule continuous_map_eq) |
71745 | 112 |
show "continuous_map (prod_topology (top_of_set {0..1}) X) Y (f \<circ> snd)" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
113 |
by (simp add: contf continuous_map_of_snd) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
114 |
qed (auto simp: fg) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
115 |
show "P (\<lambda>x. ?h (t, x))" if "t \<in> {0..1}" for t |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
116 |
by (cases "t = 1") (simp_all add: assms) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
117 |
qed auto |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
118 |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
119 |
lemma homotopic_with_imp_funspace1: |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
120 |
"homotopic_with_canon P X Y f g \<Longrightarrow> f \<in> X \<rightarrow> Y" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
121 |
using homotopic_with_imp_continuous_maps by fastforce |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
122 |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
123 |
lemma homotopic_with_imp_subset1: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
124 |
"homotopic_with_canon P X Y f g \<Longrightarrow> f ` X \<subseteq> Y" |
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
125 |
using homotopic_with_imp_funspace1 by blast |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
126 |
|
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
127 |
lemma homotopic_with_imp_funspace2: |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
128 |
"homotopic_with_canon P X Y f g \<Longrightarrow> g \<in> X \<rightarrow> Y" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
129 |
using homotopic_with_imp_continuous_maps by force |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
130 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
131 |
lemma homotopic_with_imp_subset2: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
132 |
"homotopic_with_canon P X Y f g \<Longrightarrow> g ` X \<subseteq> Y" |
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
133 |
using homotopic_with_imp_funspace2 by blast |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
134 |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
135 |
lemma homotopic_with_subset_left: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
136 |
"\<lbrakk>homotopic_with_canon P X Y f g; Z \<subseteq> X\<rbrakk> \<Longrightarrow> homotopic_with_canon P Z Y f g" |
71745 | 137 |
unfolding homotopic_with_def by (auto elim!: continuous_on_subset ex_forward) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
138 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
139 |
lemma homotopic_with_subset_right: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
140 |
"\<lbrakk>homotopic_with_canon P X Y f g; Y \<subseteq> Z\<rbrakk> \<Longrightarrow> homotopic_with_canon P X Z f g" |
71745 | 141 |
unfolding homotopic_with_def by (auto elim!: continuous_on_subset ex_forward) |
69620 | 142 |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
143 |
subsection\<open>Homotopy with P is an equivalence relation\<close> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
144 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
145 |
text \<open>(on continuous functions mapping X into Y that satisfy P, though this only affects reflexivity)\<close> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
146 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
147 |
lemma homotopic_with_refl [simp]: "homotopic_with P X Y f f \<longleftrightarrow> continuous_map X Y f \<and> P f" |
77684 | 148 |
by (metis homotopic_with_equal homotopic_with_imp_continuous_maps homotopic_with_imp_property) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
149 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
150 |
lemma homotopic_with_symD: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
151 |
assumes "homotopic_with P X Y f g" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
152 |
shows "homotopic_with P X Y g f" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
153 |
proof - |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
154 |
let ?I01 = "subtopology euclideanreal {0..1}" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
155 |
let ?j = "\<lambda>y. (1 - fst y, snd y)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
156 |
have 1: "continuous_map (prod_topology ?I01 X) (prod_topology euclideanreal X) ?j" |
71745 | 157 |
by (intro continuous_intros; simp add: continuous_map_subtopology_fst prod_topology_subtopology) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
158 |
have *: "continuous_map (prod_topology ?I01 X) (prod_topology ?I01 X) ?j" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
159 |
proof - |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
160 |
have "continuous_map (prod_topology ?I01 X) (subtopology (prod_topology euclideanreal X) ({0..1} \<times> topspace X)) ?j" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
161 |
by (simp add: continuous_map_into_subtopology [OF 1] image_subset_iff flip: image_subset_iff_funcset) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
162 |
then show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
163 |
by (simp add: prod_topology_subtopology(1)) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
164 |
qed |
69620 | 165 |
show ?thesis |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
166 |
using assms |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
167 |
apply (clarsimp simp: homotopic_with_def) |
71745 | 168 |
subgoal for h |
169 |
by (rule_tac x="h \<circ> (\<lambda>y. (1 - fst y, snd y))" in exI) (simp add: continuous_map_compose [OF *]) |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
170 |
done |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
171 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
172 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
173 |
lemma homotopic_with_sym: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
174 |
"homotopic_with P X Y f g \<longleftrightarrow> homotopic_with P X Y g f" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
175 |
by (metis homotopic_with_symD) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
176 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
177 |
proposition homotopic_with_trans: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
178 |
assumes "homotopic_with P X Y f g" "homotopic_with P X Y g h" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
179 |
shows "homotopic_with P X Y f h" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
180 |
proof - |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
181 |
let ?X01 = "prod_topology (subtopology euclideanreal {0..1}) X" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
182 |
obtain k1 k2 |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
183 |
where contk1: "continuous_map ?X01 Y k1" and contk2: "continuous_map ?X01 Y k2" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
184 |
and k12: "\<forall>x. k1 (1, x) = g x" "\<forall>x. k2 (0, x) = g x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
185 |
"\<forall>x. k1 (0, x) = f x" "\<forall>x. k2 (1, x) = h x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
186 |
and P: "\<forall>t\<in>{0..1}. P (\<lambda>x. k1 (t, x))" "\<forall>t\<in>{0..1}. P (\<lambda>x. k2 (t, x))" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
187 |
using assms by (auto simp: homotopic_with_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
188 |
define k where "k \<equiv> \<lambda>y. if fst y \<le> 1/2 |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
189 |
then (k1 \<circ> (\<lambda>x. (2 *\<^sub>R fst x, snd x))) y |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
190 |
else (k2 \<circ> (\<lambda>x. (2 *\<^sub>R fst x -1, snd x))) y" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
191 |
have keq: "k1 (2 * u, v) = k2 (2 * u -1, v)" if "u = 1/2" for u v |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
192 |
by (simp add: k12 that) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
193 |
show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
194 |
unfolding homotopic_with_def |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
195 |
proof (intro exI conjI) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
196 |
show "continuous_map ?X01 Y k" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
197 |
unfolding k_def |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
198 |
proof (rule continuous_map_cases_le) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
199 |
show fst: "continuous_map ?X01 euclideanreal fst" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
200 |
using continuous_map_fst continuous_map_in_subtopology by blast |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
201 |
show "continuous_map ?X01 euclideanreal (\<lambda>x. 1/2)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
202 |
by simp |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
203 |
show "continuous_map (subtopology ?X01 {y \<in> topspace ?X01. fst y \<le> 1/2}) Y |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
204 |
(k1 \<circ> (\<lambda>x. (2 *\<^sub>R fst x, snd x)))" |
72372 | 205 |
apply (intro fst continuous_map_compose [OF _ contk1] continuous_intros continuous_map_into_subtopology continuous_map_from_subtopology | simp)+ |
71745 | 206 |
by (force simp: prod_topology_subtopology) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
207 |
show "continuous_map (subtopology ?X01 {y \<in> topspace ?X01. 1/2 \<le> fst y}) Y |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
208 |
(k2 \<circ> (\<lambda>x. (2 *\<^sub>R fst x -1, snd x)))" |
72372 | 209 |
apply (intro fst continuous_map_compose [OF _ contk2] continuous_intros continuous_map_into_subtopology continuous_map_from_subtopology | simp)+ |
71745 | 210 |
by (force simp: prod_topology_subtopology) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
211 |
show "(k1 \<circ> (\<lambda>x. (2 *\<^sub>R fst x, snd x))) y = (k2 \<circ> (\<lambda>x. (2 *\<^sub>R fst x -1, snd x))) y" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
212 |
if "y \<in> topspace ?X01" and "fst y = 1/2" for y |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
213 |
using that by (simp add: keq) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
214 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
215 |
show "\<forall>x. k (0, x) = f x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
216 |
by (simp add: k12 k_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
217 |
show "\<forall>x. k (1, x) = h x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
218 |
by (simp add: k12 k_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
219 |
show "\<forall>t\<in>{0..1}. P (\<lambda>x. k (t, x))" |
71745 | 220 |
proof |
221 |
fix t show "t\<in>{0..1} \<Longrightarrow> P (\<lambda>x. k (t, x))" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
222 |
by (cases "t \<le> 1/2") (auto simp: k_def P) |
71745 | 223 |
qed |
69620 | 224 |
qed |
225 |
qed |
|
226 |
||
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
227 |
lemma homotopic_with_id2: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
228 |
"(\<And>x. x \<in> topspace X \<Longrightarrow> g (f x) = x) \<Longrightarrow> homotopic_with (\<lambda>x. True) X X (g \<circ> f) id" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
229 |
by (metis comp_apply continuous_map_id eq_id_iff homotopic_with_equal homotopic_with_symD) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
230 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
231 |
subsection\<open>Continuity lemmas\<close> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
232 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
233 |
lemma homotopic_with_compose_continuous_map_left: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
234 |
"\<lbrakk>homotopic_with p X1 X2 f g; continuous_map X2 X3 h; \<And>j. p j \<Longrightarrow> q(h \<circ> j)\<rbrakk> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
235 |
\<Longrightarrow> homotopic_with q X1 X3 (h \<circ> f) (h \<circ> g)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
236 |
unfolding homotopic_with_def |
69620 | 237 |
apply clarify |
71745 | 238 |
subgoal for k |
239 |
by (rule_tac x="h \<circ> k" in exI) (rule conjI continuous_map_compose | simp add: o_def)+ |
|
69620 | 240 |
done |
241 |
||
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
242 |
lemma homotopic_with_compose_continuous_map_right: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
243 |
assumes hom: "homotopic_with p X2 X3 f g" and conth: "continuous_map X1 X2 h" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
244 |
and q: "\<And>j. p j \<Longrightarrow> q(j \<circ> h)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
245 |
shows "homotopic_with q X1 X3 (f \<circ> h) (g \<circ> h)" |
69620 | 246 |
proof - |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
247 |
obtain k |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
248 |
where contk: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X2) X3 k" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
249 |
and k: "\<forall>x. k (0, x) = f x" "\<forall>x. k (1, x) = g x" and p: "\<And>t. t\<in>{0..1} \<Longrightarrow> p (\<lambda>x. k (t, x))" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
250 |
using hom unfolding homotopic_with_def by blast |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
251 |
have hsnd: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) X2 (h \<circ> snd)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
252 |
by (rule continuous_map_compose [OF continuous_map_snd conth]) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
253 |
let ?h = "k \<circ> (\<lambda>(t,x). (t,h x))" |
69620 | 254 |
show ?thesis |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
255 |
unfolding homotopic_with_def |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
256 |
proof (intro exI conjI allI ballI) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
257 |
have "continuous_map (prod_topology (top_of_set {0..1}) X1) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
258 |
(prod_topology (top_of_set {0..1::real}) X2) (\<lambda>(t, x). (t, h x))" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
259 |
by (metis (mono_tags, lifting) case_prod_beta' comp_def continuous_map_eq continuous_map_fst continuous_map_pairedI hsnd) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
260 |
then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) X3 ?h" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
261 |
by (intro conjI continuous_map_compose [OF _ contk]) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
262 |
show "q (\<lambda>x. ?h (t, x))" if "t \<in> {0..1}" for t |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
263 |
using q [OF p [OF that]] by (simp add: o_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
264 |
qed (auto simp: k) |
69620 | 265 |
qed |
266 |
||
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
267 |
corollary homotopic_compose: |
71745 | 268 |
assumes "homotopic_with (\<lambda>x. True) X Y f f'" "homotopic_with (\<lambda>x. True) Y Z g g'" |
269 |
shows "homotopic_with (\<lambda>x. True) X Z (g \<circ> f) (g' \<circ> f')" |
|
77684 | 270 |
by (metis assms homotopic_with_compose_continuous_map_left homotopic_with_compose_continuous_map_right homotopic_with_imp_continuous_maps homotopic_with_trans) |
69620 | 271 |
|
272 |
proposition homotopic_with_compose_continuous_right: |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
273 |
"\<lbrakk>homotopic_with_canon (\<lambda>f. p (f \<circ> h)) X Y f g; continuous_on W h; h \<in> W \<rightarrow> X\<rbrakk> |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
274 |
\<Longrightarrow> homotopic_with_canon p W Y (f \<circ> h) (g \<circ> h)" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
275 |
by (simp add: homotopic_with_compose_continuous_map_right image_subset_iff_funcset) |
69620 | 276 |
|
277 |
proposition homotopic_with_compose_continuous_left: |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
278 |
"\<lbrakk>homotopic_with_canon (\<lambda>f. p (h \<circ> f)) X Y f g; continuous_on Y h; h \<in> Y \<rightarrow> Z\<rbrakk> |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
279 |
\<Longrightarrow> homotopic_with_canon p X Z (h \<circ> f) (h \<circ> g)" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
280 |
by (simp add: homotopic_with_compose_continuous_map_left image_subset_iff_funcset) |
69620 | 281 |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
282 |
lemma homotopic_from_subtopology: |
77684 | 283 |
"homotopic_with P X X' f g \<Longrightarrow> homotopic_with P (subtopology X S) X' f g" |
284 |
by (metis continuous_map_id_subt homotopic_with_compose_continuous_map_right o_id) |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
285 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
286 |
lemma homotopic_on_emptyI: |
78336 | 287 |
assumes "P f" "P g" |
288 |
shows "homotopic_with P trivial_topology X f g" |
|
289 |
by (metis assms continuous_map_on_empty empty_iff homotopic_with_equal topspace_discrete_topology) |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
290 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
291 |
lemma homotopic_on_empty: |
78336 | 292 |
"(homotopic_with P trivial_topology X f g \<longleftrightarrow> P f \<and> P g)" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
293 |
using homotopic_on_emptyI homotopic_with_imp_property by metis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
294 |
|
78336 | 295 |
lemma homotopic_with_canon_on_empty: "homotopic_with_canon (\<lambda>x. True) {} t f g" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
296 |
by (auto intro: homotopic_with_equal) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
297 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
298 |
lemma homotopic_constant_maps: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
299 |
"homotopic_with (\<lambda>x. True) X X' (\<lambda>x. a) (\<lambda>x. b) \<longleftrightarrow> |
78336 | 300 |
X = trivial_topology \<or> path_component_of X' a b" (is "?lhs = ?rhs") |
301 |
proof (cases "X = trivial_topology") |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
302 |
case False |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
303 |
then obtain c where c: "c \<in> topspace X" |
78336 | 304 |
by fastforce |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
305 |
have "\<exists>g. continuous_map (top_of_set {0..1::real}) X' g \<and> g 0 = a \<and> g 1 = b" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
306 |
if "x \<in> topspace X" and hom: "homotopic_with (\<lambda>x. True) X X' (\<lambda>x. a) (\<lambda>x. b)" for x |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
307 |
proof - |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
308 |
obtain h :: "real \<times> 'a \<Rightarrow> 'b" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
309 |
where conth: "continuous_map (prod_topology (top_of_set {0..1}) X) X' h" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
310 |
and h: "\<And>x. h (0, x) = a" "\<And>x. h (1, x) = b" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
311 |
using hom by (auto simp: homotopic_with_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
312 |
have cont: "continuous_map (top_of_set {0..1}) X' (h \<circ> (\<lambda>t. (t, c)))" |
78336 | 313 |
by (rule continuous_map_compose [OF _ conth] continuous_intros | simp add: c)+ |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
314 |
then show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
315 |
by (force simp: h) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
316 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
317 |
moreover have "homotopic_with (\<lambda>x. True) X X' (\<lambda>x. g 0) (\<lambda>x. g 1)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
318 |
if "x \<in> topspace X" "a = g 0" "b = g 1" "continuous_map (top_of_set {0..1}) X' g" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
319 |
for x and g :: "real \<Rightarrow> 'b" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
320 |
unfolding homotopic_with_def |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
321 |
by (force intro!: continuous_map_compose continuous_intros c that) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
322 |
ultimately show ?thesis |
78336 | 323 |
using False |
324 |
by (metis c path_component_of_set pathin_def) |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
325 |
qed (simp add: homotopic_on_empty) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
326 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
327 |
proposition homotopic_with_eq: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
328 |
assumes h: "homotopic_with P X Y f g" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
329 |
and f': "\<And>x. x \<in> topspace X \<Longrightarrow> f' x = f x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
330 |
and g': "\<And>x. x \<in> topspace X \<Longrightarrow> g' x = g x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
331 |
and P: "(\<And>h k. (\<And>x. x \<in> topspace X \<Longrightarrow> h x = k x) \<Longrightarrow> P h \<longleftrightarrow> P k)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
332 |
shows "homotopic_with P X Y f' g'" |
77684 | 333 |
by (smt (verit, ccfv_SIG) assms homotopic_with) |
69620 | 334 |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
335 |
lemma homotopic_with_prod_topology: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
336 |
assumes "homotopic_with p X1 Y1 f f'" and "homotopic_with q X2 Y2 g g'" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
337 |
and r: "\<And>i j. \<lbrakk>p i; q j\<rbrakk> \<Longrightarrow> r(\<lambda>(x,y). (i x, j y))" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
338 |
shows "homotopic_with r (prod_topology X1 X2) (prod_topology Y1 Y2) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
339 |
(\<lambda>z. (f(fst z),g(snd z))) (\<lambda>z. (f'(fst z), g'(snd z)))" |
69620 | 340 |
proof - |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
341 |
obtain h |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
342 |
where h: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) Y1 h" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
343 |
and h0: "\<And>x. h (0, x) = f x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
344 |
and h1: "\<And>x. h (1, x) = f' x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
345 |
and p: "\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> p (\<lambda>x. h (t,x))" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
346 |
using assms unfolding homotopic_with_def by auto |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
347 |
obtain k |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
348 |
where k: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X2) Y2 k" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
349 |
and k0: "\<And>x. k (0, x) = g x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
350 |
and k1: "\<And>x. k (1, x) = g' x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
351 |
and q: "\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> q (\<lambda>x. k (t,x))" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
352 |
using assms unfolding homotopic_with_def by auto |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
353 |
let ?hk = "\<lambda>(t,x,y). (h(t,x), k(t,y))" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
354 |
show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
355 |
unfolding homotopic_with_def |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
356 |
proof (intro conjI allI exI) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
357 |
show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (prod_topology X1 X2)) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
358 |
(prod_topology Y1 Y2) ?hk" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
359 |
unfolding continuous_map_pairwise case_prod_unfold |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
360 |
by (rule conjI continuous_map_pairedI continuous_intros continuous_map_id [unfolded id_def] |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
361 |
continuous_map_fst_of [unfolded o_def] continuous_map_snd_of [unfolded o_def] |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
362 |
continuous_map_compose [OF _ h, unfolded o_def] |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
363 |
continuous_map_compose [OF _ k, unfolded o_def])+ |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
364 |
next |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
365 |
fix x |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
366 |
show "?hk (0, x) = (f (fst x), g (snd x))" "?hk (1, x) = (f' (fst x), g' (snd x))" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
367 |
by (auto simp: case_prod_beta h0 k0 h1 k1) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
368 |
qed (auto simp: p q r) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
369 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
370 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
371 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
372 |
lemma homotopic_with_product_topology: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
373 |
assumes ht: "\<And>i. i \<in> I \<Longrightarrow> homotopic_with (p i) (X i) (Y i) (f i) (g i)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
374 |
and pq: "\<And>h. (\<And>i. i \<in> I \<Longrightarrow> p i (h i)) \<Longrightarrow> q(\<lambda>x. (\<lambda>i\<in>I. h i (x i)))" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
375 |
shows "homotopic_with q (product_topology X I) (product_topology Y I) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
376 |
(\<lambda>z. (\<lambda>i\<in>I. (f i) (z i))) (\<lambda>z. (\<lambda>i\<in>I. (g i) (z i)))" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
377 |
proof - |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
378 |
obtain h |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
379 |
where h: "\<And>i. i \<in> I \<Longrightarrow> continuous_map (prod_topology (subtopology euclideanreal {0..1}) (X i)) (Y i) (h i)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
380 |
and h0: "\<And>i x. i \<in> I \<Longrightarrow> h i (0, x) = f i x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
381 |
and h1: "\<And>i x. i \<in> I \<Longrightarrow> h i (1, x) = g i x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
382 |
and p: "\<And>i t. \<lbrakk>i \<in> I; t \<in> {0..1}\<rbrakk> \<Longrightarrow> p i (\<lambda>x. h i (t,x))" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
383 |
using ht unfolding homotopic_with_def by metis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
384 |
show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
385 |
unfolding homotopic_with_def |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
386 |
proof (intro conjI allI exI) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
387 |
let ?h = "\<lambda>(t,z). \<lambda>i\<in>I. h i (t,z i)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
388 |
have "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I)) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
389 |
(Y i) (\<lambda>x. h i (fst x, snd x i))" if "i \<in> I" for i |
71745 | 390 |
proof - |
391 |
have \<section>: "continuous_map (prod_topology (top_of_set {0..1}) (product_topology X I)) (X i) (\<lambda>x. snd x i)" |
|
392 |
using continuous_map_componentwise continuous_map_snd that by fastforce |
|
393 |
show ?thesis |
|
394 |
unfolding continuous_map_pairwise case_prod_unfold |
|
395 |
by (intro conjI that \<section> continuous_intros continuous_map_compose [OF _ h, unfolded o_def]) |
|
396 |
qed |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
397 |
then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I)) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
398 |
(product_topology Y I) ?h" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
399 |
by (auto simp: continuous_map_componentwise case_prod_beta) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
400 |
show "?h (0, x) = (\<lambda>i\<in>I. f i (x i))" "?h (1, x) = (\<lambda>i\<in>I. g i (x i))" for x |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
401 |
by (auto simp: case_prod_beta h0 h1) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
402 |
show "\<forall>t\<in>{0..1}. q (\<lambda>x. ?h (t, x))" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
403 |
by (force intro: p pq) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
404 |
qed |
69620 | 405 |
qed |
406 |
||
407 |
text\<open>Homotopic triviality implicitly incorporates path-connectedness.\<close> |
|
408 |
lemma homotopic_triviality: |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
409 |
shows "(\<forall>f g. continuous_on S f \<and> f \<in> S \<rightarrow> T \<and> |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
410 |
continuous_on S g \<and> g \<in> S \<rightarrow> T |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
411 |
\<longrightarrow> homotopic_with_canon (\<lambda>x. True) S T f g) \<longleftrightarrow> |
69620 | 412 |
(S = {} \<or> path_connected T) \<and> |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
413 |
(\<forall>f. continuous_on S f \<and> f \<in> S \<rightarrow> T \<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) S T f (\<lambda>x. c)))" |
69620 | 414 |
(is "?lhs = ?rhs") |
415 |
proof (cases "S = {} \<or> T = {}") |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
416 |
case True then show ?thesis |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
417 |
by (auto simp: homotopic_on_emptyI simp flip: image_subset_iff_funcset) |
69620 | 418 |
next |
419 |
case False show ?thesis |
|
420 |
proof |
|
421 |
assume LHS [rule_format]: ?lhs |
|
422 |
have pab: "path_component T a b" if "a \<in> T" "b \<in> T" for a b |
|
423 |
proof - |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
424 |
have "homotopic_with_canon (\<lambda>x. True) S T (\<lambda>x. a) (\<lambda>x. b)" |
71172 | 425 |
by (simp add: LHS image_subset_iff that) |
69620 | 426 |
then show ?thesis |
78336 | 427 |
using False homotopic_constant_maps [of "top_of_set S" "top_of_set T" a b] |
428 |
by (metis path_component_of_canon_iff topspace_discrete_topology topspace_euclidean_subtopology) |
|
69620 | 429 |
qed |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
430 |
moreover |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
431 |
have "\<exists>c. homotopic_with_canon (\<lambda>x. True) S T f (\<lambda>x. c)" if "continuous_on S f" "f \<in> S \<rightarrow> T" for f |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
432 |
using False LHS continuous_on_const that by blast |
69620 | 433 |
ultimately show ?rhs |
434 |
by (simp add: path_connected_component) |
|
435 |
next |
|
436 |
assume RHS: ?rhs |
|
437 |
with False have T: "path_connected T" |
|
438 |
by blast |
|
439 |
show ?lhs |
|
440 |
proof clarify |
|
441 |
fix f g |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
442 |
assume "continuous_on S f" "f \<in> S \<rightarrow> T" "continuous_on S g" "g \<in> S \<rightarrow> T" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
443 |
obtain c d where c: "homotopic_with_canon (\<lambda>x. True) S T f (\<lambda>x. c)" and d: "homotopic_with_canon (\<lambda>x. True) S T g (\<lambda>x. d)" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
444 |
using RHS \<open>continuous_on S f\<close> \<open>continuous_on S g\<close> \<open>f \<in> S \<rightarrow> T\<close> \<open>g \<in> S \<rightarrow> T\<close> by presburger |
69620 | 445 |
with T have "path_component T c d" |
77684 | 446 |
by (metis False ex_in_conv homotopic_with_imp_subset2 image_subset_iff path_connected_component) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
447 |
then have "homotopic_with_canon (\<lambda>x. True) S T (\<lambda>x. c) (\<lambda>x. d)" |
69620 | 448 |
by (simp add: homotopic_constant_maps) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
449 |
with c d show "homotopic_with_canon (\<lambda>x. True) S T f g" |
69620 | 450 |
by (meson homotopic_with_symD homotopic_with_trans) |
451 |
qed |
|
452 |
qed |
|
453 |
qed |
|
454 |
||
455 |
||
456 |
subsection\<open>Homotopy of paths, maintaining the same endpoints\<close> |
|
457 |
||
458 |
||
70136 | 459 |
definition\<^marker>\<open>tag important\<close> homotopic_paths :: "['a set, real \<Rightarrow> 'a, real \<Rightarrow> 'a::topological_space] \<Rightarrow> bool" |
69620 | 460 |
where |
77684 | 461 |
"homotopic_paths S p q \<equiv> |
462 |
homotopic_with_canon (\<lambda>r. pathstart r = pathstart p \<and> pathfinish r = pathfinish p) {0..1} S p q" |
|
69620 | 463 |
|
464 |
lemma homotopic_paths: |
|
77684 | 465 |
"homotopic_paths S p q \<longleftrightarrow> |
69620 | 466 |
(\<exists>h. continuous_on ({0..1} \<times> {0..1}) h \<and> |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
467 |
h \<in> ({0..1} \<times> {0..1}) \<rightarrow> S \<and> |
69620 | 468 |
(\<forall>x \<in> {0..1}. h(0,x) = p x) \<and> |
469 |
(\<forall>x \<in> {0..1}. h(1,x) = q x) \<and> |
|
470 |
(\<forall>t \<in> {0..1::real}. pathstart(h \<circ> Pair t) = pathstart p \<and> |
|
471 |
pathfinish(h \<circ> Pair t) = pathfinish p))" |
|
472 |
by (auto simp: homotopic_paths_def homotopic_with pathstart_def pathfinish_def) |
|
473 |
||
474 |
proposition homotopic_paths_imp_pathstart: |
|
77684 | 475 |
"homotopic_paths S p q \<Longrightarrow> pathstart p = pathstart q" |
69620 | 476 |
by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) |
477 |
||
478 |
proposition homotopic_paths_imp_pathfinish: |
|
77684 | 479 |
"homotopic_paths S p q \<Longrightarrow> pathfinish p = pathfinish q" |
69620 | 480 |
by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) |
481 |
||
482 |
lemma homotopic_paths_imp_path: |
|
77684 | 483 |
"homotopic_paths S p q \<Longrightarrow> path p \<and> path q" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
484 |
using homotopic_paths_def homotopic_with_imp_continuous_maps path_def continuous_map_subtopology_eu by blast |
69620 | 485 |
|
486 |
lemma homotopic_paths_imp_subset: |
|
77684 | 487 |
"homotopic_paths S p q \<Longrightarrow> path_image p \<subseteq> S \<and> path_image q \<subseteq> S" |
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
488 |
by (simp add: homotopic_paths_def homotopic_with_imp_subset1 homotopic_with_imp_subset2 |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
489 |
path_image_def) |
69620 | 490 |
|
77684 | 491 |
proposition homotopic_paths_refl [simp]: "homotopic_paths S p p \<longleftrightarrow> path p \<and> path_image p \<subseteq> S" |
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
492 |
by (auto simp add: homotopic_paths_def path_def path_image_def) |
69620 | 493 |
|
77684 | 494 |
proposition homotopic_paths_sym: "homotopic_paths S p q \<Longrightarrow> homotopic_paths S q p" |
69620 | 495 |
by (metis (mono_tags) homotopic_paths_def homotopic_paths_imp_pathfinish homotopic_paths_imp_pathstart homotopic_with_symD) |
496 |
||
77684 | 497 |
proposition homotopic_paths_sym_eq: "homotopic_paths S p q \<longleftrightarrow> homotopic_paths S q p" |
69620 | 498 |
by (metis homotopic_paths_sym) |
499 |
||
500 |
proposition homotopic_paths_trans [trans]: |
|
77684 | 501 |
assumes "homotopic_paths S p q" "homotopic_paths S q r" |
502 |
shows "homotopic_paths S p r" |
|
503 |
using assms homotopic_paths_imp_pathfinish homotopic_paths_imp_pathstart unfolding homotopic_paths_def |
|
504 |
by (smt (verit, ccfv_SIG) homotopic_with_mono homotopic_with_trans) |
|
69620 | 505 |
|
506 |
proposition homotopic_paths_eq: |
|
77684 | 507 |
"\<lbrakk>path p; path_image p \<subseteq> S; \<And>t. t \<in> {0..1} \<Longrightarrow> p t = q t\<rbrakk> \<Longrightarrow> homotopic_paths S p q" |
508 |
by (smt (verit, best) homotopic_paths homotopic_paths_refl) |
|
69620 | 509 |
|
510 |
proposition homotopic_paths_reparametrize: |
|
511 |
assumes "path p" |
|
77684 | 512 |
and pips: "path_image p \<subseteq> S" |
69620 | 513 |
and contf: "continuous_on {0..1} f" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
514 |
and f01 :"f \<in> {0..1} \<rightarrow> {0..1}" |
69620 | 515 |
and [simp]: "f(0) = 0" "f(1) = 1" |
516 |
and q: "\<And>t. t \<in> {0..1} \<Longrightarrow> q(t) = p(f t)" |
|
77684 | 517 |
shows "homotopic_paths S p q" |
69620 | 518 |
proof - |
519 |
have contp: "continuous_on {0..1} p" |
|
520 |
by (metis \<open>path p\<close> path_def) |
|
521 |
then have "continuous_on {0..1} (p \<circ> f)" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
522 |
by (meson assms(4) contf continuous_on_compose continuous_on_subset image_subset_iff_funcset) |
69620 | 523 |
then have "path q" |
524 |
by (simp add: path_def) (metis q continuous_on_cong) |
|
77684 | 525 |
have piqs: "path_image q \<subseteq> S" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
526 |
by (smt (verit, ccfv_threshold) Pi_iff assms(2) assms(4) assms(7) image_subset_iff path_defs(4)) |
69620 | 527 |
have fb0: "\<And>a b. \<lbrakk>0 \<le> a; a \<le> 1; 0 \<le> b; b \<le> 1\<rbrakk> \<Longrightarrow> 0 \<le> (1 - a) * f b + a * b" |
528 |
using f01 by force |
|
529 |
have fb1: "\<lbrakk>0 \<le> a; a \<le> 1; 0 \<le> b; b \<le> 1\<rbrakk> \<Longrightarrow> (1 - a) * f b + a * b \<le> 1" for a b |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
530 |
by (intro convex_bound_le) (use f01 in auto) |
77684 | 531 |
have "homotopic_paths S q p" |
69620 | 532 |
proof (rule homotopic_paths_trans) |
77684 | 533 |
show "homotopic_paths S q (p \<circ> f)" |
69620 | 534 |
using q by (force intro: homotopic_paths_eq [OF \<open>path q\<close> piqs]) |
535 |
next |
|
77684 | 536 |
show "homotopic_paths S (p \<circ> f) p" |
72372 | 537 |
using pips [unfolded path_image_def] |
69620 | 538 |
apply (simp add: homotopic_paths_def homotopic_with_def) |
539 |
apply (rule_tac x="p \<circ> (\<lambda>y. (1 - (fst y)) *\<^sub>R ((f \<circ> snd) y) + (fst y) *\<^sub>R snd y)" in exI) |
|
540 |
apply (rule conjI contf continuous_intros continuous_on_subset [OF contp] | simp)+ |
|
72372 | 541 |
by (auto simp: fb0 fb1 pathstart_def pathfinish_def) |
69620 | 542 |
qed |
543 |
then show ?thesis |
|
544 |
by (simp add: homotopic_paths_sym) |
|
545 |
qed |
|
546 |
||
77684 | 547 |
lemma homotopic_paths_subset: "\<lbrakk>homotopic_paths S p q; S \<subseteq> t\<rbrakk> \<Longrightarrow> homotopic_paths t p q" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
548 |
unfolding homotopic_paths by fast |
69620 | 549 |
|
550 |
||
551 |
text\<open> A slightly ad-hoc but useful lemma in constructing homotopies.\<close> |
|
71746 | 552 |
lemma continuous_on_homotopic_join_lemma: |
69620 | 553 |
fixes q :: "[real,real] \<Rightarrow> 'a::topological_space" |
71745 | 554 |
assumes p: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. p (fst y) (snd y))" (is "continuous_on ?A ?p") |
555 |
and q: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. q (fst y) (snd y))" (is "continuous_on ?A ?q") |
|
69620 | 556 |
and pf: "\<And>t. t \<in> {0..1} \<Longrightarrow> pathfinish(p t) = pathstart(q t)" |
557 |
shows "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. (p(fst y) +++ q(fst y)) (snd y))" |
|
558 |
proof - |
|
71745 | 559 |
have \<section>: "(\<lambda>t. p (fst t) (2 * snd t)) = ?p \<circ> (\<lambda>y. (fst y, 2 * snd y))" |
560 |
"(\<lambda>t. q (fst t) (2 * snd t - 1)) = ?q \<circ> (\<lambda>y. (fst y, 2 * snd y - 1))" |
|
561 |
by force+ |
|
69620 | 562 |
show ?thesis |
71745 | 563 |
unfolding joinpaths_def |
564 |
proof (rule continuous_on_cases_le) |
|
565 |
show "continuous_on {y \<in> ?A. snd y \<le> 1/2} (\<lambda>t. p (fst t) (2 * snd t))" |
|
566 |
"continuous_on {y \<in> ?A. 1/2 \<le> snd y} (\<lambda>t. q (fst t) (2 * snd t - 1))" |
|
567 |
"continuous_on ?A snd" |
|
568 |
unfolding \<section> |
|
569 |
by (rule continuous_intros continuous_on_subset [OF p] continuous_on_subset [OF q] | force)+ |
|
570 |
qed (use pf in \<open>auto simp: mult.commute pathstart_def pathfinish_def\<close>) |
|
69620 | 571 |
qed |
572 |
||
573 |
text\<open> Congruence properties of homotopy w.r.t. path-combining operations.\<close> |
|
574 |
||
575 |
lemma homotopic_paths_reversepath_D: |
|
77684 | 576 |
assumes "homotopic_paths S p q" |
577 |
shows "homotopic_paths S (reversepath p) (reversepath q)" |
|
69620 | 578 |
using assms |
579 |
apply (simp add: homotopic_paths_def homotopic_with_def, clarify) |
|
580 |
apply (rule_tac x="h \<circ> (\<lambda>x. (fst x, 1 - snd x))" in exI) |
|
581 |
apply (rule conjI continuous_intros)+ |
|
582 |
apply (auto simp: reversepath_def pathstart_def pathfinish_def elim!: continuous_on_subset) |
|
583 |
done |
|
584 |
||
585 |
proposition homotopic_paths_reversepath: |
|
77684 | 586 |
"homotopic_paths S (reversepath p) (reversepath q) \<longleftrightarrow> homotopic_paths S p q" |
69620 | 587 |
using homotopic_paths_reversepath_D by force |
588 |
||
589 |
||
590 |
proposition homotopic_paths_join: |
|
77684 | 591 |
"\<lbrakk>homotopic_paths S p p'; homotopic_paths S q q'; pathfinish p = pathstart q\<rbrakk> \<Longrightarrow> homotopic_paths S (p +++ q) (p' +++ q')" |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
592 |
apply (clarsimp simp: homotopic_paths_def homotopic_with_def) |
69620 | 593 |
apply (rename_tac k1 k2) |
594 |
apply (rule_tac x="(\<lambda>y. ((k1 \<circ> Pair (fst y)) +++ (k2 \<circ> Pair (fst y))) (snd y))" in exI) |
|
71746 | 595 |
apply (intro conjI continuous_intros continuous_on_homotopic_join_lemma; force simp: joinpaths_def pathstart_def pathfinish_def path_image_def) |
69620 | 596 |
done |
597 |
||
598 |
proposition homotopic_paths_continuous_image: |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
599 |
"\<lbrakk>homotopic_paths S f g; continuous_on S h; h \<in> S \<rightarrow> t\<rbrakk> \<Longrightarrow> homotopic_paths t (h \<circ> f) (h \<circ> g)" |
69620 | 600 |
unfolding homotopic_paths_def |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
601 |
by (simp add: homotopic_with_compose_continuous_map_left pathfinish_compose pathstart_compose image_subset_iff_funcset) |
69620 | 602 |
|
603 |
||
604 |
subsection\<open>Group properties for homotopy of paths\<close> |
|
605 |
||
70136 | 606 |
text\<^marker>\<open>tag important\<close>\<open>So taking equivalence classes under homotopy would give the fundamental group\<close> |
69620 | 607 |
|
608 |
proposition homotopic_paths_rid: |
|
77684 | 609 |
assumes "path p" "path_image p \<subseteq> S" |
610 |
shows "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p)) p" |
|
71745 | 611 |
proof - |
612 |
have \<section>: "continuous_on {0..1} (\<lambda>t::real. if t \<le> 1/2 then 2 *\<^sub>R t else 1)" |
|
613 |
unfolding split_01 |
|
614 |
by (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+ |
|
615 |
show ?thesis |
|
72372 | 616 |
apply (rule homotopic_paths_sym) |
617 |
using assms unfolding pathfinish_def joinpaths_def |
|
618 |
by (intro \<section> continuous_on_cases continuous_intros homotopic_paths_reparametrize [where f = "\<lambda>t. if t \<le> 1/2 then 2 *\<^sub>R t else 1"]; force) |
|
71745 | 619 |
qed |
69620 | 620 |
|
621 |
proposition homotopic_paths_lid: |
|
77684 | 622 |
"\<lbrakk>path p; path_image p \<subseteq> S\<rbrakk> \<Longrightarrow> homotopic_paths S (linepath (pathstart p) (pathstart p) +++ p) p" |
623 |
using homotopic_paths_rid [of "reversepath p" S] |
|
69620 | 624 |
by (metis homotopic_paths_reversepath path_image_reversepath path_reversepath pathfinish_linepath |
625 |
pathfinish_reversepath reversepath_joinpaths reversepath_linepath) |
|
626 |
||
80052
35b2143aeec6
An assortment of new material, mostly due to Manuel
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
627 |
lemma homotopic_paths_rid': |
35b2143aeec6
An assortment of new material, mostly due to Manuel
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
628 |
assumes "path p" "path_image p \<subseteq> s" "x = pathfinish p" |
35b2143aeec6
An assortment of new material, mostly due to Manuel
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
629 |
shows "homotopic_paths s (p +++ linepath x x) p" |
35b2143aeec6
An assortment of new material, mostly due to Manuel
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
630 |
using homotopic_paths_rid[of p s] assms by simp |
35b2143aeec6
An assortment of new material, mostly due to Manuel
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
631 |
|
35b2143aeec6
An assortment of new material, mostly due to Manuel
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
632 |
lemma homotopic_paths_lid': |
35b2143aeec6
An assortment of new material, mostly due to Manuel
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
633 |
"\<lbrakk>path p; path_image p \<subseteq> s; x = pathstart p\<rbrakk> \<Longrightarrow> homotopic_paths s (linepath x x +++ p) p" |
35b2143aeec6
An assortment of new material, mostly due to Manuel
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
634 |
using homotopic_paths_lid[of p s] by simp |
35b2143aeec6
An assortment of new material, mostly due to Manuel
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
635 |
|
69620 | 636 |
proposition homotopic_paths_assoc: |
77684 | 637 |
"\<lbrakk>path p; path_image p \<subseteq> S; path q; path_image q \<subseteq> S; path r; path_image r \<subseteq> S; pathfinish p = pathstart q; |
69620 | 638 |
pathfinish q = pathstart r\<rbrakk> |
77684 | 639 |
\<Longrightarrow> homotopic_paths S (p +++ (q +++ r)) ((p +++ q) +++ r)" |
69620 | 640 |
apply (subst homotopic_paths_sym) |
641 |
apply (rule homotopic_paths_reparametrize |
|
71745 | 642 |
[where f = "\<lambda>t. if t \<le> 1/2 then inverse 2 *\<^sub>R t |
69620 | 643 |
else if t \<le> 3 / 4 then t - (1 / 4) |
644 |
else 2 *\<^sub>R t - 1"]) |
|
72372 | 645 |
apply (simp_all del: le_divide_eq_numeral1 add: subset_path_image_join) |
71745 | 646 |
apply (rule continuous_on_cases_1 continuous_intros | auto simp: joinpaths_def)+ |
69620 | 647 |
done |
648 |
||
649 |
proposition homotopic_paths_rinv: |
|
77684 | 650 |
assumes "path p" "path_image p \<subseteq> S" |
651 |
shows "homotopic_paths S (p +++ reversepath p) (linepath (pathstart p) (pathstart p))" |
|
69620 | 652 |
proof - |
71745 | 653 |
have p: "continuous_on {0..1} p" |
654 |
using assms by (auto simp: path_def) |
|
655 |
let ?A = "{0..1} \<times> {0..1}" |
|
656 |
have "continuous_on ?A (\<lambda>x. (subpath 0 (fst x) p +++ reversepath (subpath 0 (fst x) p)) (snd x))" |
|
657 |
unfolding joinpaths_def subpath_def reversepath_def path_def add_0_right diff_0_right |
|
658 |
proof (rule continuous_on_cases_le) |
|
659 |
show "continuous_on {x \<in> ?A. snd x \<le> 1/2} (\<lambda>t. p (fst t * (2 * snd t)))" |
|
660 |
"continuous_on {x \<in> ?A. 1/2 \<le> snd x} (\<lambda>t. p (fst t * (1 - (2 * snd t - 1))))" |
|
661 |
"continuous_on ?A snd" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
662 |
by (intro continuous_on_compose2 [OF p] continuous_intros; auto simp: mult_le_one)+ |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
663 |
qed (auto simp: algebra_simps) |
69620 | 664 |
then show ?thesis |
665 |
using assms |
|
666 |
apply (subst homotopic_paths_sym_eq) |
|
667 |
unfolding homotopic_paths_def homotopic_with_def |
|
668 |
apply (rule_tac x="(\<lambda>y. (subpath 0 (fst y) p +++ reversepath(subpath 0 (fst y) p)) (snd y))" in exI) |
|
71745 | 669 |
apply (force simp: mult_le_one path_defs joinpaths_def subpath_def reversepath_def) |
69620 | 670 |
done |
671 |
qed |
|
672 |
||
673 |
proposition homotopic_paths_linv: |
|
77684 | 674 |
assumes "path p" "path_image p \<subseteq> S" |
675 |
shows "homotopic_paths S (reversepath p +++ p) (linepath (pathfinish p) (pathfinish p))" |
|
676 |
using homotopic_paths_rinv [of "reversepath p" S] assms by simp |
|
69620 | 677 |
|
678 |
||
679 |
subsection\<open>Homotopy of loops without requiring preservation of endpoints\<close> |
|
680 |
||
70136 | 681 |
definition\<^marker>\<open>tag important\<close> homotopic_loops :: "'a::topological_space set \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> bool" where |
77684 | 682 |
"homotopic_loops S p q \<equiv> |
683 |
homotopic_with_canon (\<lambda>r. pathfinish r = pathstart r) {0..1} S p q" |
|
69620 | 684 |
|
685 |
lemma homotopic_loops: |
|
77684 | 686 |
"homotopic_loops S p q \<longleftrightarrow> |
69620 | 687 |
(\<exists>h. continuous_on ({0..1::real} \<times> {0..1}) h \<and> |
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
688 |
h \<in> ({0..1} \<times> {0..1}) \<rightarrow> S \<and> |
69620 | 689 |
(\<forall>x \<in> {0..1}. h(0,x) = p x) \<and> |
690 |
(\<forall>x \<in> {0..1}. h(1,x) = q x) \<and> |
|
691 |
(\<forall>t \<in> {0..1}. pathfinish(h \<circ> Pair t) = pathstart(h \<circ> Pair t)))" |
|
692 |
by (simp add: homotopic_loops_def pathstart_def pathfinish_def homotopic_with) |
|
693 |
||
694 |
proposition homotopic_loops_imp_loop: |
|
77684 | 695 |
"homotopic_loops S p q \<Longrightarrow> pathfinish p = pathstart p \<and> pathfinish q = pathstart q" |
69620 | 696 |
using homotopic_with_imp_property homotopic_loops_def by blast |
697 |
||
698 |
proposition homotopic_loops_imp_path: |
|
77684 | 699 |
"homotopic_loops S p q \<Longrightarrow> path p \<and> path q" |
69620 | 700 |
unfolding homotopic_loops_def path_def |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
701 |
using homotopic_with_imp_continuous_maps continuous_map_subtopology_eu by blast |
69620 | 702 |
|
703 |
proposition homotopic_loops_imp_subset: |
|
77684 | 704 |
"homotopic_loops S p q \<Longrightarrow> path_image p \<subseteq> S \<and> path_image q \<subseteq> S" |
69620 | 705 |
unfolding homotopic_loops_def path_image_def |
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
706 |
by (simp add: homotopic_with_imp_subset1 homotopic_with_imp_subset2) |
69620 | 707 |
|
708 |
proposition homotopic_loops_refl: |
|
77684 | 709 |
"homotopic_loops S p p \<longleftrightarrow> |
710 |
path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p" |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
711 |
by (metis (mono_tags, lifting) homotopic_loops_def homotopic_paths_def |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
712 |
homotopic_paths_refl homotopic_with_refl) |
69620 | 713 |
|
77684 | 714 |
proposition homotopic_loops_sym: "homotopic_loops S p q \<Longrightarrow> homotopic_loops S q p" |
69620 | 715 |
by (simp add: homotopic_loops_def homotopic_with_sym) |
716 |
||
77684 | 717 |
proposition homotopic_loops_sym_eq: "homotopic_loops S p q \<longleftrightarrow> homotopic_loops S q p" |
69620 | 718 |
by (metis homotopic_loops_sym) |
719 |
||
720 |
proposition homotopic_loops_trans: |
|
77684 | 721 |
"\<lbrakk>homotopic_loops S p q; homotopic_loops S q r\<rbrakk> \<Longrightarrow> homotopic_loops S p r" |
69620 | 722 |
unfolding homotopic_loops_def by (blast intro: homotopic_with_trans) |
723 |
||
724 |
proposition homotopic_loops_subset: |
|
77684 | 725 |
"\<lbrakk>homotopic_loops S p q; S \<subseteq> t\<rbrakk> \<Longrightarrow> homotopic_loops t p q" |
78474 | 726 |
by (fastforce simp: homotopic_loops) |
69620 | 727 |
|
728 |
proposition homotopic_loops_eq: |
|
77684 | 729 |
"\<lbrakk>path p; path_image p \<subseteq> S; pathfinish p = pathstart p; \<And>t. t \<in> {0..1} \<Longrightarrow> p(t) = q(t)\<rbrakk> |
730 |
\<Longrightarrow> homotopic_loops S p q" |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
731 |
unfolding homotopic_loops_def path_image_def path_def pathstart_def pathfinish_def image_subset_iff_funcset |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
732 |
using homotopic_with_eq [OF homotopic_with_refl [where f = p, THEN iffD2]] |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
733 |
by fastforce |
69620 | 734 |
|
735 |
proposition homotopic_loops_continuous_image: |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
736 |
"\<lbrakk>homotopic_loops S f g; continuous_on S h; h \<in> S \<rightarrow> t\<rbrakk> \<Longrightarrow> homotopic_loops t (h \<circ> f) (h \<circ> g)" |
69620 | 737 |
unfolding homotopic_loops_def |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
738 |
by (simp add: homotopic_with_compose_continuous_map_left pathfinish_def pathstart_def image_subset_iff_funcset) |
69620 | 739 |
|
740 |
||
741 |
subsection\<open>Relations between the two variants of homotopy\<close> |
|
742 |
||
743 |
proposition homotopic_paths_imp_homotopic_loops: |
|
77684 | 744 |
"\<lbrakk>homotopic_paths S p q; pathfinish p = pathstart p; pathfinish q = pathstart p\<rbrakk> \<Longrightarrow> homotopic_loops S p q" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
745 |
by (auto simp: homotopic_with_def homotopic_paths_def homotopic_loops_def) |
69620 | 746 |
|
747 |
proposition homotopic_loops_imp_homotopic_paths_null: |
|
77684 | 748 |
assumes "homotopic_loops S p (linepath a a)" |
749 |
shows "homotopic_paths S p (linepath (pathstart p) (pathstart p))" |
|
69620 | 750 |
proof - |
751 |
have "path p" by (metis assms homotopic_loops_imp_path) |
|
752 |
have ploop: "pathfinish p = pathstart p" by (metis assms homotopic_loops_imp_loop) |
|
77684 | 753 |
have pip: "path_image p \<subseteq> S" by (metis assms homotopic_loops_imp_subset) |
71745 | 754 |
let ?A = "{0..1::real} \<times> {0..1::real}" |
755 |
obtain h where conth: "continuous_on ?A h" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
756 |
and hs: "h \<in> ?A \<rightarrow> S" |
77684 | 757 |
and h0[simp]: "\<And>x. x \<in> {0..1} \<Longrightarrow> h(0,x) = p x" |
758 |
and h1[simp]: "\<And>x. x \<in> {0..1} \<Longrightarrow> h(1,x) = a" |
|
69620 | 759 |
and ends: "\<And>t. t \<in> {0..1} \<Longrightarrow> pathfinish (h \<circ> Pair t) = pathstart (h \<circ> Pair t)" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
760 |
using assms by (auto simp: homotopic_loops homotopic_with image_subset_iff_funcset) |
69620 | 761 |
have conth0: "path (\<lambda>u. h (u, 0))" |
762 |
unfolding path_def |
|
71745 | 763 |
proof (rule continuous_on_compose [of _ _ h, unfolded o_def]) |
764 |
show "continuous_on ((\<lambda>x. (x, 0)) ` {0..1}) h" |
|
765 |
by (force intro: continuous_on_subset [OF conth]) |
|
766 |
qed (force intro: continuous_intros) |
|
77684 | 767 |
have pih0: "path_image (\<lambda>u. h (u, 0)) \<subseteq> S" |
69620 | 768 |
using hs by (force simp: path_image_def) |
71745 | 769 |
have c1: "continuous_on ?A (\<lambda>x. h (fst x * snd x, 0))" |
770 |
proof (rule continuous_on_compose [of _ _ h, unfolded o_def]) |
|
771 |
show "continuous_on ((\<lambda>x. (fst x * snd x, 0)) ` ?A) h" |
|
772 |
by (force simp: mult_le_one intro: continuous_on_subset [OF conth]) |
|
773 |
qed (force intro: continuous_intros)+ |
|
774 |
have c2: "continuous_on ?A (\<lambda>x. h (fst x - fst x * snd x, 0))" |
|
775 |
proof (rule continuous_on_compose [of _ _ h, unfolded o_def]) |
|
776 |
show "continuous_on ((\<lambda>x. (fst x - fst x * snd x, 0)) ` ?A) h" |
|
777 |
by (auto simp: algebra_simps add_increasing2 mult_left_le intro: continuous_on_subset [OF conth]) |
|
778 |
qed (force intro: continuous_intros) |
|
69620 | 779 |
have [simp]: "\<And>t. \<lbrakk>0 \<le> t \<and> t \<le> 1\<rbrakk> \<Longrightarrow> h (t, 1) = h (t, 0)" |
780 |
using ends by (simp add: pathfinish_def pathstart_def) |
|
781 |
have adhoc_le: "c * 4 \<le> 1 + c * (d * 4)" if "\<not> d * 4 \<le> 3" "0 \<le> c" "c \<le> 1" for c d::real |
|
782 |
proof - |
|
783 |
have "c * 3 \<le> c * (d * 4)" using that less_eq_real_def by auto |
|
784 |
with \<open>c \<le> 1\<close> show ?thesis by fastforce |
|
785 |
qed |
|
71746 | 786 |
have *: "\<And>p x. \<lbrakk>path p \<and> path(reversepath p); |
77684 | 787 |
path_image p \<subseteq> S \<and> path_image(reversepath p) \<subseteq> S; |
71746 | 788 |
pathfinish p = pathstart(linepath a a +++ reversepath p) \<and> |
789 |
pathstart(reversepath p) = a \<and> pathstart p = x\<rbrakk> |
|
77684 | 790 |
\<Longrightarrow> homotopic_paths S (p +++ linepath a a +++ reversepath p) (linepath x x)" |
69620 | 791 |
by (metis homotopic_paths_lid homotopic_paths_join |
792 |
homotopic_paths_trans homotopic_paths_sym homotopic_paths_rinv) |
|
77684 | 793 |
have 1: "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))" |
69620 | 794 |
using \<open>path p\<close> homotopic_paths_rid homotopic_paths_sym pip by blast |
77684 | 795 |
moreover have "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p)) |
69620 | 796 |
(linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p))" |
77684 | 797 |
using homotopic_paths_lid [of "p +++ linepath (pathfinish p) (pathfinish p)" S] |
798 |
by (metis 1 homotopic_paths_imp_path homotopic_paths_imp_subset homotopic_paths_sym pathstart_join) |
|
71745 | 799 |
moreover |
77684 | 800 |
have "homotopic_paths S (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p)) |
69620 | 801 |
((\<lambda>u. h (u, 0)) +++ linepath a a +++ reversepath (\<lambda>u. h (u, 0)))" |
71745 | 802 |
unfolding homotopic_paths_def homotopic_with_def |
803 |
proof (intro exI strip conjI) |
|
71746 | 804 |
let ?h = "\<lambda>y. (subpath 0 (fst y) (\<lambda>u. h (u, 0)) +++ (\<lambda>u. h (Pair (fst y) u)) |
805 |
+++ subpath (fst y) 0 (\<lambda>u. h (u, 0))) (snd y)" |
|
806 |
have "continuous_on ?A ?h" |
|
807 |
by (intro continuous_on_homotopic_join_lemma; simp add: path_defs joinpaths_def subpath_def conth c1 c2) |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
808 |
moreover have "?h \<in> ?A \<rightarrow> S" |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
809 |
using hs |
71746 | 810 |
unfolding joinpaths_def subpath_def |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
811 |
by (force simp: algebra_simps mult_le_one mult_left_le intro: adhoc_le) |
71746 | 812 |
ultimately show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set {0..1})) |
77684 | 813 |
(top_of_set S) ?h" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
814 |
by (simp add: subpath_reversepath image_subset_iff_funcset) |
71745 | 815 |
qed (use ploop in \<open>simp_all add: reversepath_def path_defs joinpaths_def o_def subpath_def conth c1 c2\<close>) |
77684 | 816 |
moreover have "homotopic_paths S ((\<lambda>u. h (u, 0)) +++ linepath a a +++ reversepath (\<lambda>u. h (u, 0))) |
69620 | 817 |
(linepath (pathstart p) (pathstart p))" |
77684 | 818 |
by (rule *; simp add: pih0 pathstart_def pathfinish_def conth0; simp add: reversepath_def joinpaths_def) |
69620 | 819 |
ultimately show ?thesis |
820 |
by (blast intro: homotopic_paths_trans) |
|
821 |
qed |
|
822 |
||
823 |
proposition homotopic_loops_conjugate: |
|
77684 | 824 |
fixes S :: "'a::real_normed_vector set" |
825 |
assumes "path p" "path q" and pip: "path_image p \<subseteq> S" and piq: "path_image q \<subseteq> S" |
|
71746 | 826 |
and pq: "pathfinish p = pathstart q" and qloop: "pathfinish q = pathstart q" |
77684 | 827 |
shows "homotopic_loops S (p +++ q +++ reversepath p) q" |
69620 | 828 |
proof - |
829 |
have contp: "continuous_on {0..1} p" using \<open>path p\<close> [unfolded path_def] by blast |
|
830 |
have contq: "continuous_on {0..1} q" using \<open>path q\<close> [unfolded path_def] by blast |
|
71745 | 831 |
let ?A = "{0..1::real} \<times> {0..1::real}" |
832 |
have c1: "continuous_on ?A (\<lambda>x. p ((1 - fst x) * snd x + fst x))" |
|
833 |
proof (rule continuous_on_compose [of _ _ p, unfolded o_def]) |
|
834 |
show "continuous_on ((\<lambda>x. (1 - fst x) * snd x + fst x) ` ?A) p" |
|
835 |
by (auto intro: continuous_on_subset [OF contp] simp: algebra_simps add_increasing2 mult_right_le_one_le sum_le_prod1) |
|
836 |
qed (force intro: continuous_intros) |
|
837 |
have c2: "continuous_on ?A (\<lambda>x. p ((fst x - 1) * snd x + 1))" |
|
838 |
proof (rule continuous_on_compose [of _ _ p, unfolded o_def]) |
|
839 |
show "continuous_on ((\<lambda>x. (fst x - 1) * snd x + 1) ` ?A) p" |
|
840 |
by (auto intro: continuous_on_subset [OF contp] simp: algebra_simps add_increasing2 mult_left_le_one_le) |
|
841 |
qed (force intro: continuous_intros) |
|
842 |
||
77684 | 843 |
have ps1: "\<And>a b. \<lbrakk>b * 2 \<le> 1; 0 \<le> b; 0 \<le> a; a \<le> 1\<rbrakk> \<Longrightarrow> p ((1 - a) * (2 * b) + a) \<in> S" |
69620 | 844 |
using sum_le_prod1 |
845 |
by (force simp: algebra_simps add_increasing2 mult_left_le intro: pip [unfolded path_image_def, THEN subsetD]) |
|
77684 | 846 |
have ps2: "\<And>a b. \<lbrakk>\<not> 4 * b \<le> 3; b \<le> 1; 0 \<le> a; a \<le> 1\<rbrakk> \<Longrightarrow> p ((a - 1) * (4 * b - 3) + 1) \<in> S" |
69620 | 847 |
apply (rule pip [unfolded path_image_def, THEN subsetD]) |
848 |
apply (rule image_eqI, blast) |
|
849 |
apply (simp add: algebra_simps) |
|
77684 | 850 |
by (metis add_mono affine_ineq linear mult.commute mult.left_neutral mult_right_mono |
69620 | 851 |
add.commute zero_le_numeral) |
77684 | 852 |
have qs: "\<And>a b. \<lbrakk>4 * b \<le> 3; \<not> b * 2 \<le> 1\<rbrakk> \<Longrightarrow> q (4 * b - 2) \<in> S" |
69620 | 853 |
using path_image_def piq by fastforce |
77684 | 854 |
have "homotopic_loops S (p +++ q +++ reversepath p) |
69620 | 855 |
(linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q))" |
71746 | 856 |
unfolding homotopic_loops_def homotopic_with_def |
857 |
proof (intro exI strip conjI) |
|
858 |
let ?h = "(\<lambda>y. (subpath (fst y) 1 p +++ q +++ subpath 1 (fst y) p) (snd y))" |
|
859 |
have "continuous_on ?A (\<lambda>y. q (snd y))" |
|
860 |
by (force simp: contq intro: continuous_on_compose [of _ _ q, unfolded o_def] continuous_on_id continuous_on_snd) |
|
861 |
then have "continuous_on ?A ?h" |
|
862 |
using pq qloop |
|
72372 | 863 |
by (intro continuous_on_homotopic_join_lemma) (auto simp: path_defs joinpaths_def subpath_def c1 c2) |
77684 | 864 |
then show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set {0..1})) (top_of_set S) ?h" |
71746 | 865 |
by (auto simp: joinpaths_def subpath_def ps1 ps2 qs) |
866 |
show "?h (1,x) = (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) x" for x |
|
867 |
using pq by (simp add: pathfinish_def subpath_refl) |
|
868 |
qed (auto simp: subpath_reversepath) |
|
77684 | 869 |
moreover have "homotopic_loops S (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) q" |
69620 | 870 |
proof - |
77684 | 871 |
have "homotopic_paths S (linepath (pathfinish q) (pathfinish q) +++ q) q" |
69620 | 872 |
using \<open>path q\<close> homotopic_paths_lid qloop piq by auto |
77684 | 873 |
hence 1: "\<And>f. homotopic_paths S f q \<or> \<not> homotopic_paths S f (linepath (pathfinish q) (pathfinish q) +++ q)" |
69620 | 874 |
using homotopic_paths_trans by blast |
77684 | 875 |
hence "homotopic_paths S (linepath (pathfinish q) (pathfinish q) +++ q +++ linepath (pathfinish q) (pathfinish q)) q" |
876 |
by (smt (verit, best) \<open>path q\<close> homotopic_paths_imp_path homotopic_paths_imp_subset homotopic_paths_lid |
|
877 |
homotopic_paths_rid homotopic_paths_trans pathstart_join piq qloop) |
|
69620 | 878 |
thus ?thesis |
879 |
by (metis (no_types) qloop homotopic_loops_sym homotopic_paths_imp_homotopic_loops homotopic_paths_imp_pathfinish homotopic_paths_sym) |
|
880 |
qed |
|
881 |
ultimately show ?thesis |
|
882 |
by (blast intro: homotopic_loops_trans) |
|
883 |
qed |
|
884 |
||
885 |
lemma homotopic_paths_loop_parts: |
|
886 |
assumes loops: "homotopic_loops S (p +++ reversepath q) (linepath a a)" and "path q" |
|
887 |
shows "homotopic_paths S p q" |
|
888 |
proof - |
|
889 |
have paths: "homotopic_paths S (p +++ reversepath q) (linepath (pathstart p) (pathstart p))" |
|
890 |
using homotopic_loops_imp_homotopic_paths_null [OF loops] by simp |
|
891 |
then have "path p" |
|
892 |
using \<open>path q\<close> homotopic_loops_imp_path loops path_join path_join_path_ends path_reversepath by blast |
|
893 |
show ?thesis |
|
894 |
proof (cases "pathfinish p = pathfinish q") |
|
895 |
case True |
|
77684 | 896 |
obtain pipq: "path_image p \<subseteq> S" "path_image q \<subseteq> S" |
69620 | 897 |
by (metis Un_subset_iff paths \<open>path p\<close> \<open>path q\<close> homotopic_loops_imp_subset homotopic_paths_imp_path loops |
77684 | 898 |
path_image_join path_image_reversepath path_imp_reversepath path_join_eq) |
69620 | 899 |
have "homotopic_paths S p (p +++ (linepath (pathfinish p) (pathfinish p)))" |
900 |
using \<open>path p\<close> \<open>path_image p \<subseteq> S\<close> homotopic_paths_rid homotopic_paths_sym by blast |
|
901 |
moreover have "homotopic_paths S (p +++ (linepath (pathfinish p) (pathfinish p))) (p +++ (reversepath q +++ q))" |
|
902 |
by (simp add: True \<open>path p\<close> \<open>path q\<close> pipq homotopic_paths_join homotopic_paths_linv homotopic_paths_sym) |
|
903 |
moreover have "homotopic_paths S (p +++ (reversepath q +++ q)) ((p +++ reversepath q) +++ q)" |
|
904 |
by (simp add: True \<open>path p\<close> \<open>path q\<close> homotopic_paths_assoc pipq) |
|
905 |
moreover have "homotopic_paths S ((p +++ reversepath q) +++ q) (linepath (pathstart p) (pathstart p) +++ q)" |
|
906 |
by (simp add: \<open>path q\<close> homotopic_paths_join paths pipq) |
|
907 |
ultimately show ?thesis |
|
77684 | 908 |
by (metis \<open>path q\<close> homotopic_paths_imp_path homotopic_paths_lid homotopic_paths_trans path_join_path_ends pathfinish_linepath pipq(2)) |
69620 | 909 |
next |
910 |
case False |
|
911 |
then show ?thesis |
|
912 |
using \<open>path q\<close> homotopic_loops_imp_path loops path_join_path_ends by fastforce |
|
913 |
qed |
|
914 |
qed |
|
915 |
||
916 |
||
70136 | 917 |
subsection\<^marker>\<open>tag unimportant\<close>\<open>Homotopy of "nearby" function, paths and loops\<close> |
69620 | 918 |
|
919 |
lemma homotopic_with_linear: |
|
920 |
fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector" |
|
71746 | 921 |
assumes contf: "continuous_on S f" |
922 |
and contg:"continuous_on S g" |
|
923 |
and sub: "\<And>x. x \<in> S \<Longrightarrow> closed_segment (f x) (g x) \<subseteq> t" |
|
924 |
shows "homotopic_with_canon (\<lambda>z. True) S t f g" |
|
72372 | 925 |
unfolding homotopic_with_def |
69620 | 926 |
apply (rule_tac x="\<lambda>y. ((1 - (fst y)) *\<^sub>R f(snd y) + (fst y) *\<^sub>R g(snd y))" in exI) |
72372 | 927 |
using sub closed_segment_def |
928 |
by (fastforce intro: continuous_intros continuous_on_subset [OF contf] continuous_on_compose2 [where g=f] |
|
929 |
continuous_on_subset [OF contg] continuous_on_compose2 [where g=g]) |
|
69620 | 930 |
|
931 |
lemma homotopic_paths_linear: |
|
932 |
fixes g h :: "real \<Rightarrow> 'a::real_normed_vector" |
|
933 |
assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" |
|
71746 | 934 |
"\<And>t. t \<in> {0..1} \<Longrightarrow> closed_segment (g t) (h t) \<subseteq> S" |
935 |
shows "homotopic_paths S g h" |
|
69620 | 936 |
using assms |
937 |
unfolding path_def |
|
938 |
apply (simp add: closed_segment_def pathstart_def pathfinish_def homotopic_paths_def homotopic_with_def) |
|
939 |
apply (rule_tac x="\<lambda>y. ((1 - (fst y)) *\<^sub>R (g \<circ> snd) y + (fst y) *\<^sub>R (h \<circ> snd) y)" in exI) |
|
940 |
apply (intro conjI subsetI continuous_intros; force) |
|
941 |
done |
|
942 |
||
943 |
lemma homotopic_loops_linear: |
|
944 |
fixes g h :: "real \<Rightarrow> 'a::real_normed_vector" |
|
945 |
assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" |
|
71746 | 946 |
"\<And>t x. t \<in> {0..1} \<Longrightarrow> closed_segment (g t) (h t) \<subseteq> S" |
947 |
shows "homotopic_loops S g h" |
|
69620 | 948 |
using assms |
72372 | 949 |
unfolding path_defs homotopic_loops_def homotopic_with_def |
69620 | 950 |
apply (rule_tac x="\<lambda>y. ((1 - (fst y)) *\<^sub>R g(snd y) + (fst y) *\<^sub>R h(snd y))" in exI) |
72372 | 951 |
by (force simp: closed_segment_def intro!: continuous_intros intro: continuous_on_compose2 [where g=g] continuous_on_compose2 [where g=h]) |
69620 | 952 |
|
953 |
lemma homotopic_paths_nearby_explicit: |
|
71746 | 954 |
assumes \<section>: "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" |
955 |
and no: "\<And>t x. \<lbrakk>t \<in> {0..1}; x \<notin> S\<rbrakk> \<Longrightarrow> norm(h t - g t) < norm(g t - x)" |
|
956 |
shows "homotopic_paths S g h" |
|
77684 | 957 |
using homotopic_paths_linear [OF \<section>] by (metis linorder_not_le no norm_minus_commute segment_bound1 subsetI) |
69620 | 958 |
|
959 |
lemma homotopic_loops_nearby_explicit: |
|
71746 | 960 |
assumes \<section>: "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" |
961 |
and no: "\<And>t x. \<lbrakk>t \<in> {0..1}; x \<notin> S\<rbrakk> \<Longrightarrow> norm(h t - g t) < norm(g t - x)" |
|
962 |
shows "homotopic_loops S g h" |
|
77684 | 963 |
using homotopic_loops_linear [OF \<section>] by (metis linorder_not_le no norm_minus_commute segment_bound1 subsetI) |
69620 | 964 |
|
965 |
lemma homotopic_nearby_paths: |
|
966 |
fixes g h :: "real \<Rightarrow> 'a::euclidean_space" |
|
71746 | 967 |
assumes "path g" "open S" "path_image g \<subseteq> S" |
69620 | 968 |
shows "\<exists>e. 0 < e \<and> |
969 |
(\<forall>h. path h \<and> |
|
970 |
pathstart h = pathstart g \<and> pathfinish h = pathfinish g \<and> |
|
71746 | 971 |
(\<forall>t \<in> {0..1}. norm(h t - g t) < e) \<longrightarrow> homotopic_paths S g h)" |
69620 | 972 |
proof - |
71746 | 973 |
obtain e where "e > 0" and e: "\<And>x y. x \<in> path_image g \<Longrightarrow> y \<in> - S \<Longrightarrow> e \<le> dist x y" |
974 |
using separate_compact_closed [of "path_image g" "-S"] assms by force |
|
69620 | 975 |
show ?thesis |
71768 | 976 |
using e [unfolded dist_norm] \<open>e > 0\<close> |
72372 | 977 |
by (fastforce simp: path_image_def intro!: homotopic_paths_nearby_explicit assms exI) |
69620 | 978 |
qed |
979 |
||
980 |
lemma homotopic_nearby_loops: |
|
981 |
fixes g h :: "real \<Rightarrow> 'a::euclidean_space" |
|
71746 | 982 |
assumes "path g" "open S" "path_image g \<subseteq> S" "pathfinish g = pathstart g" |
69620 | 983 |
shows "\<exists>e. 0 < e \<and> |
984 |
(\<forall>h. path h \<and> pathfinish h = pathstart h \<and> |
|
71746 | 985 |
(\<forall>t \<in> {0..1}. norm(h t - g t) < e) \<longrightarrow> homotopic_loops S g h)" |
69620 | 986 |
proof - |
71746 | 987 |
obtain e where "e > 0" and e: "\<And>x y. x \<in> path_image g \<Longrightarrow> y \<in> - S \<Longrightarrow> e \<le> dist x y" |
988 |
using separate_compact_closed [of "path_image g" "-S"] assms by force |
|
69620 | 989 |
show ?thesis |
71768 | 990 |
using e [unfolded dist_norm] \<open>e > 0\<close> |
72372 | 991 |
by (fastforce simp: path_image_def intro!: homotopic_loops_nearby_explicit assms exI) |
69620 | 992 |
qed |
993 |
||
994 |
||
995 |
subsection\<open> Homotopy and subpaths\<close> |
|
996 |
||
997 |
lemma homotopic_join_subpaths1: |
|
77684 | 998 |
assumes "path g" and pag: "path_image g \<subseteq> S" |
69620 | 999 |
and u: "u \<in> {0..1}" and v: "v \<in> {0..1}" and w: "w \<in> {0..1}" "u \<le> v" "v \<le> w" |
77684 | 1000 |
shows "homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" |
69620 | 1001 |
proof - |
1002 |
have 1: "t * 2 \<le> 1 \<Longrightarrow> u + t * (v * 2) \<le> v + t * (u * 2)" for t |
|
1003 |
using affine_ineq \<open>u \<le> v\<close> by fastforce |
|
1004 |
have 2: "t * 2 > 1 \<Longrightarrow> u + (2*t - 1) * v \<le> v + (2*t - 1) * w" for t |
|
1005 |
by (metis add_mono_thms_linordered_semiring(1) diff_gt_0_iff_gt less_eq_real_def mult.commute mult_right_mono \<open>u \<le> v\<close> \<open>v \<le> w\<close>) |
|
1006 |
have t2: "\<And>t::real. t*2 = 1 \<Longrightarrow> t = 1/2" by auto |
|
71746 | 1007 |
have "homotopic_paths (path_image g) (subpath u v g +++ subpath v w g) (subpath u w g)" |
1008 |
proof (cases "w = u") |
|
1009 |
case True |
|
1010 |
then show ?thesis |
|
1011 |
by (metis \<open>path g\<close> homotopic_paths_rinv path_image_subpath_subset path_subpath pathstart_subpath reversepath_subpath subpath_refl u v) |
|
1012 |
next |
|
1013 |
case False |
|
1014 |
let ?f = "\<lambda>t. if t \<le> 1/2 then inverse((w - u)) *\<^sub>R (2 * (v - u)) *\<^sub>R t |
|
1015 |
else inverse((w - u)) *\<^sub>R ((v - u) + (w - v) *\<^sub>R (2 *\<^sub>R t - 1))" |
|
1016 |
show ?thesis |
|
1017 |
proof (rule homotopic_paths_sym [OF homotopic_paths_reparametrize [where f = ?f]]) |
|
1018 |
show "path (subpath u w g)" |
|
1019 |
using assms(1) path_subpath u w(1) by blast |
|
1020 |
show "path_image (subpath u w g) \<subseteq> path_image g" |
|
1021 |
by (meson path_image_subpath_subset u w(1)) |
|
1022 |
show "continuous_on {0..1} ?f" |
|
1023 |
unfolding split_01 |
|
1024 |
by (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def dest!: t2)+ |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1025 |
show "?f \<in> {0..1} \<rightarrow> {0..1}" |
71746 | 1026 |
using False assms |
1027 |
by (force simp: field_simps not_le mult_left_mono affine_ineq dest!: 1 2) |
|
1028 |
show "(subpath u v g +++ subpath v w g) t = subpath u w g (?f t)" if "t \<in> {0..1}" for t |
|
1029 |
using assms |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1030 |
unfolding joinpaths_def subpath_def by (auto simp: divide_simps add.commute mult.commute mult.left_commute) |
71746 | 1031 |
qed (use False in auto) |
1032 |
qed |
|
1033 |
then show ?thesis |
|
1034 |
by (rule homotopic_paths_subset [OF _ pag]) |
|
69620 | 1035 |
qed |
1036 |
||
1037 |
lemma homotopic_join_subpaths2: |
|
77684 | 1038 |
assumes "homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" |
1039 |
shows "homotopic_paths S (subpath w v g +++ subpath v u g) (subpath w u g)" |
|
1040 |
by (metis assms homotopic_paths_reversepath_D pathfinish_subpath pathstart_subpath reversepath_joinpaths reversepath_subpath) |
|
69620 | 1041 |
|
1042 |
lemma homotopic_join_subpaths3: |
|
77684 | 1043 |
assumes hom: "homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" |
1044 |
and "path g" and pag: "path_image g \<subseteq> S" |
|
69620 | 1045 |
and u: "u \<in> {0..1}" and v: "v \<in> {0..1}" and w: "w \<in> {0..1}" |
77684 | 1046 |
shows "homotopic_paths S (subpath v w g +++ subpath w u g) (subpath v u g)" |
69620 | 1047 |
proof - |
77684 | 1048 |
obtain wvg: "path (subpath w v g)" "path_image (subpath w v g) \<subseteq> S" |
1049 |
and wug: "path (subpath w u g)" "path_image (subpath w u g) \<subseteq> S" |
|
1050 |
and vug: "path (subpath v u g)" "path_image (subpath v u g) \<subseteq> S" |
|
1051 |
by (meson \<open>path g\<close> pag path_image_subpath_subset path_subpath subset_trans u v w) |
|
1052 |
have "homotopic_paths S (subpath u w g +++ subpath w v g) |
|
1053 |
((subpath u v g +++ subpath v w g) +++ subpath w v g)" |
|
1054 |
by (simp add: hom homotopic_paths_join homotopic_paths_sym wvg) |
|
78474 | 1055 |
also have "homotopic_paths S \<dots> (subpath u v g +++ subpath v w g +++ subpath w v g)" |
77684 | 1056 |
using wvg vug \<open>path g\<close> |
1057 |
by (metis homotopic_paths_assoc homotopic_paths_sym path_image_subpath_commute path_subpath |
|
1058 |
pathfinish_subpath pathstart_subpath u v w) |
|
78474 | 1059 |
also have "homotopic_paths S \<dots> (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))" |
77684 | 1060 |
using wvg vug \<open>path g\<close> |
1061 |
by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl path_image_subpath_commute |
|
1062 |
path_subpath pathfinish_subpath pathstart_join pathstart_subpath reversepath_subpath u v) |
|
78474 | 1063 |
also have "homotopic_paths S \<dots> (subpath u v g)" |
77684 | 1064 |
using vug \<open>path g\<close> by (metis homotopic_paths_rid path_image_subpath_commute path_subpath u v) |
1065 |
finally have "homotopic_paths S (subpath u w g +++ subpath w v g) (subpath u v g)" . |
|
69620 | 1066 |
then show ?thesis |
1067 |
using homotopic_join_subpaths2 by blast |
|
1068 |
qed |
|
1069 |
||
1070 |
proposition homotopic_join_subpaths: |
|
77684 | 1071 |
"\<lbrakk>path g; path_image g \<subseteq> S; u \<in> {0..1}; v \<in> {0..1}; w \<in> {0..1}\<rbrakk> |
1072 |
\<Longrightarrow> homotopic_paths S (subpath u v g +++ subpath v w g) (subpath u w g)" |
|
1073 |
by (smt (verit, del_insts) homotopic_join_subpaths1 homotopic_join_subpaths2 homotopic_join_subpaths3) |
|
69620 | 1074 |
|
1075 |
text\<open>Relating homotopy of trivial loops to path-connectedness.\<close> |
|
1076 |
||
1077 |
lemma path_component_imp_homotopic_points: |
|
71768 | 1078 |
assumes "path_component S a b" |
1079 |
shows "homotopic_loops S (linepath a a) (linepath b b)" |
|
1080 |
proof - |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1081 |
obtain g :: "real \<Rightarrow> 'a" where g: "continuous_on {0..1} g" "g \<in> {0..1} \<rightarrow> S" "g 0 = a" "g 1 = b" |
71768 | 1082 |
using assms by (auto simp: path_defs) |
1083 |
then have "continuous_on ({0..1} \<times> {0..1}) (g \<circ> fst)" |
|
1084 |
by (fastforce intro!: continuous_intros)+ |
|
1085 |
with g show ?thesis |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1086 |
by (auto simp: homotopic_loops_def homotopic_with_def path_defs Pi_iff) |
71768 | 1087 |
qed |
69620 | 1088 |
|
1089 |
lemma homotopic_loops_imp_path_component_value: |
|
78474 | 1090 |
"\<lbrakk>homotopic_loops S p q; 0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> path_component S (p t) (q t)" |
1091 |
apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs) |
|
1092 |
apply (rule_tac x="h \<circ> (\<lambda>u. (u, t))" in exI) |
|
1093 |
apply (fastforce elim!: continuous_on_subset intro!: continuous_intros) |
|
1094 |
done |
|
69620 | 1095 |
|
1096 |
lemma homotopic_points_eq_path_component: |
|
71768 | 1097 |
"homotopic_loops S (linepath a a) (linepath b b) \<longleftrightarrow> path_component S a b" |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1098 |
using homotopic_loops_imp_path_component_value path_component_imp_homotopic_points by fastforce |
69620 | 1099 |
|
1100 |
lemma path_connected_eq_homotopic_points: |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1101 |
"path_connected S \<longleftrightarrow> |
69620 | 1102 |
(\<forall>a b. a \<in> S \<and> b \<in> S \<longrightarrow> homotopic_loops S (linepath a a) (linepath b b))" |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1103 |
by (auto simp: path_connected_def path_component_def homotopic_points_eq_path_component) |
69620 | 1104 |
|
1105 |
||
1106 |
subsection\<open>Simply connected sets\<close> |
|
1107 |
||
70136 | 1108 |
text\<^marker>\<open>tag important\<close>\<open>defined as "all loops are homotopic (as loops)\<close> |
1109 |
||
1110 |
definition\<^marker>\<open>tag important\<close> simply_connected where |
|
69620 | 1111 |
"simply_connected S \<equiv> |
1112 |
\<forall>p q. path p \<and> pathfinish p = pathstart p \<and> path_image p \<subseteq> S \<and> |
|
1113 |
path q \<and> pathfinish q = pathstart q \<and> path_image q \<subseteq> S |
|
1114 |
\<longrightarrow> homotopic_loops S p q" |
|
1115 |
||
1116 |
lemma simply_connected_empty [iff]: "simply_connected {}" |
|
1117 |
by (simp add: simply_connected_def) |
|
1118 |
||
1119 |
lemma simply_connected_imp_path_connected: |
|
1120 |
fixes S :: "_::real_normed_vector set" |
|
1121 |
shows "simply_connected S \<Longrightarrow> path_connected S" |
|
78474 | 1122 |
by (simp add: simply_connected_def path_connected_eq_homotopic_points) |
69620 | 1123 |
|
1124 |
lemma simply_connected_imp_connected: |
|
1125 |
fixes S :: "_::real_normed_vector set" |
|
1126 |
shows "simply_connected S \<Longrightarrow> connected S" |
|
78474 | 1127 |
by (simp add: path_connected_imp_connected simply_connected_imp_path_connected) |
69620 | 1128 |
|
1129 |
lemma simply_connected_eq_contractible_loop_any: |
|
1130 |
fixes S :: "_::real_normed_vector set" |
|
1131 |
shows "simply_connected S \<longleftrightarrow> |
|
71768 | 1132 |
(\<forall>p a. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p \<and> a \<in> S |
69620 | 1133 |
\<longrightarrow> homotopic_loops S p (linepath a a))" |
71768 | 1134 |
(is "?lhs = ?rhs") |
1135 |
proof |
|
1136 |
assume ?rhs then show ?lhs |
|
1137 |
unfolding simply_connected_def |
|
1138 |
by (metis pathfinish_in_path_image subsetD homotopic_loops_trans homotopic_loops_sym) |
|
78474 | 1139 |
qed (force simp: simply_connected_def) |
69620 | 1140 |
|
1141 |
lemma simply_connected_eq_contractible_loop_some: |
|
1142 |
fixes S :: "_::real_normed_vector set" |
|
1143 |
shows "simply_connected S \<longleftrightarrow> |
|
1144 |
path_connected S \<and> |
|
1145 |
(\<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p |
|
1146 |
\<longrightarrow> (\<exists>a. a \<in> S \<and> homotopic_loops S p (linepath a a)))" |
|
71746 | 1147 |
(is "?lhs = ?rhs") |
1148 |
proof |
|
1149 |
assume ?lhs |
|
1150 |
then show ?rhs |
|
1151 |
using simply_connected_eq_contractible_loop_any by (blast intro: simply_connected_imp_path_connected) |
|
1152 |
next |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1153 |
assume ?rhs |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1154 |
then show ?lhs |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1155 |
by (meson homotopic_loops_trans path_connected_eq_homotopic_points simply_connected_eq_contractible_loop_any) |
71746 | 1156 |
qed |
69620 | 1157 |
|
1158 |
lemma simply_connected_eq_contractible_loop_all: |
|
1159 |
fixes S :: "_::real_normed_vector set" |
|
1160 |
shows "simply_connected S \<longleftrightarrow> |
|
1161 |
S = {} \<or> |
|
1162 |
(\<exists>a \<in> S. \<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p |
|
1163 |
\<longrightarrow> homotopic_loops S p (linepath a a))" |
|
78474 | 1164 |
by (meson ex_in_conv homotopic_loops_sym homotopic_loops_trans simply_connected_def simply_connected_eq_contractible_loop_any) |
69620 | 1165 |
|
1166 |
lemma simply_connected_eq_contractible_path: |
|
1167 |
fixes S :: "_::real_normed_vector set" |
|
1168 |
shows "simply_connected S \<longleftrightarrow> |
|
1169 |
path_connected S \<and> |
|
1170 |
(\<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p |
|
1171 |
\<longrightarrow> homotopic_paths S p (linepath (pathstart p) (pathstart p)))" |
|
71746 | 1172 |
(is "?lhs = ?rhs") |
1173 |
proof |
|
1174 |
assume ?lhs |
|
1175 |
then show ?rhs |
|
1176 |
unfolding simply_connected_imp_path_connected |
|
1177 |
by (metis simply_connected_eq_contractible_loop_some homotopic_loops_imp_homotopic_paths_null) |
|
1178 |
next |
|
1179 |
assume ?rhs |
|
1180 |
then show ?lhs |
|
1181 |
using homotopic_paths_imp_homotopic_loops simply_connected_eq_contractible_loop_some by fastforce |
|
1182 |
qed |
|
69620 | 1183 |
|
1184 |
lemma simply_connected_eq_homotopic_paths: |
|
1185 |
fixes S :: "_::real_normed_vector set" |
|
1186 |
shows "simply_connected S \<longleftrightarrow> |
|
1187 |
path_connected S \<and> |
|
1188 |
(\<forall>p q. path p \<and> path_image p \<subseteq> S \<and> |
|
1189 |
path q \<and> path_image q \<subseteq> S \<and> |
|
1190 |
pathstart q = pathstart p \<and> pathfinish q = pathfinish p |
|
1191 |
\<longrightarrow> homotopic_paths S p q)" |
|
1192 |
(is "?lhs = ?rhs") |
|
1193 |
proof |
|
1194 |
assume ?lhs |
|
1195 |
then have pc: "path_connected S" |
|
1196 |
and *: "\<And>p. \<lbrakk>path p; path_image p \<subseteq> S; |
|
1197 |
pathfinish p = pathstart p\<rbrakk> |
|
1198 |
\<Longrightarrow> homotopic_paths S p (linepath (pathstart p) (pathstart p))" |
|
1199 |
by (auto simp: simply_connected_eq_contractible_path) |
|
1200 |
have "homotopic_paths S p q" |
|
1201 |
if "path p" "path_image p \<subseteq> S" "path q" |
|
1202 |
"path_image q \<subseteq> S" "pathstart q = pathstart p" |
|
1203 |
"pathfinish q = pathfinish p" for p q |
|
1204 |
proof - |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1205 |
have "homotopic_paths S p (p +++ reversepath q +++ q)" |
69620 | 1206 |
using that |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1207 |
by (smt (verit, best) homotopic_paths_join homotopic_paths_linv homotopic_paths_rid homotopic_paths_sym |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1208 |
homotopic_paths_trans pathstart_linepath) |
78474 | 1209 |
also have "homotopic_paths S \<dots> ((p +++ reversepath q) +++ q)" |
69620 | 1210 |
by (simp add: that homotopic_paths_assoc) |
78474 | 1211 |
also have "homotopic_paths S \<dots> (linepath (pathstart q) (pathstart q) +++ q)" |
69620 | 1212 |
using * [of "p +++ reversepath q"] that |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1213 |
by (simp add: homotopic_paths_assoc homotopic_paths_join path_image_join) |
78474 | 1214 |
also have "homotopic_paths S \<dots> q" |
69620 | 1215 |
using that homotopic_paths_lid by blast |
1216 |
finally show ?thesis . |
|
1217 |
qed |
|
1218 |
then show ?rhs |
|
1219 |
by (blast intro: pc *) |
|
1220 |
next |
|
1221 |
assume ?rhs |
|
1222 |
then show ?lhs |
|
1223 |
by (force simp: simply_connected_eq_contractible_path) |
|
1224 |
qed |
|
1225 |
||
1226 |
proposition simply_connected_Times: |
|
1227 |
fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" |
|
1228 |
assumes S: "simply_connected S" and T: "simply_connected T" |
|
1229 |
shows "simply_connected(S \<times> T)" |
|
1230 |
proof - |
|
1231 |
have "homotopic_loops (S \<times> T) p (linepath (a, b) (a, b))" |
|
1232 |
if "path p" "path_image p \<subseteq> S \<times> T" "p 1 = p 0" "a \<in> S" "b \<in> T" |
|
1233 |
for p a b |
|
1234 |
proof - |
|
1235 |
have "path (fst \<circ> p)" |
|
71746 | 1236 |
by (simp add: continuous_on_fst Path_Connected.path_continuous_image [OF \<open>path p\<close>]) |
69620 | 1237 |
moreover have "path_image (fst \<circ> p) \<subseteq> S" |
78474 | 1238 |
using that by (force simp: path_image_def) |
69620 | 1239 |
ultimately have p1: "homotopic_loops S (fst \<circ> p) (linepath a a)" |
1240 |
using S that |
|
71746 | 1241 |
by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) |
69620 | 1242 |
have "path (snd \<circ> p)" |
71746 | 1243 |
by (simp add: continuous_on_snd Path_Connected.path_continuous_image [OF \<open>path p\<close>]) |
69620 | 1244 |
moreover have "path_image (snd \<circ> p) \<subseteq> T" |
71768 | 1245 |
using that by (force simp: path_image_def) |
69620 | 1246 |
ultimately have p2: "homotopic_loops T (snd \<circ> p) (linepath b b)" |
1247 |
using T that |
|
71746 | 1248 |
by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) |
69620 | 1249 |
show ?thesis |
71746 | 1250 |
using p1 p2 unfolding homotopic_loops |
1251 |
apply clarify |
|
72372 | 1252 |
subgoal for h k |
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
1253 |
by (rule_tac x="\<lambda>z. (h z, k z)" in exI) (auto intro: continuous_intros simp: path_defs) |
69620 | 1254 |
done |
1255 |
qed |
|
1256 |
with assms show ?thesis |
|
1257 |
by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) |
|
1258 |
qed |
|
1259 |
||
1260 |
||
1261 |
subsection\<open>Contractible sets\<close> |
|
1262 |
||
70136 | 1263 |
definition\<^marker>\<open>tag important\<close> contractible where |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1264 |
"contractible S \<equiv> \<exists>a. homotopic_with_canon (\<lambda>x. True) S S id (\<lambda>x. a)" |
69620 | 1265 |
|
1266 |
proposition contractible_imp_simply_connected: |
|
1267 |
fixes S :: "_::real_normed_vector set" |
|
1268 |
assumes "contractible S" shows "simply_connected S" |
|
1269 |
proof (cases "S = {}") |
|
1270 |
case True then show ?thesis by force |
|
1271 |
next |
|
1272 |
case False |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1273 |
obtain a where a: "homotopic_with_canon (\<lambda>x. True) S S id (\<lambda>x. a)" |
69620 | 1274 |
using assms by (force simp: contractible_def) |
1275 |
then have "a \<in> S" |
|
78336 | 1276 |
using False homotopic_with_imp_funspace2 by fastforce |
71746 | 1277 |
have "\<forall>p. path p \<and> |
1278 |
path_image p \<subseteq> S \<and> pathfinish p = pathstart p \<longrightarrow> |
|
1279 |
homotopic_loops S p (linepath a a)" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1280 |
using a apply (clarsimp simp: homotopic_loops_def homotopic_with_def path_defs) |
69620 | 1281 |
apply (rule_tac x="(h \<circ> (\<lambda>y. (fst y, (p \<circ> snd) y)))" in exI) |
71746 | 1282 |
apply (intro conjI continuous_on_compose continuous_intros; force elim: continuous_on_subset) |
69620 | 1283 |
done |
71746 | 1284 |
with \<open>a \<in> S\<close> show ?thesis |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1285 |
by (auto simp: simply_connected_eq_contractible_loop_all False) |
69620 | 1286 |
qed |
1287 |
||
1288 |
corollary contractible_imp_connected: |
|
1289 |
fixes S :: "_::real_normed_vector set" |
|
1290 |
shows "contractible S \<Longrightarrow> connected S" |
|
78474 | 1291 |
by (simp add: contractible_imp_simply_connected simply_connected_imp_connected) |
69620 | 1292 |
|
1293 |
lemma contractible_imp_path_connected: |
|
1294 |
fixes S :: "_::real_normed_vector set" |
|
1295 |
shows "contractible S \<Longrightarrow> path_connected S" |
|
78474 | 1296 |
by (simp add: contractible_imp_simply_connected simply_connected_imp_path_connected) |
69620 | 1297 |
|
1298 |
lemma nullhomotopic_through_contractible: |
|
1299 |
fixes S :: "_::topological_space set" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1300 |
assumes f: "continuous_on S f" "f \<in> S \<rightarrow> T" |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1301 |
and g: "continuous_on T g" "g \<in> T \<rightarrow> U" |
69620 | 1302 |
and T: "contractible T" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1303 |
obtains c where "homotopic_with_canon (\<lambda>h. True) S U (g \<circ> f) (\<lambda>x. c)" |
69620 | 1304 |
proof - |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1305 |
obtain b where b: "homotopic_with_canon (\<lambda>x. True) T T id (\<lambda>x. b)" |
69620 | 1306 |
using assms by (force simp: contractible_def) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1307 |
have "homotopic_with_canon (\<lambda>f. True) T U (g \<circ> id) (g \<circ> (\<lambda>x. b))" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1308 |
by (metis b continuous_map_subtopology_eu g homotopic_with_compose_continuous_map_left image_subset_iff_funcset) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1309 |
then have "homotopic_with_canon (\<lambda>f. True) S U (g \<circ> id \<circ> f) (g \<circ> (\<lambda>x. b) \<circ> f)" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1310 |
by (simp add: f homotopic_with_compose_continuous_map_right image_subset_iff_funcset) |
69620 | 1311 |
then show ?thesis |
1312 |
by (simp add: comp_def that) |
|
1313 |
qed |
|
1314 |
||
1315 |
lemma nullhomotopic_into_contractible: |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1316 |
assumes f: "continuous_on S f" "f \<in> S \<rightarrow> T" |
69620 | 1317 |
and T: "contractible T" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1318 |
obtains c where "homotopic_with_canon (\<lambda>h. True) S T f (\<lambda>x. c)" |
71746 | 1319 |
by (rule nullhomotopic_through_contractible [OF f, of id T]) (use assms in auto) |
69620 | 1320 |
|
1321 |
lemma nullhomotopic_from_contractible: |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1322 |
assumes f: "continuous_on S f" "f \<in> S \<rightarrow> T" |
69620 | 1323 |
and S: "contractible S" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1324 |
obtains c where "homotopic_with_canon (\<lambda>h. True) S T f (\<lambda>x. c)" |
71768 | 1325 |
by (auto simp: comp_def intro: nullhomotopic_through_contractible [OF continuous_on_id _ f S]) |
69620 | 1326 |
|
1327 |
lemma homotopic_through_contractible: |
|
1328 |
fixes S :: "_::real_normed_vector set" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1329 |
assumes "continuous_on S f1" "f1 \<in> S \<rightarrow> T" |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1330 |
"continuous_on T g1" "g1 \<in> T \<rightarrow> U" |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1331 |
"continuous_on S f2" "f2 \<in> S \<rightarrow> T" |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1332 |
"continuous_on T g2" "g2 \<in> T \<rightarrow> U" |
69620 | 1333 |
"contractible T" "path_connected U" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1334 |
shows "homotopic_with_canon (\<lambda>h. True) S U (g1 \<circ> f1) (g2 \<circ> f2)" |
69620 | 1335 |
proof - |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1336 |
obtain c1 where c1: "homotopic_with_canon (\<lambda>h. True) S U (g1 \<circ> f1) (\<lambda>x. c1)" |
71746 | 1337 |
by (rule nullhomotopic_through_contractible [of S f1 T g1 U]) (use assms in auto) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1338 |
obtain c2 where c2: "homotopic_with_canon (\<lambda>h. True) S U (g2 \<circ> f2) (\<lambda>x. c2)" |
71746 | 1339 |
by (rule nullhomotopic_through_contractible [of S f2 T g2 U]) (use assms in auto) |
1340 |
have "S = {} \<or> (\<exists>t. path_connected t \<and> t \<subseteq> U \<and> c2 \<in> t \<and> c1 \<in> t)" |
|
69620 | 1341 |
proof (cases "S = {}") |
1342 |
case True then show ?thesis by force |
|
1343 |
next |
|
1344 |
case False |
|
1345 |
with c1 c2 have "c1 \<in> U" "c2 \<in> U" |
|
78336 | 1346 |
using homotopic_with_imp_continuous_maps |
1347 |
by (metis PiE equals0I homotopic_with_imp_funspace2)+ |
|
69620 | 1348 |
with \<open>path_connected U\<close> show ?thesis by blast |
1349 |
qed |
|
71746 | 1350 |
then have "homotopic_with_canon (\<lambda>h. True) S U (\<lambda>x. c2) (\<lambda>x. c1)" |
78474 | 1351 |
by (auto simp: path_component homotopic_constant_maps) |
71746 | 1352 |
then show ?thesis |
1353 |
using c1 c2 homotopic_with_symD homotopic_with_trans by blast |
|
69620 | 1354 |
qed |
1355 |
||
1356 |
lemma homotopic_into_contractible: |
|
1357 |
fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1358 |
assumes f: "continuous_on S f" "f \<in> S \<rightarrow> T" |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1359 |
and g: "continuous_on S g" "g \<in> S \<rightarrow> T" |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1360 |
and T: "contractible T" |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1361 |
shows "homotopic_with_canon (\<lambda>h. True) S T f g" |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1362 |
using homotopic_through_contractible [of S f T id T g id] |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1363 |
by (simp add: assms contractible_imp_path_connected) |
69620 | 1364 |
|
1365 |
lemma homotopic_from_contractible: |
|
1366 |
fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1367 |
assumes f: "continuous_on S f" "f \<in> S \<rightarrow> T" |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1368 |
and g: "continuous_on S g" "g \<in> S \<rightarrow> T" |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1369 |
and "contractible S" "path_connected T" |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1370 |
shows "homotopic_with_canon (\<lambda>h. True) S T f g" |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1371 |
using homotopic_through_contractible [of S id S f T id g] |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1372 |
by (simp add: assms contractible_imp_path_connected) |
69620 | 1373 |
|
71233 | 1374 |
subsection\<open>Starlike sets\<close> |
1375 |
||
1376 |
definition\<^marker>\<open>tag important\<close> "starlike S \<longleftrightarrow> (\<exists>a\<in>S. \<forall>x\<in>S. closed_segment a x \<subseteq> S)" |
|
1377 |
||
1378 |
lemma starlike_UNIV [simp]: "starlike UNIV" |
|
1379 |
by (simp add: starlike_def) |
|
1380 |
||
1381 |
lemma convex_imp_starlike: |
|
1382 |
"convex S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> starlike S" |
|
1383 |
unfolding convex_contains_segment starlike_def by auto |
|
1384 |
||
1385 |
lemma starlike_convex_tweak_boundary_points: |
|
1386 |
fixes S :: "'a::euclidean_space set" |
|
1387 |
assumes "convex S" "S \<noteq> {}" and ST: "rel_interior S \<subseteq> T" and TS: "T \<subseteq> closure S" |
|
1388 |
shows "starlike T" |
|
1389 |
proof - |
|
1390 |
have "rel_interior S \<noteq> {}" |
|
1391 |
by (simp add: assms rel_interior_eq_empty) |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1392 |
with ST obtain a where a: "a \<in> rel_interior S" and "a \<in> T" by blast |
71746 | 1393 |
have "\<And>x. x \<in> T \<Longrightarrow> open_segment a x \<subseteq> rel_interior S" |
1394 |
by (rule rel_interior_closure_convex_segment [OF \<open>convex S\<close> a]) (use assms in auto) |
|
1395 |
then have "\<forall>x\<in>T. a \<in> T \<and> open_segment a x \<subseteq> T" |
|
71768 | 1396 |
using ST by (blast intro: a \<open>a \<in> T\<close> rel_interior_closure_convex_segment [OF \<open>convex S\<close> a]) |
71746 | 1397 |
then show ?thesis |
1398 |
unfolding starlike_def using bexI [OF _ \<open>a \<in> T\<close>] |
|
1399 |
by (simp add: closed_segment_eq_open) |
|
71233 | 1400 |
qed |
1401 |
||
69620 | 1402 |
lemma starlike_imp_contractible_gen: |
1403 |
fixes S :: "'a::real_normed_vector set" |
|
1404 |
assumes S: "starlike S" |
|
1405 |
and P: "\<And>a T. \<lbrakk>a \<in> S; 0 \<le> T; T \<le> 1\<rbrakk> \<Longrightarrow> P(\<lambda>x. (1 - T) *\<^sub>R x + T *\<^sub>R a)" |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1406 |
obtains a where "homotopic_with_canon P S S (\<lambda>x. x) (\<lambda>x. a)" |
69620 | 1407 |
proof - |
1408 |
obtain a where "a \<in> S" and a: "\<And>x. x \<in> S \<Longrightarrow> closed_segment a x \<subseteq> S" |
|
1409 |
using S by (auto simp: starlike_def) |
|
71768 | 1410 |
have "\<And>t b. 0 \<le> t \<and> t \<le> 1 \<Longrightarrow> |
1411 |
\<exists>u. (1 - t) *\<^sub>R b + t *\<^sub>R a = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1" |
|
1412 |
by (metis add_diff_cancel_right' diff_ge_0_iff_ge le_add_diff_inverse pth_c(1)) |
|
1413 |
then have "(\<lambda>y. (1 - fst y) *\<^sub>R snd y + fst y *\<^sub>R a) ` ({0..1} \<times> S) \<subseteq> S" |
|
1414 |
using a [unfolded closed_segment_def] by force |
|
71746 | 1415 |
then have "homotopic_with_canon P S S (\<lambda>x. x) (\<lambda>x. a)" |
69620 | 1416 |
using \<open>a \<in> S\<close> |
71768 | 1417 |
unfolding homotopic_with_def |
69620 | 1418 |
apply (rule_tac x="\<lambda>y. (1 - (fst y)) *\<^sub>R snd y + (fst y) *\<^sub>R a" in exI) |
78474 | 1419 |
apply (force simp: P intro: continuous_intros) |
69620 | 1420 |
done |
71746 | 1421 |
then show ?thesis |
1422 |
using that by blast |
|
69620 | 1423 |
qed |
1424 |
||
1425 |
lemma starlike_imp_contractible: |
|
1426 |
fixes S :: "'a::real_normed_vector set" |
|
1427 |
shows "starlike S \<Longrightarrow> contractible S" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1428 |
using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def) |
69620 | 1429 |
|
1430 |
lemma contractible_UNIV [simp]: "contractible (UNIV :: 'a::real_normed_vector set)" |
|
1431 |
by (simp add: starlike_imp_contractible) |
|
1432 |
||
1433 |
lemma starlike_imp_simply_connected: |
|
1434 |
fixes S :: "'a::real_normed_vector set" |
|
1435 |
shows "starlike S \<Longrightarrow> simply_connected S" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1436 |
by (simp add: contractible_imp_simply_connected starlike_imp_contractible) |
69620 | 1437 |
|
1438 |
lemma convex_imp_simply_connected: |
|
1439 |
fixes S :: "'a::real_normed_vector set" |
|
1440 |
shows "convex S \<Longrightarrow> simply_connected S" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1441 |
using convex_imp_starlike starlike_imp_simply_connected by blast |
69620 | 1442 |
|
1443 |
lemma starlike_imp_path_connected: |
|
1444 |
fixes S :: "'a::real_normed_vector set" |
|
1445 |
shows "starlike S \<Longrightarrow> path_connected S" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1446 |
by (simp add: simply_connected_imp_path_connected starlike_imp_simply_connected) |
69620 | 1447 |
|
1448 |
lemma starlike_imp_connected: |
|
1449 |
fixes S :: "'a::real_normed_vector set" |
|
1450 |
shows "starlike S \<Longrightarrow> connected S" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1451 |
by (simp add: path_connected_imp_connected starlike_imp_path_connected) |
69620 | 1452 |
|
1453 |
lemma is_interval_simply_connected_1: |
|
1454 |
fixes S :: "real set" |
|
1455 |
shows "is_interval S \<longleftrightarrow> simply_connected S" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1456 |
by (meson convex_imp_simply_connected is_interval_connected_1 is_interval_convex_1 simply_connected_imp_connected) |
69620 | 1457 |
|
1458 |
lemma contractible_empty [simp]: "contractible {}" |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1459 |
by (simp add: contractible_def homotopic_on_emptyI) |
69620 | 1460 |
|
1461 |
lemma contractible_convex_tweak_boundary_points: |
|
1462 |
fixes S :: "'a::euclidean_space set" |
|
1463 |
assumes "convex S" and TS: "rel_interior S \<subseteq> T" "T \<subseteq> closure S" |
|
1464 |
shows "contractible T" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1465 |
by (metis assms closure_eq_empty contractible_empty empty_subsetI |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1466 |
starlike_convex_tweak_boundary_points starlike_imp_contractible subset_antisym) |
69620 | 1467 |
|
1468 |
lemma convex_imp_contractible: |
|
1469 |
fixes S :: "'a::real_normed_vector set" |
|
1470 |
shows "convex S \<Longrightarrow> contractible S" |
|
1471 |
using contractible_empty convex_imp_starlike starlike_imp_contractible by blast |
|
1472 |
||
1473 |
lemma contractible_sing [simp]: |
|
1474 |
fixes a :: "'a::real_normed_vector" |
|
1475 |
shows "contractible {a}" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1476 |
by (rule convex_imp_contractible [OF convex_singleton]) |
69620 | 1477 |
|
1478 |
lemma is_interval_contractible_1: |
|
1479 |
fixes S :: "real set" |
|
1480 |
shows "is_interval S \<longleftrightarrow> contractible S" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1481 |
using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1 |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1482 |
is_interval_simply_connected_1 by auto |
69620 | 1483 |
|
1484 |
lemma contractible_Times: |
|
1485 |
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" |
|
1486 |
assumes S: "contractible S" and T: "contractible T" |
|
1487 |
shows "contractible (S \<times> T)" |
|
1488 |
proof - |
|
1489 |
obtain a h where conth: "continuous_on ({0..1} \<times> S) h" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1490 |
and hsub: "h \<in> ({0..1} \<times> S) \<rightarrow> S" |
69620 | 1491 |
and [simp]: "\<And>x. x \<in> S \<Longrightarrow> h (0, x) = x" |
1492 |
and [simp]: "\<And>x. x \<in> S \<Longrightarrow> h (1::real, x) = a" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1493 |
using S by (force simp: contractible_def homotopic_with) |
69620 | 1494 |
obtain b k where contk: "continuous_on ({0..1} \<times> T) k" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1495 |
and ksub: "k \<in> ({0..1} \<times> T) \<rightarrow> T" |
69620 | 1496 |
and [simp]: "\<And>x. x \<in> T \<Longrightarrow> k (0, x) = x" |
1497 |
and [simp]: "\<And>x. x \<in> T \<Longrightarrow> k (1::real, x) = b" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
1498 |
using T by (force simp: contractible_def homotopic_with) |
69620 | 1499 |
show ?thesis |
1500 |
apply (simp add: contractible_def homotopic_with) |
|
1501 |
apply (rule exI [where x=a]) |
|
1502 |
apply (rule exI [where x=b]) |
|
1503 |
apply (rule exI [where x = "\<lambda>z. (h (fst z, fst(snd z)), k (fst z, snd(snd z)))"]) |
|
1504 |
using hsub ksub |
|
71768 | 1505 |
apply (fastforce intro!: continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF contk]) |
69620 | 1506 |
done |
1507 |
qed |
|
1508 |
||
1509 |
||
1510 |
subsection\<open>Local versions of topological properties in general\<close> |
|
1511 |
||
70136 | 1512 |
definition\<^marker>\<open>tag important\<close> locally :: "('a::topological_space set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool" |
69620 | 1513 |
where |
1514 |
"locally P S \<equiv> |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1515 |
\<forall>w x. openin (top_of_set S) w \<and> x \<in> w |
78474 | 1516 |
\<longrightarrow> (\<exists>U V. openin (top_of_set S) U \<and> P V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> w)" |
69620 | 1517 |
|
1518 |
lemma locallyI: |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1519 |
assumes "\<And>w x. \<lbrakk>openin (top_of_set S) w; x \<in> w\<rbrakk> |
78474 | 1520 |
\<Longrightarrow> \<exists>U V. openin (top_of_set S) U \<and> P V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> w" |
69620 | 1521 |
shows "locally P S" |
1522 |
using assms by (force simp: locally_def) |
|
1523 |
||
1524 |
lemma locallyE: |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1525 |
assumes "locally P S" "openin (top_of_set S) w" "x \<in> w" |
78474 | 1526 |
obtains U V where "openin (top_of_set S) U" "P V" "x \<in> U" "U \<subseteq> V" "V \<subseteq> w" |
69620 | 1527 |
using assms unfolding locally_def by meson |
1528 |
||
1529 |
lemma locally_mono: |
|
71769 | 1530 |
assumes "locally P S" "\<And>T. P T \<Longrightarrow> Q T" |
69620 | 1531 |
shows "locally Q S" |
1532 |
by (metis assms locally_def) |
|
1533 |
||
1534 |
lemma locally_open_subset: |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1535 |
assumes "locally P S" "openin (top_of_set S) t" |
69620 | 1536 |
shows "locally P t" |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1537 |
by (smt (verit, ccfv_SIG) assms order.trans locally_def openin_imp_subset openin_subset_trans openin_trans) |
69620 | 1538 |
|
1539 |
lemma locally_diff_closed: |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1540 |
"\<lbrakk>locally P S; closedin (top_of_set S) t\<rbrakk> \<Longrightarrow> locally P (S - t)" |
69620 | 1541 |
using locally_open_subset closedin_def by fastforce |
1542 |
||
1543 |
lemma locally_empty [iff]: "locally P {}" |
|
1544 |
by (simp add: locally_def openin_subtopology) |
|
1545 |
||
1546 |
lemma locally_singleton [iff]: |
|
1547 |
fixes a :: "'a::metric_space" |
|
1548 |
shows "locally P {a} \<longleftrightarrow> P {a}" |
|
71768 | 1549 |
proof - |
1550 |
have "\<forall>x::real. \<not> 0 < x \<Longrightarrow> P {a}" |
|
1551 |
using zero_less_one by blast |
|
1552 |
then show ?thesis |
|
1553 |
unfolding locally_def |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1554 |
by (auto simp: openin_euclidean_subtopology_iff subset_singleton_iff conj_disj_distribR) |
71768 | 1555 |
qed |
69620 | 1556 |
|
1557 |
lemma locally_iff: |
|
1558 |
"locally P S \<longleftrightarrow> |
|
71746 | 1559 |
(\<forall>T x. open T \<and> x \<in> S \<inter> T \<longrightarrow> (\<exists>U. open U \<and> (\<exists>V. P V \<and> x \<in> S \<inter> U \<and> S \<inter> U \<subseteq> V \<and> V \<subseteq> S \<inter> T)))" |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1560 |
by (smt (verit) locally_def openin_open) |
69620 | 1561 |
|
1562 |
lemma locally_Int: |
|
71746 | 1563 |
assumes S: "locally P S" and T: "locally P T" |
1564 |
and P: "\<And>S T. P S \<and> P T \<Longrightarrow> P(S \<inter> T)" |
|
1565 |
shows "locally P (S \<inter> T)" |
|
1566 |
unfolding locally_iff |
|
1567 |
proof clarify |
|
1568 |
fix A x |
|
1569 |
assume "open A" "x \<in> A" "x \<in> S" "x \<in> T" |
|
1570 |
then obtain U1 V1 U2 V2 |
|
1571 |
where "open U1" "P V1" "x \<in> S \<inter> U1" "S \<inter> U1 \<subseteq> V1 \<and> V1 \<subseteq> S \<inter> A" |
|
1572 |
"open U2" "P V2" "x \<in> T \<inter> U2" "T \<inter> U2 \<subseteq> V2 \<and> V2 \<subseteq> T \<inter> A" |
|
1573 |
using S T unfolding locally_iff by (meson IntI) |
|
1574 |
then have "S \<inter> T \<inter> (U1 \<inter> U2) \<subseteq> V1 \<inter> V2" "V1 \<inter> V2 \<subseteq> S \<inter> T \<inter> A" "x \<in> S \<inter> T \<inter> (U1 \<inter> U2)" |
|
1575 |
by blast+ |
|
1576 |
moreover have "P (V1 \<inter> V2)" |
|
1577 |
by (simp add: P \<open>P V1\<close> \<open>P V2\<close>) |
|
1578 |
ultimately show "\<exists>U. open U \<and> (\<exists>V. P V \<and> x \<in> S \<inter> T \<inter> U \<and> S \<inter> T \<inter> U \<subseteq> V \<and> V \<subseteq> S \<inter> T \<inter> A)" |
|
1579 |
using \<open>open U1\<close> \<open>open U2\<close> by blast |
|
1580 |
qed |
|
1581 |
||
69620 | 1582 |
|
1583 |
lemma locally_Times: |
|
1584 |
fixes S :: "('a::metric_space) set" and T :: "('b::metric_space) set" |
|
1585 |
assumes PS: "locally P S" and QT: "locally Q T" and R: "\<And>S T. P S \<and> Q T \<Longrightarrow> R(S \<times> T)" |
|
1586 |
shows "locally R (S \<times> T)" |
|
1587 |
unfolding locally_def |
|
1588 |
proof (clarify) |
|
1589 |
fix W x y |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1590 |
assume W: "openin (top_of_set (S \<times> T)) W" and xy: "(x, y) \<in> W" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1591 |
then obtain U V where "openin (top_of_set S) U" "x \<in> U" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1592 |
"openin (top_of_set T) V" "y \<in> V" "U \<times> V \<subseteq> W" |
69620 | 1593 |
using Times_in_interior_subtopology by metis |
1594 |
then obtain U1 U2 V1 V2 |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1595 |
where opeS: "openin (top_of_set S) U1 \<and> P U2 \<and> x \<in> U1 \<and> U1 \<subseteq> U2 \<and> U2 \<subseteq> U" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1596 |
and opeT: "openin (top_of_set T) V1 \<and> Q V2 \<and> y \<in> V1 \<and> V1 \<subseteq> V2 \<and> V2 \<subseteq> V" |
69620 | 1597 |
by (meson PS QT locallyE) |
71768 | 1598 |
then have "openin (top_of_set (S \<times> T)) (U1 \<times> V1)" |
1599 |
by (simp add: openin_Times) |
|
1600 |
moreover have "R (U2 \<times> V2)" |
|
1601 |
by (simp add: R opeS opeT) |
|
1602 |
moreover have "U1 \<times> V1 \<subseteq> U2 \<times> V2 \<and> U2 \<times> V2 \<subseteq> W" |
|
1603 |
using opeS opeT \<open>U \<times> V \<subseteq> W\<close> by auto |
|
1604 |
ultimately show "\<exists>U V. openin (top_of_set (S \<times> T)) U \<and> R V \<and> (x,y) \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W" |
|
1605 |
using opeS opeT by auto |
|
69620 | 1606 |
qed |
1607 |
||
1608 |
||
1609 |
proposition homeomorphism_locally_imp: |
|
71746 | 1610 |
fixes S :: "'a::metric_space set" and T :: "'b::t2_space set" |
1611 |
assumes S: "locally P S" and hom: "homeomorphism S T f g" |
|
69620 | 1612 |
and Q: "\<And>S S'. \<lbrakk>P S; homeomorphism S S' f g\<rbrakk> \<Longrightarrow> Q S'" |
71746 | 1613 |
shows "locally Q T" |
69620 | 1614 |
proof (clarsimp simp: locally_def) |
1615 |
fix W y |
|
71746 | 1616 |
assume "y \<in> W" and "openin (top_of_set T) W" |
1617 |
then obtain A where T: "open A" "W = T \<inter> A" |
|
69620 | 1618 |
by (force simp: openin_open) |
71746 | 1619 |
then have "W \<subseteq> T" by auto |
1620 |
have f: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" "f ` S = T" "continuous_on S f" |
|
1621 |
and g: "\<And>y. y \<in> T \<Longrightarrow> f(g y) = y" "g ` T = S" "continuous_on T g" |
|
69620 | 1622 |
using hom by (auto simp: homeomorphism_def) |
1623 |
have gw: "g ` W = S \<inter> f -` W" |
|
71746 | 1624 |
using \<open>W \<subseteq> T\<close> g by force |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1625 |
have "openin (top_of_set S) (g ` W)" |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1626 |
using \<open>openin (top_of_set T) W\<close> continuous_on_open f gw by auto |
71746 | 1627 |
then obtain U V |
1628 |
where osu: "openin (top_of_set S) U" and uv: "P V" "g y \<in> U" "U \<subseteq> V" "V \<subseteq> g ` W" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1629 |
by (metis S \<open>y \<in> W\<close> image_eqI locallyE) |
71746 | 1630 |
have "V \<subseteq> S" using uv by (simp add: gw) |
1631 |
have fv: "f ` V = T \<inter> {x. g x \<in> V}" |
|
1632 |
using \<open>f ` S = T\<close> f \<open>V \<subseteq> S\<close> by auto |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1633 |
have contvf: "continuous_on V f" |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1634 |
using \<open>V \<subseteq> S\<close> continuous_on_subset f(3) by blast |
71746 | 1635 |
have "openin (top_of_set (g ` T)) U" |
1636 |
using \<open>g ` T = S\<close> by (simp add: osu) |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1637 |
then have "openin (top_of_set T) (T \<inter> g -` U)" |
71746 | 1638 |
using \<open>continuous_on T g\<close> continuous_on_open [THEN iffD1] by blast |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1639 |
moreover have "\<exists>V. Q V \<and> y \<in> (T \<inter> g -` U) \<and> (T \<inter> g -` U) \<subseteq> V \<and> V \<subseteq> W" |
71768 | 1640 |
proof (intro exI conjI) |
78474 | 1641 |
show "f ` V \<subseteq> W" |
1642 |
using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto |
|
1643 |
then have contvg: "continuous_on (f ` V) g" |
|
1644 |
using \<open>W \<subseteq> T\<close> continuous_on_subset [OF g(3)] by blast |
|
1645 |
have "V \<subseteq> g ` f ` V" |
|
1646 |
by (metis \<open>V \<subseteq> S\<close> hom homeomorphism_def homeomorphism_of_subsets order_refl) |
|
1647 |
then have homv: "homeomorphism V (f ` V) f g" |
|
1648 |
using \<open>V \<subseteq> S\<close> f by (auto simp: homeomorphism_def contvf contvg) |
|
71768 | 1649 |
show "Q (f ` V)" |
1650 |
using Q homv \<open>P V\<close> by blast |
|
1651 |
show "y \<in> T \<inter> g -` U" |
|
1652 |
using T(2) \<open>y \<in> W\<close> \<open>g y \<in> U\<close> by blast |
|
1653 |
show "T \<inter> g -` U \<subseteq> f ` V" |
|
1654 |
using g(1) image_iff uv(3) by fastforce |
|
1655 |
qed |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1656 |
ultimately show "\<exists>U. openin (top_of_set T) U \<and> (\<exists>v. Q v \<and> y \<in> U \<and> U \<subseteq> v \<and> v \<subseteq> W)" |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1657 |
by meson |
69620 | 1658 |
qed |
1659 |
||
1660 |
lemma homeomorphism_locally: |
|
1661 |
fixes f:: "'a::metric_space \<Rightarrow> 'b::metric_space" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1662 |
assumes "homeomorphism S T f g" |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1663 |
and "\<And>S T. homeomorphism S T f g \<Longrightarrow> (P S \<longleftrightarrow> Q T)" |
71746 | 1664 |
shows "locally P S \<longleftrightarrow> locally Q T" |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1665 |
by (smt (verit) assms homeomorphism_locally_imp homeomorphism_symD) |
69620 | 1666 |
|
1667 |
lemma homeomorphic_locally: |
|
1668 |
fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" |
|
1669 |
assumes hom: "S homeomorphic T" |
|
1670 |
and iff: "\<And>X Y. X homeomorphic Y \<Longrightarrow> (P X \<longleftrightarrow> Q Y)" |
|
1671 |
shows "locally P S \<longleftrightarrow> locally Q T" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1672 |
by (smt (verit, ccfv_SIG) hom homeomorphic_def homeomorphism_locally homeomorphism_locally_imp iff) |
69620 | 1673 |
|
1674 |
lemma homeomorphic_local_compactness: |
|
1675 |
fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" |
|
1676 |
shows "S homeomorphic T \<Longrightarrow> locally compact S \<longleftrightarrow> locally compact T" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1677 |
by (simp add: homeomorphic_compactness homeomorphic_locally) |
69620 | 1678 |
|
1679 |
lemma locally_translation: |
|
1680 |
fixes P :: "'a :: real_normed_vector set \<Rightarrow> bool" |
|
71746 | 1681 |
shows "(\<And>S. P ((+) a ` S) = P S) \<Longrightarrow> locally P ((+) a ` S) = locally P S" |
71768 | 1682 |
using homeomorphism_locally [OF homeomorphism_translation] |
1683 |
by (metis (full_types) homeomorphism_image2) |
|
69620 | 1684 |
|
1685 |
lemma locally_injective_linear_image: |
|
1686 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
1687 |
assumes f: "linear f" "inj f" and iff: "\<And>S. P (f ` S) \<longleftrightarrow> Q S" |
|
71746 | 1688 |
shows "locally P (f ` S) \<longleftrightarrow> locally Q S" |
78474 | 1689 |
by (smt (verit) f homeomorphism_image2 homeomorphism_locally iff linear_homeomorphism_image) |
69620 | 1690 |
|
1691 |
lemma locally_open_map_image: |
|
1692 |
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" |
|
1693 |
assumes P: "locally P S" |
|
1694 |
and f: "continuous_on S f" |
|
71746 | 1695 |
and oo: "\<And>T. openin (top_of_set S) T \<Longrightarrow> openin (top_of_set (f ` S)) (f ` T)" |
1696 |
and Q: "\<And>T. \<lbrakk>T \<subseteq> S; P T\<rbrakk> \<Longrightarrow> Q(f ` T)" |
|
69620 | 1697 |
shows "locally Q (f ` S)" |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1698 |
proof (clarsimp simp: locally_def) |
69620 | 1699 |
fix W y |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1700 |
assume oiw: "openin (top_of_set (f ` S)) W" and "y \<in> W" |
69620 | 1701 |
then have "W \<subseteq> f ` S" by (simp add: openin_euclidean_subtopology_iff) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1702 |
have oivf: "openin (top_of_set S) (S \<inter> f -` W)" |
69620 | 1703 |
by (rule continuous_on_open [THEN iffD1, rule_format, OF f oiw]) |
1704 |
then obtain x where "x \<in> S" "f x = y" |
|
1705 |
using \<open>W \<subseteq> f ` S\<close> \<open>y \<in> W\<close> by blast |
|
1706 |
then obtain U V |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1707 |
where "openin (top_of_set S) U" "P V" "x \<in> U" "U \<subseteq> V" "V \<subseteq> S \<inter> f -` W" |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1708 |
by (metis IntI P \<open>y \<in> W\<close> locallyE oivf vimageI) |
71746 | 1709 |
then have "openin (top_of_set (f ` S)) (f ` U)" |
1710 |
by (simp add: oo) |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1711 |
then show "\<exists>X. openin (top_of_set (f ` S)) X \<and> (\<exists>Y. Q Y \<and> y \<in> X \<and> X \<subseteq> Y \<and> Y \<subseteq> W)" |
71746 | 1712 |
using Q \<open>P V\<close> \<open>U \<subseteq> V\<close> \<open>V \<subseteq> S \<inter> f -` W\<close> \<open>f x = y\<close> \<open>x \<in> U\<close> by blast |
69620 | 1713 |
qed |
1714 |
||
1715 |
||
1716 |
subsection\<open>An induction principle for connected sets\<close> |
|
1717 |
||
1718 |
proposition connected_induction: |
|
1719 |
assumes "connected S" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1720 |
and opD: "\<And>T a. \<lbrakk>openin (top_of_set S) T; a \<in> T\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> T \<and> P z" |
69620 | 1721 |
and opI: "\<And>a. a \<in> S |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1722 |
\<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> |
69620 | 1723 |
(\<forall>x \<in> T. \<forall>y \<in> T. P x \<and> P y \<and> Q x \<longrightarrow> Q y)" |
1724 |
and etc: "a \<in> S" "b \<in> S" "P a" "P b" "Q a" |
|
1725 |
shows "Q b" |
|
1726 |
proof - |
|
71746 | 1727 |
let ?A = "{b. \<exists>T. openin (top_of_set S) T \<and> b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> Q x)}" |
1728 |
let ?B = "{b. \<exists>T. openin (top_of_set S) T \<and> b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> \<not> Q x)}" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1729 |
have "?A \<inter> ?B = {}" |
73932
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
72372
diff
changeset
|
1730 |
by (clarsimp simp: set_eq_iff) (metis (no_types, opaque_lifting) Int_iff opD openin_Int) |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1731 |
moreover have "S \<subseteq> ?A \<union> ?B" |
71746 | 1732 |
by clarsimp (meson opI) |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1733 |
moreover have "openin (top_of_set S) ?A" |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1734 |
by (subst openin_subopen, blast) |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1735 |
moreover have "openin (top_of_set S) ?B" |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1736 |
by (subst openin_subopen, blast) |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1737 |
ultimately have "?A = {} \<or> ?B = {}" |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1738 |
by (metis (no_types, lifting) \<open>connected S\<close> connected_openin) |
71746 | 1739 |
then show ?thesis |
1740 |
by clarsimp (meson opI etc) |
|
69620 | 1741 |
qed |
1742 |
||
1743 |
lemma connected_equivalence_relation_gen: |
|
1744 |
assumes "connected S" |
|
1745 |
and etc: "a \<in> S" "b \<in> S" "P a" "P b" |
|
1746 |
and trans: "\<And>x y z. \<lbrakk>R x y; R y z\<rbrakk> \<Longrightarrow> R x z" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1747 |
and opD: "\<And>T a. \<lbrakk>openin (top_of_set S) T; a \<in> T\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> T \<and> P z" |
69620 | 1748 |
and opI: "\<And>a. a \<in> S |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1749 |
\<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> |
69620 | 1750 |
(\<forall>x \<in> T. \<forall>y \<in> T. P x \<and> P y \<longrightarrow> R x y)" |
1751 |
shows "R a b" |
|
1752 |
proof - |
|
1753 |
have "\<And>a b c. \<lbrakk>a \<in> S; P a; b \<in> S; c \<in> S; P b; P c; R a b\<rbrakk> \<Longrightarrow> R a c" |
|
1754 |
apply (rule connected_induction [OF \<open>connected S\<close> opD], simp_all) |
|
1755 |
by (meson trans opI) |
|
1756 |
then show ?thesis by (metis etc opI) |
|
1757 |
qed |
|
1758 |
||
1759 |
lemma connected_induction_simple: |
|
1760 |
assumes "connected S" |
|
1761 |
and etc: "a \<in> S" "b \<in> S" "P a" |
|
1762 |
and opI: "\<And>a. a \<in> S |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1763 |
\<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> |
69620 | 1764 |
(\<forall>x \<in> T. \<forall>y \<in> T. P x \<longrightarrow> P y)" |
1765 |
shows "P b" |
|
72372 | 1766 |
by (rule connected_induction [OF \<open>connected S\<close> _, where P = "\<lambda>x. True"]) |
1767 |
(use opI etc in auto) |
|
69620 | 1768 |
|
1769 |
lemma connected_equivalence_relation: |
|
1770 |
assumes "connected S" |
|
1771 |
and etc: "a \<in> S" "b \<in> S" |
|
1772 |
and sym: "\<And>x y. \<lbrakk>R x y; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> R y x" |
|
1773 |
and trans: "\<And>x y z. \<lbrakk>R x y; R y z; x \<in> S; y \<in> S; z \<in> S\<rbrakk> \<Longrightarrow> R x z" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1774 |
and opI: "\<And>a. a \<in> S \<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> (\<forall>x \<in> T. R a x)" |
69620 | 1775 |
shows "R a b" |
1776 |
proof - |
|
1777 |
have "\<And>a b c. \<lbrakk>a \<in> S; b \<in> S; c \<in> S; R a b\<rbrakk> \<Longrightarrow> R a c" |
|
78474 | 1778 |
by (smt (verit, ccfv_threshold) connected_induction_simple [OF \<open>connected S\<close>] |
1779 |
assms openin_imp_subset subset_eq) |
|
69620 | 1780 |
then show ?thesis by (metis etc opI) |
1781 |
qed |
|
1782 |
||
1783 |
lemma locally_constant_imp_constant: |
|
1784 |
assumes "connected S" |
|
1785 |
and opI: "\<And>a. a \<in> S |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
1786 |
\<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> (\<forall>x \<in> T. f x = f a)" |
69620 | 1787 |
shows "f constant_on S" |
1788 |
proof - |
|
1789 |
have "\<And>x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> f x = f y" |
|
1790 |
apply (rule connected_equivalence_relation [OF \<open>connected S\<close>], simp_all) |
|
1791 |
by (metis opI) |
|
1792 |
then show ?thesis |
|
1793 |
by (metis constant_on_def) |
|
1794 |
qed |
|
1795 |
||
1796 |
lemma locally_constant: |
|
71768 | 1797 |
assumes "connected S" |
1798 |
shows "locally (\<lambda>U. f constant_on U) S \<longleftrightarrow> f constant_on S" (is "?lhs = ?rhs") |
|
1799 |
proof |
|
1800 |
assume ?lhs |
|
1801 |
then show ?rhs |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1802 |
by (smt (verit, del_insts) assms constant_on_def locally_constant_imp_constant locally_def openin_subtopology_self subset_iff) |
71768 | 1803 |
next |
1804 |
assume ?rhs then show ?lhs |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1805 |
by (metis constant_on_subset locallyI openin_imp_subset order_refl) |
71768 | 1806 |
qed |
69620 | 1807 |
|
1808 |
||
1809 |
subsection\<open>Basic properties of local compactness\<close> |
|
1810 |
||
1811 |
proposition locally_compact: |
|
77684 | 1812 |
fixes S :: "'a :: metric_space set" |
69620 | 1813 |
shows |
77684 | 1814 |
"locally compact S \<longleftrightarrow> |
1815 |
(\<forall>x \<in> S. \<exists>u v. x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> S \<and> |
|
1816 |
openin (top_of_set S) u \<and> compact v)" |
|
69620 | 1817 |
(is "?lhs = ?rhs") |
1818 |
proof |
|
1819 |
assume ?lhs |
|
1820 |
then show ?rhs |
|
71768 | 1821 |
by (meson locallyE openin_subtopology_self) |
69620 | 1822 |
next |
1823 |
assume r [rule_format]: ?rhs |
|
1824 |
have *: "\<exists>u v. |
|
77684 | 1825 |
openin (top_of_set S) u \<and> |
1826 |
compact v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> S \<inter> T" |
|
1827 |
if "open T" "x \<in> S" "x \<in> T" for x T |
|
69620 | 1828 |
proof - |
78474 | 1829 |
obtain U V where uv: "x \<in> U" "U \<subseteq> V" "V \<subseteq> S" "compact V" "openin (top_of_set S) U" |
77684 | 1830 |
using r [OF \<open>x \<in> S\<close>] by auto |
69620 | 1831 |
obtain e where "e>0" and e: "cball x e \<subseteq> T" |
1832 |
using open_contains_cball \<open>open T\<close> \<open>x \<in> T\<close> by blast |
|
1833 |
show ?thesis |
|
78474 | 1834 |
apply (rule_tac x="(S \<inter> ball x e) \<inter> U" in exI) |
1835 |
apply (rule_tac x="cball x e \<inter> V" in exI) |
|
69620 | 1836 |
using that \<open>e > 0\<close> e uv |
1837 |
apply auto |
|
1838 |
done |
|
1839 |
qed |
|
1840 |
show ?lhs |
|
71768 | 1841 |
by (rule locallyI) (metis "*" Int_iff openin_open) |
69620 | 1842 |
qed |
1843 |
||
1844 |
lemma locally_compactE: |
|
71746 | 1845 |
fixes S :: "'a :: metric_space set" |
1846 |
assumes "locally compact S" |
|
1847 |
obtains u v where "\<And>x. x \<in> S \<Longrightarrow> x \<in> u x \<and> u x \<subseteq> v x \<and> v x \<subseteq> S \<and> |
|
1848 |
openin (top_of_set S) (u x) \<and> compact (v x)" |
|
1849 |
using assms unfolding locally_compact by metis |
|
69620 | 1850 |
|
1851 |
lemma locally_compact_alt: |
|
71746 | 1852 |
fixes S :: "'a :: heine_borel set" |
1853 |
shows "locally compact S \<longleftrightarrow> |
|
71768 | 1854 |
(\<forall>x \<in> S. \<exists>U. x \<in> U \<and> |
1855 |
openin (top_of_set S) U \<and> compact(closure U) \<and> closure U \<subseteq> S)" |
|
78474 | 1856 |
by (smt (verit, ccfv_threshold) bounded_subset closure_closed closure_mono closure_subset |
1857 |
compact_closure compact_imp_closed order.trans locally_compact) |
|
69620 | 1858 |
|
1859 |
lemma locally_compact_Int_cball: |
|
71746 | 1860 |
fixes S :: "'a :: heine_borel set" |
1861 |
shows "locally compact S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>e. 0 < e \<and> closed(cball x e \<inter> S))" |
|
69620 | 1862 |
(is "?lhs = ?rhs") |
1863 |
proof |
|
72372 | 1864 |
assume L: ?lhs |
1865 |
then have "\<And>x U V e. \<lbrakk>U \<subseteq> V; V \<subseteq> S; compact V; 0 < e; cball x e \<inter> S \<subseteq> U\<rbrakk> |
|
1866 |
\<Longrightarrow> closed (cball x e \<inter> S)" |
|
1867 |
by (metis compact_Int compact_cball compact_imp_closed inf.absorb_iff2 inf.assoc inf.orderE) |
|
1868 |
with L show ?rhs |
|
1869 |
by (meson locally_compactE openin_contains_cball) |
|
69620 | 1870 |
next |
72372 | 1871 |
assume R: ?rhs |
1872 |
show ?lhs unfolding locally_compact |
|
1873 |
proof |
|
1874 |
fix x |
|
1875 |
assume "x \<in> S" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1876 |
then obtain e where "e>0" and "compact (cball x e \<inter> S)" |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1877 |
by (metis Int_commute compact_Int_closed compact_cball inf.right_idem R) |
72372 | 1878 |
moreover have "\<forall>y\<in>ball x e \<inter> S. \<exists>\<epsilon>>0. cball y \<epsilon> \<inter> S \<subseteq> ball x e" |
1879 |
by (meson Elementary_Metric_Spaces.open_ball IntD1 le_infI1 open_contains_cball_eq) |
|
1880 |
moreover have "openin (top_of_set S) (ball x e \<inter> S)" |
|
1881 |
by (simp add: inf_commute openin_open_Int) |
|
1882 |
ultimately show "\<exists>U V. x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> S \<and> openin (top_of_set S) U \<and> compact V" |
|
1883 |
by (metis Int_iff \<open>0 < e\<close> \<open>x \<in> S\<close> ball_subset_cball centre_in_ball inf_commute inf_le1 inf_mono order_refl) |
|
1884 |
qed |
|
69620 | 1885 |
qed |
1886 |
||
1887 |
lemma locally_compact_compact: |
|
71746 | 1888 |
fixes S :: "'a :: heine_borel set" |
1889 |
shows "locally compact S \<longleftrightarrow> |
|
72372 | 1890 |
(\<forall>K. K \<subseteq> S \<and> compact K |
1891 |
\<longrightarrow> (\<exists>U V. K \<subseteq> U \<and> U \<subseteq> V \<and> V \<subseteq> S \<and> |
|
1892 |
openin (top_of_set S) U \<and> compact V))" |
|
69620 | 1893 |
(is "?lhs = ?rhs") |
1894 |
proof |
|
1895 |
assume ?lhs |
|
1896 |
then obtain u v where |
|
71746 | 1897 |
uv: "\<And>x. x \<in> S \<Longrightarrow> x \<in> u x \<and> u x \<subseteq> v x \<and> v x \<subseteq> S \<and> |
1898 |
openin (top_of_set S) (u x) \<and> compact (v x)" |
|
69620 | 1899 |
by (metis locally_compactE) |
72372 | 1900 |
have *: "\<exists>U V. K \<subseteq> U \<and> U \<subseteq> V \<and> V \<subseteq> S \<and> openin (top_of_set S) U \<and> compact V" |
1901 |
if "K \<subseteq> S" "compact K" for K |
|
69620 | 1902 |
proof - |
72372 | 1903 |
have "\<And>C. (\<forall>c\<in>C. openin (top_of_set K) c) \<and> K \<subseteq> \<Union>C \<Longrightarrow> |
1904 |
\<exists>D\<subseteq>C. finite D \<and> K \<subseteq> \<Union>D" |
|
69620 | 1905 |
using that by (simp add: compact_eq_openin_cover) |
72372 | 1906 |
moreover have "\<forall>c \<in> (\<lambda>x. K \<inter> u x) ` K. openin (top_of_set K) c" |
69620 | 1907 |
using that by clarify (metis subsetD inf.absorb_iff2 openin_subset openin_subtopology_Int_subset topspace_euclidean_subtopology uv) |
72372 | 1908 |
moreover have "K \<subseteq> \<Union>((\<lambda>x. K \<inter> u x) ` K)" |
69620 | 1909 |
using that by clarsimp (meson subsetCE uv) |
72372 | 1910 |
ultimately obtain D where "D \<subseteq> (\<lambda>x. K \<inter> u x) ` K" "finite D" "K \<subseteq> \<Union>D" |
69620 | 1911 |
by metis |
72372 | 1912 |
then obtain T where T: "T \<subseteq> K" "finite T" "K \<subseteq> \<Union>((\<lambda>x. K \<inter> u x) ` T)" |
69620 | 1913 |
by (metis finite_subset_image) |
1914 |
have Tuv: "\<Union>(u ` T) \<subseteq> \<Union>(v ` T)" |
|
72372 | 1915 |
using T that by (force dest!: uv) |
1916 |
moreover |
|
1917 |
have "openin (top_of_set S) (\<Union> (u ` T))" |
|
1918 |
using T that uv by fastforce |
|
1919 |
moreover |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1920 |
obtain "compact (\<Union> (v ` T))" "\<Union> (v ` T) \<subseteq> S" |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1921 |
by (metis T UN_subset_iff \<open>K \<subseteq> S\<close> compact_UN subset_iff uv) |
72372 | 1922 |
ultimately show ?thesis |
1923 |
using T by auto |
|
69620 | 1924 |
qed |
1925 |
show ?rhs |
|
1926 |
by (blast intro: *) |
|
1927 |
next |
|
1928 |
assume ?rhs |
|
1929 |
then show ?lhs |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1930 |
apply (clarsimp simp: locally_compact) |
69620 | 1931 |
apply (drule_tac x="{x}" in spec, simp) |
1932 |
done |
|
1933 |
qed |
|
1934 |
||
1935 |
lemma open_imp_locally_compact: |
|
71746 | 1936 |
fixes S :: "'a :: heine_borel set" |
1937 |
assumes "open S" |
|
1938 |
shows "locally compact S" |
|
69620 | 1939 |
proof - |
71746 | 1940 |
have *: "\<exists>U V. x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> S \<and> openin (top_of_set S) U \<and> compact V" |
1941 |
if "x \<in> S" for x |
|
69620 | 1942 |
proof - |
71746 | 1943 |
obtain e where "e>0" and e: "cball x e \<subseteq> S" |
1944 |
using open_contains_cball assms \<open>x \<in> S\<close> by blast |
|
1945 |
have ope: "openin (top_of_set S) (ball x e)" |
|
69620 | 1946 |
by (meson e open_ball ball_subset_cball dual_order.trans open_subset) |
1947 |
show ?thesis |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1948 |
by (meson \<open>0 < e\<close> ball_subset_cball centre_in_ball compact_cball e ope) |
69620 | 1949 |
qed |
1950 |
show ?thesis |
|
71746 | 1951 |
unfolding locally_compact by (blast intro: *) |
69620 | 1952 |
qed |
1953 |
||
1954 |
lemma closed_imp_locally_compact: |
|
71746 | 1955 |
fixes S :: "'a :: heine_borel set" |
1956 |
assumes "closed S" |
|
1957 |
shows "locally compact S" |
|
69620 | 1958 |
proof - |
71746 | 1959 |
have *: "\<exists>U V. x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> S \<and> openin (top_of_set S) U \<and> compact V" |
1960 |
if "x \<in> S" for x |
|
72372 | 1961 |
apply (rule_tac x = "S \<inter> ball x 1" in exI, rule_tac x = "S \<inter> cball x 1" in exI) |
1962 |
using \<open>x \<in> S\<close> assms by auto |
|
69620 | 1963 |
show ?thesis |
72372 | 1964 |
unfolding locally_compact by (blast intro: *) |
69620 | 1965 |
qed |
1966 |
||
1967 |
lemma locally_compact_UNIV: "locally compact (UNIV :: 'a :: heine_borel set)" |
|
1968 |
by (simp add: closed_imp_locally_compact) |
|
1969 |
||
1970 |
lemma locally_compact_Int: |
|
71746 | 1971 |
fixes S :: "'a :: t2_space set" |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1972 |
shows "\<lbrakk>locally compact S; locally compact T\<rbrakk> \<Longrightarrow> locally compact (S \<inter> T)" |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1973 |
by (simp add: compact_Int locally_Int) |
69620 | 1974 |
|
1975 |
lemma locally_compact_closedin: |
|
71746 | 1976 |
fixes S :: "'a :: heine_borel set" |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1977 |
shows "\<lbrakk>closedin (top_of_set S) T; locally compact S\<rbrakk> |
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
1978 |
\<Longrightarrow> locally compact T" |
71746 | 1979 |
unfolding closedin_closed |
1980 |
using closed_imp_locally_compact locally_compact_Int by blast |
|
69620 | 1981 |
|
1982 |
lemma locally_compact_delete: |
|
71746 | 1983 |
fixes S :: "'a :: t1_space set" |
1984 |
shows "locally compact S \<Longrightarrow> locally compact (S - {a})" |
|
69620 | 1985 |
by (auto simp: openin_delete locally_open_subset) |
1986 |
||
1987 |
lemma locally_closed: |
|
71746 | 1988 |
fixes S :: "'a :: heine_borel set" |
1989 |
shows "locally closed S \<longleftrightarrow> locally compact S" |
|
69620 | 1990 |
(is "?lhs = ?rhs") |
1991 |
proof |
|
1992 |
assume ?lhs |
|
1993 |
then show ?rhs |
|
72372 | 1994 |
unfolding locally_def |
1995 |
apply (elim all_forward imp_forward asm_rl exE) |
|
78474 | 1996 |
apply (rename_tac U V) |
1997 |
apply (rule_tac x = "U \<inter> ball x 1" in exI) |
|
1998 |
apply (rule_tac x = "V \<inter> cball x 1" in exI) |
|
69620 | 1999 |
apply (force intro: openin_trans) |
2000 |
done |
|
2001 |
next |
|
2002 |
assume ?rhs then show ?lhs |
|
2003 |
using compact_eq_bounded_closed locally_mono by blast |
|
2004 |
qed |
|
2005 |
||
2006 |
lemma locally_compact_openin_Un: |
|
2007 |
fixes S :: "'a::euclidean_space set" |
|
78474 | 2008 |
assumes LCS: "locally compact S" and LCT: "locally compact T" |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2009 |
and opS: "openin (top_of_set (S \<union> T)) S" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2010 |
and opT: "openin (top_of_set (S \<union> T)) T" |
69620 | 2011 |
shows "locally compact (S \<union> T)" |
2012 |
proof - |
|
2013 |
have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> S" for x |
|
2014 |
proof - |
|
2015 |
obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)" |
|
2016 |
using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast |
|
2017 |
moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \<inter> (S \<union> T) \<subseteq> S" |
|
2018 |
by (meson \<open>x \<in> S\<close> opS openin_contains_cball) |
|
2019 |
then have "cball x e2 \<inter> (S \<union> T) = cball x e2 \<inter> S" |
|
2020 |
by force |
|
71768 | 2021 |
ultimately have "closed (cball x (min e1 e2) \<inter> (S \<union> T))" |
2022 |
by (metis (no_types, lifting) cball_min_Int closed_Int closed_cball inf_assoc inf_commute) |
|
2023 |
then show ?thesis |
|
2024 |
by (metis \<open>0 < e1\<close> \<open>0 < e2\<close> min_def) |
|
69620 | 2025 |
qed |
2026 |
moreover have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> T" for x |
|
2027 |
proof - |
|
2028 |
obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> T)" |
|
2029 |
using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast |
|
2030 |
moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \<inter> (S \<union> T) \<subseteq> T" |
|
2031 |
by (meson \<open>x \<in> T\<close> opT openin_contains_cball) |
|
2032 |
then have "cball x e2 \<inter> (S \<union> T) = cball x e2 \<inter> T" |
|
2033 |
by force |
|
72372 | 2034 |
moreover have "closed (cball x e1 \<inter> (cball x e2 \<inter> T))" |
2035 |
by (metis closed_Int closed_cball e1 inf_left_commute) |
|
69620 | 2036 |
ultimately show ?thesis |
72372 | 2037 |
by (rule_tac x="min e1 e2" in exI) (simp add: \<open>0 < e2\<close> cball_min_Int inf_assoc) |
69620 | 2038 |
qed |
2039 |
ultimately show ?thesis |
|
2040 |
by (force simp: locally_compact_Int_cball) |
|
2041 |
qed |
|
2042 |
||
2043 |
lemma locally_compact_closedin_Un: |
|
2044 |
fixes S :: "'a::euclidean_space set" |
|
2045 |
assumes LCS: "locally compact S" and LCT:"locally compact T" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2046 |
and clS: "closedin (top_of_set (S \<union> T)) S" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2047 |
and clT: "closedin (top_of_set (S \<union> T)) T" |
69620 | 2048 |
shows "locally compact (S \<union> T)" |
2049 |
proof - |
|
2050 |
have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> S" "x \<in> T" for x |
|
2051 |
proof - |
|
2052 |
obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)" |
|
2053 |
using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast |
|
2054 |
moreover |
|
2055 |
obtain e2 where "e2 > 0" and e2: "closed (cball x e2 \<inter> T)" |
|
2056 |
using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast |
|
72372 | 2057 |
moreover have "closed (cball x (min e1 e2) \<inter> (S \<union> T))" |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
2058 |
by (smt (verit) Int_Un_distrib2 Int_commute cball_min_Int closed_Int closed_Un closed_cball e1 e2 inf_left_commute) |
69620 | 2059 |
ultimately show ?thesis |
72372 | 2060 |
by (rule_tac x="min e1 e2" in exI) linarith |
69620 | 2061 |
qed |
2062 |
moreover |
|
2063 |
have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if x: "x \<in> S" "x \<notin> T" for x |
|
2064 |
proof - |
|
2065 |
obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)" |
|
2066 |
using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast |
|
2067 |
moreover |
|
2068 |
obtain e2 where "e2>0" and "cball x e2 \<inter> (S \<union> T) \<subseteq> S - T" |
|
2069 |
using clT x by (fastforce simp: openin_contains_cball closedin_def) |
|
2070 |
then have "closed (cball x e2 \<inter> T)" |
|
2071 |
proof - |
|
2072 |
have "{} = T - (T - cball x e2)" |
|
2073 |
using Diff_subset Int_Diff \<open>cball x e2 \<inter> (S \<union> T) \<subseteq> S - T\<close> by auto |
|
2074 |
then show ?thesis |
|
2075 |
by (simp add: Diff_Diff_Int inf_commute) |
|
2076 |
qed |
|
72372 | 2077 |
with e1 have "closed ((cball x e1 \<inter> cball x e2) \<inter> (S \<union> T))" |
2078 |
apply (simp add: inf_commute inf_sup_distrib2) |
|
2079 |
by (metis closed_Int closed_Un closed_cball inf_assoc inf_left_commute) |
|
2080 |
then have "closed (cball x (min e1 e2) \<inter> (S \<union> T))" |
|
2081 |
by (simp add: cball_min_Int inf_commute) |
|
69620 | 2082 |
ultimately show ?thesis |
72372 | 2083 |
using \<open>0 < e2\<close> by (rule_tac x="min e1 e2" in exI) linarith |
69620 | 2084 |
qed |
2085 |
moreover |
|
2086 |
have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if x: "x \<notin> S" "x \<in> T" for x |
|
2087 |
proof - |
|
2088 |
obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> T)" |
|
2089 |
using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast |
|
2090 |
moreover |
|
2091 |
obtain e2 where "e2>0" and "cball x e2 \<inter> (S \<union> T) \<subseteq> S \<union> T - S" |
|
2092 |
using clS x by (fastforce simp: openin_contains_cball closedin_def) |
|
2093 |
then have "closed (cball x e2 \<inter> S)" |
|
2094 |
by (metis Diff_disjoint Int_empty_right closed_empty inf.left_commute inf.orderE inf_sup_absorb) |
|
72372 | 2095 |
with e1 have "closed ((cball x e1 \<inter> cball x e2) \<inter> (S \<union> T))" |
2096 |
apply (simp add: inf_commute inf_sup_distrib2) |
|
2097 |
by (metis closed_Int closed_Un closed_cball inf_assoc inf_left_commute) |
|
2098 |
then have "closed (cball x (min e1 e2) \<inter> (S \<union> T))" |
|
2099 |
by (auto simp: cball_min_Int) |
|
69620 | 2100 |
ultimately show ?thesis |
72372 | 2101 |
using \<open>0 < e2\<close> by (rule_tac x="min e1 e2" in exI) linarith |
69620 | 2102 |
qed |
2103 |
ultimately show ?thesis |
|
2104 |
by (auto simp: locally_compact_Int_cball) |
|
2105 |
qed |
|
2106 |
||
2107 |
lemma locally_compact_Times: |
|
2108 |
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" |
|
2109 |
shows "\<lbrakk>locally compact S; locally compact T\<rbrakk> \<Longrightarrow> locally compact (S \<times> T)" |
|
2110 |
by (auto simp: compact_Times locally_Times) |
|
2111 |
||
2112 |
lemma locally_compact_compact_subopen: |
|
2113 |
fixes S :: "'a :: heine_borel set" |
|
2114 |
shows |
|
2115 |
"locally compact S \<longleftrightarrow> |
|
2116 |
(\<forall>K T. K \<subseteq> S \<and> compact K \<and> open T \<and> K \<subseteq> T |
|
2117 |
\<longrightarrow> (\<exists>U V. K \<subseteq> U \<and> U \<subseteq> V \<and> U \<subseteq> T \<and> V \<subseteq> S \<and> |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2118 |
openin (top_of_set S) U \<and> compact V))" |
69620 | 2119 |
(is "?lhs = ?rhs") |
2120 |
proof |
|
2121 |
assume L: ?lhs |
|
2122 |
show ?rhs |
|
2123 |
proof clarify |
|
2124 |
fix K :: "'a set" and T :: "'a set" |
|
2125 |
assume "K \<subseteq> S" and "compact K" and "open T" and "K \<subseteq> T" |
|
2126 |
obtain U V where "K \<subseteq> U" "U \<subseteq> V" "V \<subseteq> S" "compact V" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2127 |
and ope: "openin (top_of_set S) U" |
69620 | 2128 |
using L unfolding locally_compact_compact by (meson \<open>K \<subseteq> S\<close> \<open>compact K\<close>) |
2129 |
show "\<exists>U V. K \<subseteq> U \<and> U \<subseteq> V \<and> U \<subseteq> T \<and> V \<subseteq> S \<and> |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2130 |
openin (top_of_set S) U \<and> compact V" |
69620 | 2131 |
proof (intro exI conjI) |
2132 |
show "K \<subseteq> U \<inter> T" |
|
2133 |
by (simp add: \<open>K \<subseteq> T\<close> \<open>K \<subseteq> U\<close>) |
|
2134 |
show "U \<inter> T \<subseteq> closure(U \<inter> T)" |
|
2135 |
by (rule closure_subset) |
|
2136 |
show "closure (U \<inter> T) \<subseteq> S" |
|
2137 |
by (metis \<open>U \<subseteq> V\<close> \<open>V \<subseteq> S\<close> \<open>compact V\<close> closure_closed closure_mono compact_imp_closed inf.cobounded1 subset_trans) |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2138 |
show "openin (top_of_set S) (U \<inter> T)" |
69620 | 2139 |
by (simp add: \<open>open T\<close> ope openin_Int_open) |
2140 |
show "compact (closure (U \<inter> T))" |
|
2141 |
by (meson Int_lower1 \<open>U \<subseteq> V\<close> \<open>compact V\<close> bounded_subset compact_closure compact_eq_bounded_closed) |
|
2142 |
qed auto |
|
2143 |
qed |
|
2144 |
next |
|
2145 |
assume ?rhs then show ?lhs |
|
2146 |
unfolding locally_compact_compact |
|
2147 |
by (metis open_openin openin_topspace subtopology_superset top.extremum topspace_euclidean_subtopology) |
|
2148 |
qed |
|
2149 |
||
2150 |
||
2151 |
subsection\<open>Sura-Bura's results about compact components of sets\<close> |
|
2152 |
||
2153 |
proposition Sura_Bura_compact: |
|
2154 |
fixes S :: "'a::euclidean_space set" |
|
2155 |
assumes "compact S" and C: "C \<in> components S" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2156 |
shows "C = \<Inter>{T. C \<subseteq> T \<and> openin (top_of_set S) T \<and> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2157 |
closedin (top_of_set S) T}" |
69620 | 2158 |
(is "C = \<Inter>?\<T>") |
2159 |
proof |
|
2160 |
obtain x where x: "C = connected_component_set S x" and "x \<in> S" |
|
2161 |
using C by (auto simp: components_def) |
|
2162 |
have "C \<subseteq> S" |
|
2163 |
by (simp add: C in_components_subset) |
|
2164 |
have "\<Inter>?\<T> \<subseteq> connected_component_set S x" |
|
2165 |
proof (rule connected_component_maximal) |
|
2166 |
have "x \<in> C" |
|
2167 |
by (simp add: \<open>x \<in> S\<close> x) |
|
2168 |
then show "x \<in> \<Inter>?\<T>" |
|
2169 |
by blast |
|
2170 |
have clo: "closed (\<Inter>?\<T>)" |
|
2171 |
by (simp add: \<open>compact S\<close> closed_Inter closedin_compact_eq compact_imp_closed) |
|
2172 |
have False |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2173 |
if K1: "closedin (top_of_set (\<Inter>?\<T>)) K1" and |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2174 |
K2: "closedin (top_of_set (\<Inter>?\<T>)) K2" and |
69620 | 2175 |
K12_Int: "K1 \<inter> K2 = {}" and K12_Un: "K1 \<union> K2 = \<Inter>?\<T>" and "K1 \<noteq> {}" "K2 \<noteq> {}" |
2176 |
for K1 K2 |
|
2177 |
proof - |
|
2178 |
have "closed K1" "closed K2" |
|
2179 |
using closedin_closed_trans clo K1 K2 by blast+ |
|
2180 |
then obtain V1 V2 where "open V1" "open V2" "K1 \<subseteq> V1" "K2 \<subseteq> V2" and V12: "V1 \<inter> V2 = {}" |
|
2181 |
using separation_normal \<open>K1 \<inter> K2 = {}\<close> by metis |
|
2182 |
have SV12_ne: "(S - (V1 \<union> V2)) \<inter> (\<Inter>?\<T>) \<noteq> {}" |
|
2183 |
proof (rule compact_imp_fip) |
|
2184 |
show "compact (S - (V1 \<union> V2))" |
|
2185 |
by (simp add: \<open>open V1\<close> \<open>open V2\<close> \<open>compact S\<close> compact_diff open_Un) |
|
2186 |
show clo\<T>: "closed T" if "T \<in> ?\<T>" for T |
|
2187 |
using that \<open>compact S\<close> |
|
2188 |
by (force intro: closedin_closed_trans simp add: compact_imp_closed) |
|
2189 |
show "(S - (V1 \<union> V2)) \<inter> \<Inter>\<F> \<noteq> {}" if "finite \<F>" and \<F>: "\<F> \<subseteq> ?\<T>" for \<F> |
|
2190 |
proof |
|
2191 |
assume djo: "(S - (V1 \<union> V2)) \<inter> \<Inter>\<F> = {}" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2192 |
obtain D where opeD: "openin (top_of_set S) D" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2193 |
and cloD: "closedin (top_of_set S) D" |
69620 | 2194 |
and "C \<subseteq> D" and DV12: "D \<subseteq> V1 \<union> V2" |
2195 |
proof (cases "\<F> = {}") |
|
2196 |
case True |
|
2197 |
with \<open>C \<subseteq> S\<close> djo that show ?thesis |
|
2198 |
by force |
|
2199 |
next |
|
2200 |
case False show ?thesis |
|
2201 |
proof |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2202 |
show ope: "openin (top_of_set S) (\<Inter>\<F>)" |
69620 | 2203 |
using openin_Inter \<open>finite \<F>\<close> False \<F> by blast |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2204 |
then show "closedin (top_of_set S) (\<Inter>\<F>)" |
69620 | 2205 |
by (meson clo\<T> \<F> closed_Inter closed_subset openin_imp_subset subset_eq) |
2206 |
show "C \<subseteq> \<Inter>\<F>" |
|
2207 |
using \<F> by auto |
|
2208 |
show "\<Inter>\<F> \<subseteq> V1 \<union> V2" |
|
2209 |
using ope djo openin_imp_subset by fastforce |
|
2210 |
qed |
|
2211 |
qed |
|
2212 |
have "connected C" |
|
2213 |
by (simp add: x) |
|
2214 |
have "closed D" |
|
2215 |
using \<open>compact S\<close> cloD closedin_closed_trans compact_imp_closed by blast |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2216 |
have cloV1: "closedin (top_of_set D) (D \<inter> closure V1)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2217 |
and cloV2: "closedin (top_of_set D) (D \<inter> closure V2)" |
69620 | 2218 |
by (simp_all add: closedin_closed_Int) |
2219 |
moreover have "D \<inter> closure V1 = D \<inter> V1" "D \<inter> closure V2 = D \<inter> V2" |
|
2220 |
using \<open>D \<subseteq> V1 \<union> V2\<close> \<open>open V1\<close> \<open>open V2\<close> V12 |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
2221 |
by (auto simp: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2222 |
ultimately have cloDV1: "closedin (top_of_set D) (D \<inter> V1)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2223 |
and cloDV2: "closedin (top_of_set D) (D \<inter> V2)" |
69620 | 2224 |
by metis+ |
2225 |
then obtain U1 U2 where "closed U1" "closed U2" |
|
2226 |
and D1: "D \<inter> V1 = D \<inter> U1" and D2: "D \<inter> V2 = D \<inter> U2" |
|
2227 |
by (auto simp: closedin_closed) |
|
2228 |
have "D \<inter> U1 \<inter> C \<noteq> {}" |
|
2229 |
proof |
|
2230 |
assume "D \<inter> U1 \<inter> C = {}" |
|
2231 |
then have *: "C \<subseteq> D \<inter> V2" |
|
2232 |
using D1 DV12 \<open>C \<subseteq> D\<close> by auto |
|
71746 | 2233 |
have 1: "openin (top_of_set S) (D \<inter> V2)" |
2234 |
by (simp add: \<open>open V2\<close> opeD openin_Int_open) |
|
2235 |
have 2: "closedin (top_of_set S) (D \<inter> V2)" |
|
2236 |
using cloD cloDV2 closedin_trans by blast |
|
2237 |
have "\<Inter> ?\<T> \<subseteq> D \<inter> V2" |
|
2238 |
by (rule Inter_lower) (use * 1 2 in simp) |
|
69620 | 2239 |
then show False |
2240 |
using K1 V12 \<open>K1 \<noteq> {}\<close> \<open>K1 \<subseteq> V1\<close> closedin_imp_subset by blast |
|
2241 |
qed |
|
2242 |
moreover have "D \<inter> U2 \<inter> C \<noteq> {}" |
|
2243 |
proof |
|
2244 |
assume "D \<inter> U2 \<inter> C = {}" |
|
2245 |
then have *: "C \<subseteq> D \<inter> V1" |
|
2246 |
using D2 DV12 \<open>C \<subseteq> D\<close> by auto |
|
71746 | 2247 |
have 1: "openin (top_of_set S) (D \<inter> V1)" |
2248 |
by (simp add: \<open>open V1\<close> opeD openin_Int_open) |
|
2249 |
have 2: "closedin (top_of_set S) (D \<inter> V1)" |
|
2250 |
using cloD cloDV1 closedin_trans by blast |
|
69620 | 2251 |
have "\<Inter>?\<T> \<subseteq> D \<inter> V1" |
71746 | 2252 |
by (rule Inter_lower) (use * 1 2 in simp) |
69620 | 2253 |
then show False |
2254 |
using K2 V12 \<open>K2 \<noteq> {}\<close> \<open>K2 \<subseteq> V2\<close> closedin_imp_subset by blast |
|
2255 |
qed |
|
2256 |
ultimately show False |
|
71746 | 2257 |
using \<open>connected C\<close> [unfolded connected_closed, simplified, rule_format, of concl: "D \<inter> U1" "D \<inter> U2"] |
69620 | 2258 |
using \<open>C \<subseteq> D\<close> D1 D2 V12 DV12 \<open>closed U1\<close> \<open>closed U2\<close> \<open>closed D\<close> |
2259 |
by blast |
|
2260 |
qed |
|
2261 |
qed |
|
2262 |
show False |
|
2263 |
by (metis (full_types) DiffE UnE Un_upper2 SV12_ne \<open>K1 \<subseteq> V1\<close> \<open>K2 \<subseteq> V2\<close> disjoint_iff_not_equal subsetCE sup_ge1 K12_Un) |
|
2264 |
qed |
|
2265 |
then show "connected (\<Inter>?\<T>)" |
|
2266 |
by (auto simp: connected_closedin_eq) |
|
2267 |
show "\<Inter>?\<T> \<subseteq> S" |
|
2268 |
by (fastforce simp: C in_components_subset) |
|
2269 |
qed |
|
2270 |
with x show "\<Inter>?\<T> \<subseteq> C" by simp |
|
2271 |
qed auto |
|
2272 |
||
2273 |
||
2274 |
corollary Sura_Bura_clopen_subset: |
|
2275 |
fixes S :: "'a::euclidean_space set" |
|
2276 |
assumes S: "locally compact S" and C: "C \<in> components S" and "compact C" |
|
2277 |
and U: "open U" "C \<subseteq> U" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2278 |
obtains K where "openin (top_of_set S) K" "compact K" "C \<subseteq> K" "K \<subseteq> U" |
69620 | 2279 |
proof (rule ccontr) |
2280 |
assume "\<not> thesis" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2281 |
with that have neg: "\<nexists>K. openin (top_of_set S) K \<and> compact K \<and> C \<subseteq> K \<and> K \<subseteq> U" |
69620 | 2282 |
by metis |
2283 |
obtain V K where "C \<subseteq> V" "V \<subseteq> U" "V \<subseteq> K" "K \<subseteq> S" "compact K" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2284 |
and opeSV: "openin (top_of_set S) V" |
72372 | 2285 |
using S U \<open>compact C\<close> by (meson C in_components_subset locally_compact_compact_subopen) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2286 |
let ?\<T> = "{T. C \<subseteq> T \<and> openin (top_of_set K) T \<and> compact T \<and> T \<subseteq> K}" |
69620 | 2287 |
have CK: "C \<in> components K" |
2288 |
by (meson C \<open>C \<subseteq> V\<close> \<open>K \<subseteq> S\<close> \<open>V \<subseteq> K\<close> components_intermediate_subset subset_trans) |
|
2289 |
with \<open>compact K\<close> |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2290 |
have "C = \<Inter>{T. C \<subseteq> T \<and> openin (top_of_set K) T \<and> closedin (top_of_set K) T}" |
69620 | 2291 |
by (simp add: Sura_Bura_compact) |
2292 |
then have Ceq: "C = \<Inter>?\<T>" |
|
2293 |
by (simp add: closedin_compact_eq \<open>compact K\<close>) |
|
2294 |
obtain W where "open W" and W: "V = S \<inter> W" |
|
2295 |
using opeSV by (auto simp: openin_open) |
|
2296 |
have "-(U \<inter> W) \<inter> \<Inter>?\<T> \<noteq> {}" |
|
2297 |
proof (rule closed_imp_fip_compact) |
|
2298 |
show "- (U \<inter> W) \<inter> \<Inter>\<F> \<noteq> {}" |
|
2299 |
if "finite \<F>" and \<F>: "\<F> \<subseteq> ?\<T>" for \<F> |
|
2300 |
proof (cases "\<F> = {}") |
|
2301 |
case True |
|
2302 |
have False if "U = UNIV" "W = UNIV" |
|
2303 |
proof - |
|
2304 |
have "V = S" |
|
2305 |
by (simp add: W \<open>W = UNIV\<close>) |
|
2306 |
with neg show False |
|
2307 |
using \<open>C \<subseteq> V\<close> \<open>K \<subseteq> S\<close> \<open>V \<subseteq> K\<close> \<open>V \<subseteq> U\<close> \<open>compact K\<close> by auto |
|
2308 |
qed |
|
2309 |
with True show ?thesis |
|
2310 |
by auto |
|
2311 |
next |
|
2312 |
case False |
|
2313 |
show ?thesis |
|
2314 |
proof |
|
2315 |
assume "- (U \<inter> W) \<inter> \<Inter>\<F> = {}" |
|
2316 |
then have FUW: "\<Inter>\<F> \<subseteq> U \<inter> W" |
|
2317 |
by blast |
|
2318 |
have "C \<subseteq> \<Inter>\<F>" |
|
2319 |
using \<F> by auto |
|
2320 |
moreover have "compact (\<Inter>\<F>)" |
|
2321 |
by (metis (no_types, lifting) compact_Inter False mem_Collect_eq subsetCE \<F>) |
|
2322 |
moreover have "\<Inter>\<F> \<subseteq> K" |
|
2323 |
using False that(2) by fastforce |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2324 |
moreover have opeKF: "openin (top_of_set K) (\<Inter>\<F>)" |
69620 | 2325 |
using False \<F> \<open>finite \<F>\<close> by blast |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2326 |
then have opeVF: "openin (top_of_set V) (\<Inter>\<F>)" |
69620 | 2327 |
using W \<open>K \<subseteq> S\<close> \<open>V \<subseteq> K\<close> opeKF \<open>\<Inter>\<F> \<subseteq> K\<close> FUW openin_subset_trans by fastforce |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2328 |
then have "openin (top_of_set S) (\<Inter>\<F>)" |
69620 | 2329 |
by (metis opeSV openin_trans) |
2330 |
moreover have "\<Inter>\<F> \<subseteq> U" |
|
2331 |
by (meson \<open>V \<subseteq> U\<close> opeVF dual_order.trans openin_imp_subset) |
|
2332 |
ultimately show False |
|
2333 |
using neg by blast |
|
2334 |
qed |
|
2335 |
qed |
|
2336 |
qed (use \<open>open W\<close> \<open>open U\<close> in auto) |
|
2337 |
with W Ceq \<open>C \<subseteq> V\<close> \<open>C \<subseteq> U\<close> show False |
|
2338 |
by auto |
|
2339 |
qed |
|
2340 |
||
2341 |
||
2342 |
corollary Sura_Bura_clopen_subset_alt: |
|
2343 |
fixes S :: "'a::euclidean_space set" |
|
2344 |
assumes S: "locally compact S" and C: "C \<in> components S" and "compact C" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2345 |
and opeSU: "openin (top_of_set S) U" and "C \<subseteq> U" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2346 |
obtains K where "openin (top_of_set S) K" "compact K" "C \<subseteq> K" "K \<subseteq> U" |
69620 | 2347 |
proof - |
2348 |
obtain V where "open V" "U = S \<inter> V" |
|
2349 |
using opeSU by (auto simp: openin_open) |
|
2350 |
with \<open>C \<subseteq> U\<close> have "C \<subseteq> V" |
|
2351 |
by auto |
|
2352 |
then show ?thesis |
|
2353 |
using Sura_Bura_clopen_subset [OF S C \<open>compact C\<close> \<open>open V\<close>] |
|
2354 |
by (metis \<open>U = S \<inter> V\<close> inf.bounded_iff openin_imp_subset that) |
|
2355 |
qed |
|
2356 |
||
2357 |
corollary Sura_Bura: |
|
2358 |
fixes S :: "'a::euclidean_space set" |
|
2359 |
assumes "locally compact S" "C \<in> components S" "compact C" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2360 |
shows "C = \<Inter> {K. C \<subseteq> K \<and> compact K \<and> openin (top_of_set S) K}" |
69620 | 2361 |
(is "C = ?rhs") |
2362 |
proof |
|
2363 |
show "?rhs \<subseteq> C" |
|
2364 |
proof (clarsimp, rule ccontr) |
|
2365 |
fix x |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2366 |
assume *: "\<forall>X. C \<subseteq> X \<and> compact X \<and> openin (top_of_set S) X \<longrightarrow> x \<in> X" |
69620 | 2367 |
and "x \<notin> C" |
2368 |
obtain U V where "open U" "open V" "{x} \<subseteq> U" "C \<subseteq> V" "U \<inter> V = {}" |
|
2369 |
using separation_normal [of "{x}" C] |
|
2370 |
by (metis Int_empty_left \<open>x \<notin> C\<close> \<open>compact C\<close> closed_empty closed_insert compact_imp_closed insert_disjoint(1)) |
|
2371 |
have "x \<notin> V" |
|
2372 |
using \<open>U \<inter> V = {}\<close> \<open>{x} \<subseteq> U\<close> by blast |
|
2373 |
then show False |
|
2374 |
by (meson "*" Sura_Bura_clopen_subset \<open>C \<subseteq> V\<close> \<open>open V\<close> assms(1) assms(2) assms(3) subsetCE) |
|
2375 |
qed |
|
2376 |
qed blast |
|
2377 |
||
2378 |
||
2379 |
subsection\<open>Special cases of local connectedness and path connectedness\<close> |
|
2380 |
||
2381 |
lemma locally_connected_1: |
|
2382 |
assumes |
|
72372 | 2383 |
"\<And>V x. \<lbrakk>openin (top_of_set S) V; x \<in> V\<rbrakk> \<Longrightarrow> \<exists>U. openin (top_of_set S) U \<and> connected U \<and> x \<in> U \<and> U \<subseteq> V" |
69620 | 2384 |
shows "locally connected S" |
72372 | 2385 |
by (metis assms locally_def) |
69620 | 2386 |
|
2387 |
lemma locally_connected_2: |
|
2388 |
assumes "locally connected S" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2389 |
"openin (top_of_set S) t" |
69620 | 2390 |
"x \<in> t" |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2391 |
shows "openin (top_of_set S) (connected_component_set t x)" |
69620 | 2392 |
proof - |
2393 |
{ fix y :: 'a |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2394 |
let ?SS = "top_of_set S" |
69620 | 2395 |
assume 1: "openin ?SS t" |
2396 |
"\<forall>w x. openin ?SS w \<and> x \<in> w \<longrightarrow> (\<exists>u. openin ?SS u \<and> (\<exists>v. connected v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w))" |
|
2397 |
and "connected_component t x y" |
|
2398 |
then have "y \<in> t" and y: "y \<in> connected_component_set t x" |
|
2399 |
using connected_component_subset by blast+ |
|
2400 |
obtain F where |
|
2401 |
"\<forall>x y. (\<exists>w. openin ?SS w \<and> (\<exists>u. connected u \<and> x \<in> w \<and> w \<subseteq> u \<and> u \<subseteq> y)) = (openin ?SS (F x y) \<and> (\<exists>u. connected u \<and> x \<in> F x y \<and> F x y \<subseteq> u \<and> u \<subseteq> y))" |
|
2402 |
by moura |
|
2403 |
then obtain G where |
|
2404 |
"\<forall>a A. (\<exists>U. openin ?SS U \<and> (\<exists>V. connected V \<and> a \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> A)) = (openin ?SS (F a A) \<and> connected (G a A) \<and> a \<in> F a A \<and> F a A \<subseteq> G a A \<and> G a A \<subseteq> A)" |
|
2405 |
by moura |
|
2406 |
then have *: "openin ?SS (F y t) \<and> connected (G y t) \<and> y \<in> F y t \<and> F y t \<subseteq> G y t \<and> G y t \<subseteq> t" |
|
2407 |
using 1 \<open>y \<in> t\<close> by presburger |
|
2408 |
have "G y t \<subseteq> connected_component_set t y" |
|
2409 |
by (metis (no_types) * connected_component_eq_self connected_component_mono contra_subsetD) |
|
2410 |
then have "\<exists>A. openin ?SS A \<and> y \<in> A \<and> A \<subseteq> connected_component_set t x" |
|
2411 |
by (metis (no_types) * connected_component_eq dual_order.trans y) |
|
2412 |
} |
|
2413 |
then show ?thesis |
|
2414 |
using assms openin_subopen by (force simp: locally_def) |
|
2415 |
qed |
|
2416 |
||
2417 |
lemma locally_connected_3: |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2418 |
assumes "\<And>t x. \<lbrakk>openin (top_of_set S) t; x \<in> t\<rbrakk> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2419 |
\<Longrightarrow> openin (top_of_set S) |
69620 | 2420 |
(connected_component_set t x)" |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2421 |
"openin (top_of_set S) v" "x \<in> v" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2422 |
shows "\<exists>u. openin (top_of_set S) u \<and> connected u \<and> x \<in> u \<and> u \<subseteq> v" |
69620 | 2423 |
using assms connected_component_subset by fastforce |
2424 |
||
2425 |
lemma locally_connected: |
|
2426 |
"locally connected S \<longleftrightarrow> |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2427 |
(\<forall>v x. openin (top_of_set S) v \<and> x \<in> v |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2428 |
\<longrightarrow> (\<exists>u. openin (top_of_set S) u \<and> connected u \<and> x \<in> u \<and> u \<subseteq> v))" |
69620 | 2429 |
by (metis locally_connected_1 locally_connected_2 locally_connected_3) |
2430 |
||
2431 |
lemma locally_connected_open_connected_component: |
|
2432 |
"locally connected S \<longleftrightarrow> |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2433 |
(\<forall>t x. openin (top_of_set S) t \<and> x \<in> t |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2434 |
\<longrightarrow> openin (top_of_set S) (connected_component_set t x))" |
69620 | 2435 |
by (metis locally_connected_1 locally_connected_2 locally_connected_3) |
2436 |
||
2437 |
lemma locally_path_connected_1: |
|
2438 |
assumes |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2439 |
"\<And>v x. \<lbrakk>openin (top_of_set S) v; x \<in> v\<rbrakk> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2440 |
\<Longrightarrow> \<exists>u. openin (top_of_set S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> v" |
69620 | 2441 |
shows "locally path_connected S" |
78474 | 2442 |
by (force simp: locally_def dest: assms) |
69620 | 2443 |
|
2444 |
lemma locally_path_connected_2: |
|
2445 |
assumes "locally path_connected S" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2446 |
"openin (top_of_set S) t" |
69620 | 2447 |
"x \<in> t" |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2448 |
shows "openin (top_of_set S) (path_component_set t x)" |
69620 | 2449 |
proof - |
2450 |
{ fix y :: 'a |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2451 |
let ?SS = "top_of_set S" |
69620 | 2452 |
assume 1: "openin ?SS t" |
2453 |
"\<forall>w x. openin ?SS w \<and> x \<in> w \<longrightarrow> (\<exists>u. openin ?SS u \<and> (\<exists>v. path_connected v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w))" |
|
2454 |
and "path_component t x y" |
|
2455 |
then have "y \<in> t" and y: "y \<in> path_component_set t x" |
|
2456 |
using path_component_mem(2) by blast+ |
|
2457 |
obtain F where |
|
2458 |
"\<forall>x y. (\<exists>w. openin ?SS w \<and> (\<exists>u. path_connected u \<and> x \<in> w \<and> w \<subseteq> u \<and> u \<subseteq> y)) = (openin ?SS (F x y) \<and> (\<exists>u. path_connected u \<and> x \<in> F x y \<and> F x y \<subseteq> u \<and> u \<subseteq> y))" |
|
2459 |
by moura |
|
2460 |
then obtain G where |
|
2461 |
"\<forall>a A. (\<exists>U. openin ?SS U \<and> (\<exists>V. path_connected V \<and> a \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> A)) = (openin ?SS (F a A) \<and> path_connected (G a A) \<and> a \<in> F a A \<and> F a A \<subseteq> G a A \<and> G a A \<subseteq> A)" |
|
2462 |
by moura |
|
2463 |
then have *: "openin ?SS (F y t) \<and> path_connected (G y t) \<and> y \<in> F y t \<and> F y t \<subseteq> G y t \<and> G y t \<subseteq> t" |
|
2464 |
using 1 \<open>y \<in> t\<close> by presburger |
|
2465 |
have "G y t \<subseteq> path_component_set t y" |
|
69712 | 2466 |
using * path_component_maximal rev_subsetD by blast |
69620 | 2467 |
then have "\<exists>A. openin ?SS A \<and> y \<in> A \<and> A \<subseteq> path_component_set t x" |
2468 |
by (metis "*" \<open>G y t \<subseteq> path_component_set t y\<close> dual_order.trans path_component_eq y) |
|
2469 |
} |
|
2470 |
then show ?thesis |
|
2471 |
using assms openin_subopen by (force simp: locally_def) |
|
2472 |
qed |
|
2473 |
||
2474 |
lemma locally_path_connected_3: |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2475 |
assumes "\<And>t x. \<lbrakk>openin (top_of_set S) t; x \<in> t\<rbrakk> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2476 |
\<Longrightarrow> openin (top_of_set S) (path_component_set t x)" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2477 |
"openin (top_of_set S) v" "x \<in> v" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2478 |
shows "\<exists>u. openin (top_of_set S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> v" |
69620 | 2479 |
proof - |
2480 |
have "path_component v x x" |
|
2481 |
by (meson assms(3) path_component_refl) |
|
2482 |
then show ?thesis |
|
71746 | 2483 |
by (metis assms mem_Collect_eq path_component_subset path_connected_path_component) |
69620 | 2484 |
qed |
2485 |
||
2486 |
proposition locally_path_connected: |
|
2487 |
"locally path_connected S \<longleftrightarrow> |
|
71769 | 2488 |
(\<forall>V x. openin (top_of_set S) V \<and> x \<in> V |
2489 |
\<longrightarrow> (\<exists>U. openin (top_of_set S) U \<and> path_connected U \<and> x \<in> U \<and> U \<subseteq> V))" |
|
69620 | 2490 |
by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) |
2491 |
||
2492 |
proposition locally_path_connected_open_path_component: |
|
2493 |
"locally path_connected S \<longleftrightarrow> |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2494 |
(\<forall>t x. openin (top_of_set S) t \<and> x \<in> t |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2495 |
\<longrightarrow> openin (top_of_set S) (path_component_set t x))" |
69620 | 2496 |
by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) |
2497 |
||
2498 |
lemma locally_connected_open_component: |
|
2499 |
"locally connected S \<longleftrightarrow> |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2500 |
(\<forall>t c. openin (top_of_set S) t \<and> c \<in> components t |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2501 |
\<longrightarrow> openin (top_of_set S) c)" |
69620 | 2502 |
by (metis components_iff locally_connected_open_connected_component) |
2503 |
||
2504 |
proposition locally_connected_im_kleinen: |
|
2505 |
"locally connected S \<longleftrightarrow> |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2506 |
(\<forall>v x. openin (top_of_set S) v \<and> x \<in> v |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2507 |
\<longrightarrow> (\<exists>u. openin (top_of_set S) u \<and> |
69620 | 2508 |
x \<in> u \<and> u \<subseteq> v \<and> |
2509 |
(\<forall>y. y \<in> u \<longrightarrow> (\<exists>c. connected c \<and> c \<subseteq> v \<and> x \<in> c \<and> y \<in> c))))" |
|
2510 |
(is "?lhs = ?rhs") |
|
2511 |
proof |
|
2512 |
assume ?lhs |
|
2513 |
then show ?rhs |
|
78474 | 2514 |
by (fastforce simp: locally_connected) |
69620 | 2515 |
next |
2516 |
assume ?rhs |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2517 |
have *: "\<exists>T. openin (top_of_set S) T \<and> x \<in> T \<and> T \<subseteq> c" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2518 |
if "openin (top_of_set S) t" and c: "c \<in> components t" and "x \<in> c" for t c x |
69620 | 2519 |
proof - |
2520 |
from that \<open>?rhs\<close> [rule_format, of t x] |
|
2521 |
obtain u where u: |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2522 |
"openin (top_of_set S) u \<and> x \<in> u \<and> u \<subseteq> t \<and> |
69620 | 2523 |
(\<forall>y. y \<in> u \<longrightarrow> (\<exists>c. connected c \<and> c \<subseteq> t \<and> x \<in> c \<and> y \<in> c))" |
2524 |
using in_components_subset by auto |
|
2525 |
obtain F :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a" where |
|
2526 |
"\<forall>x y. (\<exists>z. z \<in> x \<and> y = connected_component_set x z) = (F x y \<in> x \<and> y = connected_component_set x (F x y))" |
|
2527 |
by moura |
|
2528 |
then have F: "F t c \<in> t \<and> c = connected_component_set t (F t c)" |
|
2529 |
by (meson components_iff c) |
|
2530 |
obtain G :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a" where |
|
2531 |
G: "\<forall>x y. (\<exists>z. z \<in> y \<and> z \<notin> x) = (G x y \<in> y \<and> G x y \<notin> x)" |
|
2532 |
by moura |
|
2533 |
have "G c u \<notin> u \<or> G c u \<in> c" |
|
2534 |
using F by (metis (full_types) u connected_componentI connected_component_eq mem_Collect_eq that(3)) |
|
2535 |
then show ?thesis |
|
2536 |
using G u by auto |
|
2537 |
qed |
|
2538 |
show ?lhs |
|
71769 | 2539 |
unfolding locally_connected_open_component by (meson "*" openin_subopen) |
69620 | 2540 |
qed |
2541 |
||
2542 |
proposition locally_path_connected_im_kleinen: |
|
2543 |
"locally path_connected S \<longleftrightarrow> |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2544 |
(\<forall>v x. openin (top_of_set S) v \<and> x \<in> v |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2545 |
\<longrightarrow> (\<exists>u. openin (top_of_set S) u \<and> |
69620 | 2546 |
x \<in> u \<and> u \<subseteq> v \<and> |
2547 |
(\<forall>y. y \<in> u \<longrightarrow> (\<exists>p. path p \<and> path_image p \<subseteq> v \<and> |
|
2548 |
pathstart p = x \<and> pathfinish p = y))))" |
|
2549 |
(is "?lhs = ?rhs") |
|
2550 |
proof |
|
2551 |
assume ?lhs |
|
2552 |
then show ?rhs |
|
2553 |
apply (simp add: locally_path_connected path_connected_def) |
|
2554 |
apply (erule all_forward ex_forward imp_forward conjE | simp)+ |
|
2555 |
by (meson dual_order.trans) |
|
2556 |
next |
|
2557 |
assume ?rhs |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2558 |
have *: "\<exists>T. openin (top_of_set S) T \<and> |
69620 | 2559 |
x \<in> T \<and> T \<subseteq> path_component_set u z" |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2560 |
if "openin (top_of_set S) u" and "z \<in> u" and c: "path_component u z x" for u z x |
69620 | 2561 |
proof - |
2562 |
have "x \<in> u" |
|
2563 |
by (meson c path_component_mem(2)) |
|
2564 |
with that \<open>?rhs\<close> [rule_format, of u x] |
|
2565 |
obtain U where U: |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2566 |
"openin (top_of_set S) U \<and> x \<in> U \<and> U \<subseteq> u \<and> |
69620 | 2567 |
(\<forall>y. y \<in> U \<longrightarrow> (\<exists>p. path p \<and> path_image p \<subseteq> u \<and> pathstart p = x \<and> pathfinish p = y))" |
2568 |
by blast |
|
2569 |
show ?thesis |
|
71769 | 2570 |
by (metis U c mem_Collect_eq path_component_def path_component_eq subsetI) |
69620 | 2571 |
qed |
2572 |
show ?lhs |
|
71769 | 2573 |
unfolding locally_path_connected_open_path_component |
2574 |
using "*" openin_subopen by fastforce |
|
69620 | 2575 |
qed |
2576 |
||
2577 |
lemma locally_path_connected_imp_locally_connected: |
|
2578 |
"locally path_connected S \<Longrightarrow> locally connected S" |
|
2579 |
using locally_mono path_connected_imp_connected by blast |
|
2580 |
||
2581 |
lemma locally_connected_components: |
|
2582 |
"\<lbrakk>locally connected S; c \<in> components S\<rbrakk> \<Longrightarrow> locally connected c" |
|
2583 |
by (meson locally_connected_open_component locally_open_subset openin_subtopology_self) |
|
2584 |
||
2585 |
lemma locally_path_connected_components: |
|
2586 |
"\<lbrakk>locally path_connected S; c \<in> components S\<rbrakk> \<Longrightarrow> locally path_connected c" |
|
2587 |
by (meson locally_connected_open_component locally_open_subset locally_path_connected_imp_locally_connected openin_subtopology_self) |
|
2588 |
||
2589 |
lemma locally_path_connected_connected_component: |
|
2590 |
"locally path_connected S \<Longrightarrow> locally path_connected (connected_component_set S x)" |
|
2591 |
by (metis components_iff connected_component_eq_empty locally_empty locally_path_connected_components) |
|
2592 |
||
2593 |
lemma open_imp_locally_path_connected: |
|
2594 |
fixes S :: "'a :: real_normed_vector set" |
|
71769 | 2595 |
assumes "open S" |
2596 |
shows "locally path_connected S" |
|
2597 |
proof (rule locally_mono) |
|
2598 |
show "locally convex S" |
|
2599 |
using assms unfolding locally_def |
|
2600 |
by (meson open_ball centre_in_ball convex_ball openE open_subset openin_imp_subset openin_open_trans subset_trans) |
|
2601 |
show "\<And>T::'a set. convex T \<Longrightarrow> path_connected T" |
|
2602 |
using convex_imp_path_connected by blast |
|
2603 |
qed |
|
69620 | 2604 |
|
2605 |
lemma open_imp_locally_connected: |
|
2606 |
fixes S :: "'a :: real_normed_vector set" |
|
2607 |
shows "open S \<Longrightarrow> locally connected S" |
|
2608 |
by (simp add: locally_path_connected_imp_locally_connected open_imp_locally_path_connected) |
|
2609 |
||
2610 |
lemma locally_path_connected_UNIV: "locally path_connected (UNIV::'a :: real_normed_vector set)" |
|
2611 |
by (simp add: open_imp_locally_path_connected) |
|
2612 |
||
2613 |
lemma locally_connected_UNIV: "locally connected (UNIV::'a :: real_normed_vector set)" |
|
2614 |
by (simp add: open_imp_locally_connected) |
|
2615 |
||
2616 |
lemma openin_connected_component_locally_connected: |
|
2617 |
"locally connected S |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2618 |
\<Longrightarrow> openin (top_of_set S) (connected_component_set S x)" |
71769 | 2619 |
by (metis connected_component_eq_empty locally_connected_2 openin_empty openin_subtopology_self) |
69620 | 2620 |
|
2621 |
lemma openin_components_locally_connected: |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2622 |
"\<lbrakk>locally connected S; c \<in> components S\<rbrakk> \<Longrightarrow> openin (top_of_set S) c" |
69620 | 2623 |
using locally_connected_open_component openin_subtopology_self by blast |
2624 |
||
2625 |
lemma openin_path_component_locally_path_connected: |
|
2626 |
"locally path_connected S |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2627 |
\<Longrightarrow> openin (top_of_set S) (path_component_set S x)" |
69620 | 2628 |
by (metis (no_types) empty_iff locally_path_connected_2 openin_subopen openin_subtopology_self path_component_eq_empty) |
2629 |
||
2630 |
lemma closedin_path_component_locally_path_connected: |
|
71769 | 2631 |
assumes "locally path_connected S" |
2632 |
shows "closedin (top_of_set S) (path_component_set S x)" |
|
2633 |
proof - |
|
2634 |
have "openin (top_of_set S) (\<Union> ({path_component_set S y |y. y \<in> S} - {path_component_set S x}))" |
|
2635 |
using locally_path_connected_2 assms by fastforce |
|
2636 |
then show ?thesis |
|
2637 |
by (simp add: closedin_def path_component_subset complement_path_component_Union) |
|
2638 |
qed |
|
69620 | 2639 |
|
2640 |
lemma convex_imp_locally_path_connected: |
|
2641 |
fixes S :: "'a:: real_normed_vector set" |
|
71769 | 2642 |
assumes "convex S" |
2643 |
shows "locally path_connected S" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
2644 |
proof (clarsimp simp: locally_path_connected) |
71769 | 2645 |
fix V x |
2646 |
assume "openin (top_of_set S) V" and "x \<in> V" |
|
2647 |
then obtain T e where "V = S \<inter> T" "x \<in> S" "0 < e" "ball x e \<subseteq> T" |
|
2648 |
by (metis Int_iff openE openin_open) |
|
2649 |
then have "openin (top_of_set S) (S \<inter> ball x e)" "path_connected (S \<inter> ball x e)" |
|
2650 |
by (simp_all add: assms convex_Int convex_imp_path_connected openin_open_Int) |
|
2651 |
then show "\<exists>U. openin (top_of_set S) U \<and> path_connected U \<and> x \<in> U \<and> U \<subseteq> V" |
|
2652 |
using \<open>0 < e\<close> \<open>V = S \<inter> T\<close> \<open>ball x e \<subseteq> T\<close> \<open>x \<in> S\<close> by auto |
|
2653 |
qed |
|
69620 | 2654 |
|
2655 |
lemma convex_imp_locally_connected: |
|
2656 |
fixes S :: "'a:: real_normed_vector set" |
|
2657 |
shows "convex S \<Longrightarrow> locally connected S" |
|
2658 |
by (simp add: locally_path_connected_imp_locally_connected convex_imp_locally_path_connected) |
|
2659 |
||
2660 |
||
2661 |
subsection\<open>Relations between components and path components\<close> |
|
2662 |
||
2663 |
lemma path_component_eq_connected_component: |
|
2664 |
assumes "locally path_connected S" |
|
2665 |
shows "(path_component S x = connected_component S x)" |
|
2666 |
proof (cases "x \<in> S") |
|
2667 |
case True |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2668 |
have "openin (top_of_set (connected_component_set S x)) (path_component_set S x)" |
71769 | 2669 |
proof (rule openin_subset_trans) |
2670 |
show "openin (top_of_set S) (path_component_set S x)" |
|
2671 |
by (simp add: True assms locally_path_connected_2) |
|
2672 |
show "connected_component_set S x \<subseteq> S" |
|
2673 |
by (simp add: connected_component_subset) |
|
2674 |
qed (simp add: path_component_subset_connected_component) |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2675 |
moreover have "closedin (top_of_set (connected_component_set S x)) (path_component_set S x)" |
71769 | 2676 |
proof (rule closedin_subset_trans [of S]) |
2677 |
show "closedin (top_of_set S) (path_component_set S x)" |
|
2678 |
by (simp add: assms closedin_path_component_locally_path_connected) |
|
2679 |
show "connected_component_set S x \<subseteq> S" |
|
2680 |
by (simp add: connected_component_subset) |
|
2681 |
qed (simp add: path_component_subset_connected_component) |
|
69620 | 2682 |
ultimately have *: "path_component_set S x = connected_component_set S x" |
2683 |
by (metis connected_connected_component connected_clopen True path_component_eq_empty) |
|
2684 |
then show ?thesis |
|
2685 |
by blast |
|
2686 |
next |
|
2687 |
case False then show ?thesis |
|
2688 |
by (metis Collect_empty_eq_bot connected_component_eq_empty path_component_eq_empty) |
|
2689 |
qed |
|
2690 |
||
2691 |
lemma path_component_eq_connected_component_set: |
|
2692 |
"locally path_connected S \<Longrightarrow> (path_component_set S x = connected_component_set S x)" |
|
2693 |
by (simp add: path_component_eq_connected_component) |
|
2694 |
||
2695 |
lemma locally_path_connected_path_component: |
|
2696 |
"locally path_connected S \<Longrightarrow> locally path_connected (path_component_set S x)" |
|
2697 |
using locally_path_connected_connected_component path_component_eq_connected_component by fastforce |
|
2698 |
||
2699 |
lemma open_path_connected_component: |
|
2700 |
fixes S :: "'a :: real_normed_vector set" |
|
2701 |
shows "open S \<Longrightarrow> path_component S x = connected_component S x" |
|
2702 |
by (simp add: path_component_eq_connected_component open_imp_locally_path_connected) |
|
2703 |
||
2704 |
lemma open_path_connected_component_set: |
|
2705 |
fixes S :: "'a :: real_normed_vector set" |
|
2706 |
shows "open S \<Longrightarrow> path_component_set S x = connected_component_set S x" |
|
2707 |
by (simp add: open_path_connected_component) |
|
2708 |
||
2709 |
proposition locally_connected_quotient_image: |
|
2710 |
assumes lcS: "locally connected S" |
|
2711 |
and oo: "\<And>T. T \<subseteq> f ` S |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2712 |
\<Longrightarrow> openin (top_of_set S) (S \<inter> f -` T) \<longleftrightarrow> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2713 |
openin (top_of_set (f ` S)) T" |
69620 | 2714 |
shows "locally connected (f ` S)" |
2715 |
proof (clarsimp simp: locally_connected_open_component) |
|
2716 |
fix U C |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2717 |
assume opefSU: "openin (top_of_set (f ` S)) U" and "C \<in> components U" |
69620 | 2718 |
then have "C \<subseteq> U" "U \<subseteq> f ` S" |
2719 |
by (meson in_components_subset openin_imp_subset)+ |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2720 |
then have "openin (top_of_set (f ` S)) C \<longleftrightarrow> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2721 |
openin (top_of_set S) (S \<inter> f -` C)" |
69620 | 2722 |
by (auto simp: oo) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2723 |
moreover have "openin (top_of_set S) (S \<inter> f -` C)" |
69620 | 2724 |
proof (subst openin_subopen, clarify) |
2725 |
fix x |
|
2726 |
assume "x \<in> S" "f x \<in> C" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2727 |
show "\<exists>T. openin (top_of_set S) T \<and> x \<in> T \<and> T \<subseteq> (S \<inter> f -` C)" |
69620 | 2728 |
proof (intro conjI exI) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2729 |
show "openin (top_of_set S) (connected_component_set (S \<inter> f -` U) x)" |
69620 | 2730 |
proof (rule ccontr) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2731 |
assume **: "\<not> openin (top_of_set S) (connected_component_set (S \<inter> f -` U) x)" |
69620 | 2732 |
then have "x \<notin> (S \<inter> f -` U)" |
2733 |
using \<open>U \<subseteq> f ` S\<close> opefSU lcS locally_connected_2 oo by blast |
|
2734 |
with ** show False |
|
2735 |
by (metis (no_types) connected_component_eq_empty empty_iff openin_subopen) |
|
2736 |
qed |
|
2737 |
next |
|
2738 |
show "x \<in> connected_component_set (S \<inter> f -` U) x" |
|
2739 |
using \<open>C \<subseteq> U\<close> \<open>f x \<in> C\<close> \<open>x \<in> S\<close> by auto |
|
2740 |
next |
|
2741 |
have contf: "continuous_on S f" |
|
2742 |
by (simp add: continuous_on_open oo openin_imp_subset) |
|
2743 |
then have "continuous_on (connected_component_set (S \<inter> f -` U) x) f" |
|
71769 | 2744 |
by (meson connected_component_subset continuous_on_subset inf.boundedE) |
69620 | 2745 |
then have "connected (f ` connected_component_set (S \<inter> f -` U) x)" |
2746 |
by (rule connected_continuous_image [OF _ connected_connected_component]) |
|
2747 |
moreover have "f ` connected_component_set (S \<inter> f -` U) x \<subseteq> U" |
|
2748 |
using connected_component_in by blast |
|
2749 |
moreover have "C \<inter> f ` connected_component_set (S \<inter> f -` U) x \<noteq> {}" |
|
2750 |
using \<open>C \<subseteq> U\<close> \<open>f x \<in> C\<close> \<open>x \<in> S\<close> by fastforce |
|
2751 |
ultimately have fC: "f ` (connected_component_set (S \<inter> f -` U) x) \<subseteq> C" |
|
2752 |
by (rule components_maximal [OF \<open>C \<in> components U\<close>]) |
|
2753 |
have cUC: "connected_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` C)" |
|
2754 |
using connected_component_subset fC by blast |
|
2755 |
have "connected_component_set (S \<inter> f -` U) x \<subseteq> connected_component_set (S \<inter> f -` C) x" |
|
2756 |
proof - |
|
2757 |
{ assume "x \<in> connected_component_set (S \<inter> f -` U) x" |
|
2758 |
then have ?thesis |
|
2759 |
using cUC connected_component_idemp connected_component_mono by blast } |
|
2760 |
then show ?thesis |
|
2761 |
using connected_component_eq_empty by auto |
|
2762 |
qed |
|
2763 |
also have "\<dots> \<subseteq> (S \<inter> f -` C)" |
|
2764 |
by (rule connected_component_subset) |
|
2765 |
finally show "connected_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` C)" . |
|
2766 |
qed |
|
2767 |
qed |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2768 |
ultimately show "openin (top_of_set (f ` S)) C" |
69620 | 2769 |
by metis |
2770 |
qed |
|
2771 |
||
2772 |
text\<open>The proof resembles that above but is not identical!\<close> |
|
2773 |
proposition locally_path_connected_quotient_image: |
|
2774 |
assumes lcS: "locally path_connected S" |
|
2775 |
and oo: "\<And>T. T \<subseteq> f ` S |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2776 |
\<Longrightarrow> openin (top_of_set S) (S \<inter> f -` T) \<longleftrightarrow> openin (top_of_set (f ` S)) T" |
69620 | 2777 |
shows "locally path_connected (f ` S)" |
2778 |
proof (clarsimp simp: locally_path_connected_open_path_component) |
|
2779 |
fix U y |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2780 |
assume opefSU: "openin (top_of_set (f ` S)) U" and "y \<in> U" |
69620 | 2781 |
then have "path_component_set U y \<subseteq> U" "U \<subseteq> f ` S" |
2782 |
by (meson path_component_subset openin_imp_subset)+ |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2783 |
then have "openin (top_of_set (f ` S)) (path_component_set U y) \<longleftrightarrow> |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2784 |
openin (top_of_set S) (S \<inter> f -` path_component_set U y)" |
69620 | 2785 |
proof - |
2786 |
have "path_component_set U y \<subseteq> f ` S" |
|
2787 |
using \<open>U \<subseteq> f ` S\<close> \<open>path_component_set U y \<subseteq> U\<close> by blast |
|
2788 |
then show ?thesis |
|
2789 |
using oo by blast |
|
2790 |
qed |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2791 |
moreover have "openin (top_of_set S) (S \<inter> f -` path_component_set U y)" |
69620 | 2792 |
proof (subst openin_subopen, clarify) |
2793 |
fix x |
|
2794 |
assume "x \<in> S" and Uyfx: "path_component U y (f x)" |
|
2795 |
then have "f x \<in> U" |
|
2796 |
using path_component_mem by blast |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2797 |
show "\<exists>T. openin (top_of_set S) T \<and> x \<in> T \<and> T \<subseteq> (S \<inter> f -` path_component_set U y)" |
69620 | 2798 |
proof (intro conjI exI) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2799 |
show "openin (top_of_set S) (path_component_set (S \<inter> f -` U) x)" |
69620 | 2800 |
proof (rule ccontr) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2801 |
assume **: "\<not> openin (top_of_set S) (path_component_set (S \<inter> f -` U) x)" |
69620 | 2802 |
then have "x \<notin> (S \<inter> f -` U)" |
2803 |
by (metis (no_types, lifting) \<open>U \<subseteq> f ` S\<close> opefSU lcS oo locally_path_connected_open_path_component) |
|
2804 |
then show False |
|
2805 |
using ** \<open>path_component_set U y \<subseteq> U\<close> \<open>x \<in> S\<close> \<open>path_component U y (f x)\<close> by blast |
|
2806 |
qed |
|
2807 |
next |
|
2808 |
show "x \<in> path_component_set (S \<inter> f -` U) x" |
|
2809 |
by (simp add: \<open>f x \<in> U\<close> \<open>x \<in> S\<close> path_component_refl) |
|
2810 |
next |
|
2811 |
have contf: "continuous_on S f" |
|
2812 |
by (simp add: continuous_on_open oo openin_imp_subset) |
|
2813 |
then have "continuous_on (path_component_set (S \<inter> f -` U) x) f" |
|
71769 | 2814 |
by (meson Int_lower1 continuous_on_subset path_component_subset) |
69620 | 2815 |
then have "path_connected (f ` path_component_set (S \<inter> f -` U) x)" |
2816 |
by (simp add: path_connected_continuous_image) |
|
2817 |
moreover have "f ` path_component_set (S \<inter> f -` U) x \<subseteq> U" |
|
2818 |
using path_component_mem by fastforce |
|
2819 |
moreover have "f x \<in> f ` path_component_set (S \<inter> f -` U) x" |
|
2820 |
by (force simp: \<open>x \<in> S\<close> \<open>f x \<in> U\<close> path_component_refl_eq) |
|
2821 |
ultimately have "f ` (path_component_set (S \<inter> f -` U) x) \<subseteq> path_component_set U (f x)" |
|
2822 |
by (meson path_component_maximal) |
|
2823 |
also have "\<dots> \<subseteq> path_component_set U y" |
|
2824 |
by (simp add: Uyfx path_component_maximal path_component_subset path_component_sym) |
|
2825 |
finally have fC: "f ` (path_component_set (S \<inter> f -` U) x) \<subseteq> path_component_set U y" . |
|
2826 |
have cUC: "path_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` path_component_set U y)" |
|
2827 |
using path_component_subset fC by blast |
|
2828 |
have "path_component_set (S \<inter> f -` U) x \<subseteq> path_component_set (S \<inter> f -` path_component_set U y) x" |
|
2829 |
proof - |
|
2830 |
have "\<And>a. path_component_set (path_component_set (S \<inter> f -` U) x) a \<subseteq> path_component_set (S \<inter> f -` path_component_set U y) a" |
|
2831 |
using cUC path_component_mono by blast |
|
2832 |
then show ?thesis |
|
2833 |
using path_component_path_component by blast |
|
2834 |
qed |
|
2835 |
also have "\<dots> \<subseteq> (S \<inter> f -` path_component_set U y)" |
|
2836 |
by (rule path_component_subset) |
|
2837 |
finally show "path_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` path_component_set U y)" . |
|
2838 |
qed |
|
2839 |
qed |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2840 |
ultimately show "openin (top_of_set (f ` S)) (path_component_set U y)" |
69620 | 2841 |
by metis |
2842 |
qed |
|
2843 |
||
70136 | 2844 |
subsection\<^marker>\<open>tag unimportant\<close>\<open>Components, continuity, openin, closedin\<close> |
69620 | 2845 |
|
2846 |
lemma continuous_on_components_gen: |
|
2847 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" |
|
71769 | 2848 |
assumes "\<And>C. C \<in> components S \<Longrightarrow> |
2849 |
openin (top_of_set S) C \<and> continuous_on C f" |
|
69620 | 2850 |
shows "continuous_on S f" |
2851 |
proof (clarsimp simp: continuous_openin_preimage_eq) |
|
2852 |
fix t :: "'b set" |
|
2853 |
assume "open t" |
|
2854 |
have *: "S \<inter> f -` t = (\<Union>c \<in> components S. c \<inter> f -` t)" |
|
2855 |
by auto |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2856 |
show "openin (top_of_set S) (S \<inter> f -` t)" |
69620 | 2857 |
unfolding * using \<open>open t\<close> assms continuous_openin_preimage_gen openin_trans openin_Union by blast |
2858 |
qed |
|
2859 |
||
2860 |
lemma continuous_on_components: |
|
2861 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" |
|
71769 | 2862 |
assumes "locally connected S " "\<And>C. C \<in> components S \<Longrightarrow> continuous_on C f" |
2863 |
shows "continuous_on S f" |
|
2864 |
proof (rule continuous_on_components_gen) |
|
2865 |
fix C |
|
2866 |
assume "C \<in> components S" |
|
2867 |
then show "openin (top_of_set S) C \<and> continuous_on C f" |
|
2868 |
by (simp add: assms openin_components_locally_connected) |
|
2869 |
qed |
|
69620 | 2870 |
|
2871 |
lemma continuous_on_components_eq: |
|
2872 |
"locally connected S |
|
2873 |
\<Longrightarrow> (continuous_on S f \<longleftrightarrow> (\<forall>c \<in> components S. continuous_on c f))" |
|
2874 |
by (meson continuous_on_components continuous_on_subset in_components_subset) |
|
2875 |
||
2876 |
lemma continuous_on_components_open: |
|
2877 |
fixes S :: "'a::real_normed_vector set" |
|
2878 |
assumes "open S " |
|
2879 |
"\<And>c. c \<in> components S \<Longrightarrow> continuous_on c f" |
|
2880 |
shows "continuous_on S f" |
|
2881 |
using continuous_on_components open_imp_locally_connected assms by blast |
|
2882 |
||
2883 |
lemma continuous_on_components_open_eq: |
|
2884 |
fixes S :: "'a::real_normed_vector set" |
|
2885 |
shows "open S \<Longrightarrow> (continuous_on S f \<longleftrightarrow> (\<forall>c \<in> components S. continuous_on c f))" |
|
2886 |
using continuous_on_subset in_components_subset |
|
2887 |
by (blast intro: continuous_on_components_open) |
|
2888 |
||
2889 |
lemma closedin_union_complement_components: |
|
71769 | 2890 |
assumes U: "locally connected U" |
2891 |
and S: "closedin (top_of_set U) S" |
|
2892 |
and cuS: "c \<subseteq> components(U - S)" |
|
2893 |
shows "closedin (top_of_set U) (S \<union> \<Union>c)" |
|
69620 | 2894 |
proof - |
72372 | 2895 |
have di: "(\<And>S T. S \<in> c \<and> T \<in> c' \<Longrightarrow> disjnt S T) \<Longrightarrow> disjnt (\<Union> c) (\<Union> c')" for c' |
69620 | 2896 |
by (simp add: disjnt_def) blast |
71769 | 2897 |
have "S \<subseteq> U" |
69620 | 2898 |
using S closedin_imp_subset by blast |
71769 | 2899 |
moreover have "U - S = \<Union>c \<union> \<Union>(components (U - S) - c)" |
69620 | 2900 |
by (metis Diff_partition Union_components Union_Un_distrib assms(3)) |
71769 | 2901 |
moreover have "disjnt (\<Union>c) (\<Union>(components (U - S) - c))" |
69620 | 2902 |
apply (rule di) |
72372 | 2903 |
by (metis di DiffD1 DiffD2 assms(3) components_nonoverlap disjnt_def subsetCE) |
71769 | 2904 |
ultimately have eq: "S \<union> \<Union>c = U - (\<Union>(components(U - S) - c))" |
69620 | 2905 |
by (auto simp: disjnt_def) |
71769 | 2906 |
have *: "openin (top_of_set U) (\<Union>(components (U - S) - c))" |
2907 |
proof (rule openin_Union [OF openin_trans [of "U - S"]]) |
|
2908 |
show "openin (top_of_set (U - S)) T" if "T \<in> components (U - S) - c" for T |
|
2909 |
using that by (simp add: U S locally_diff_closed openin_components_locally_connected) |
|
2910 |
show "openin (top_of_set U) (U - S)" if "T \<in> components (U - S) - c" for T |
|
2911 |
using that by (simp add: openin_diff S) |
|
2912 |
qed |
|
2913 |
have "closedin (top_of_set U) (U - \<Union> (components (U - S) - c))" |
|
2914 |
by (metis closedin_diff closedin_topspace topspace_euclidean_subtopology *) |
|
2915 |
then have "openin (top_of_set U) (U - (U - \<Union>(components (U - S) - c)))" |
|
2916 |
by (simp add: openin_diff) |
|
69620 | 2917 |
then show ?thesis |
2918 |
by (force simp: eq closedin_def) |
|
2919 |
qed |
|
2920 |
||
2921 |
lemma closed_union_complement_components: |
|
2922 |
fixes S :: "'a::real_normed_vector set" |
|
2923 |
assumes S: "closed S" and c: "c \<subseteq> components(- S)" |
|
2924 |
shows "closed(S \<union> \<Union> c)" |
|
2925 |
proof - |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2926 |
have "closedin (top_of_set UNIV) (S \<union> \<Union>c)" |
71769 | 2927 |
by (metis Compl_eq_Diff_UNIV S c closed_closedin closedin_union_complement_components locally_connected_UNIV subtopology_UNIV) |
69620 | 2928 |
then show ?thesis by simp |
2929 |
qed |
|
2930 |
||
2931 |
lemma closedin_Un_complement_component: |
|
2932 |
fixes S :: "'a::real_normed_vector set" |
|
2933 |
assumes u: "locally connected u" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2934 |
and S: "closedin (top_of_set u) S" |
69620 | 2935 |
and c: " c \<in> components(u - S)" |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2936 |
shows "closedin (top_of_set u) (S \<union> c)" |
69620 | 2937 |
proof - |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
2938 |
have "closedin (top_of_set u) (S \<union> \<Union>{c})" |
69620 | 2939 |
using c by (blast intro: closedin_union_complement_components [OF u S]) |
2940 |
then show ?thesis |
|
2941 |
by simp |
|
2942 |
qed |
|
2943 |
||
2944 |
lemma closed_Un_complement_component: |
|
2945 |
fixes S :: "'a::real_normed_vector set" |
|
2946 |
assumes S: "closed S" and c: " c \<in> components(-S)" |
|
2947 |
shows "closed (S \<union> c)" |
|
2948 |
by (metis Compl_eq_Diff_UNIV S c closed_closedin closedin_Un_complement_component |
|
2949 |
locally_connected_UNIV subtopology_UNIV) |
|
2950 |
||
2951 |
||
2952 |
subsection\<open>Existence of isometry between subspaces of same dimension\<close> |
|
2953 |
||
2954 |
lemma isometry_subset_subspace: |
|
2955 |
fixes S :: "'a::euclidean_space set" |
|
2956 |
and T :: "'b::euclidean_space set" |
|
2957 |
assumes S: "subspace S" |
|
2958 |
and T: "subspace T" |
|
2959 |
and d: "dim S \<le> dim T" |
|
78457 | 2960 |
obtains f where "linear f" "f \<in> S \<rightarrow> T" "\<And>x. x \<in> S \<Longrightarrow> norm(f x) = norm x" |
69620 | 2961 |
proof - |
2962 |
obtain B where "B \<subseteq> S" and Borth: "pairwise orthogonal B" |
|
2963 |
and B1: "\<And>x. x \<in> B \<Longrightarrow> norm x = 1" |
|
2964 |
and "independent B" "finite B" "card B = dim S" "span B = S" |
|
78516
56a408fa2440
substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents:
78474
diff
changeset
|
2965 |
by (metis orthonormal_basis_subspace [OF S] independent_imp_finite) |
69620 | 2966 |
obtain C where "C \<subseteq> T" and Corth: "pairwise orthogonal C" |
2967 |
and C1:"\<And>x. x \<in> C \<Longrightarrow> norm x = 1" |
|
2968 |
and "independent C" "finite C" "card C = dim T" "span C = T" |
|
78516
56a408fa2440
substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents:
78474
diff
changeset
|
2969 |
by (metis orthonormal_basis_subspace [OF T] independent_imp_finite) |
69620 | 2970 |
obtain fb where "fb ` B \<subseteq> C" "inj_on fb B" |
2971 |
by (metis \<open>card B = dim S\<close> \<open>card C = dim T\<close> \<open>finite B\<close> \<open>finite C\<close> card_le_inj d) |
|
2972 |
then have pairwise_orth_fb: "pairwise (\<lambda>v j. orthogonal (fb v) (fb j)) B" |
|
72372 | 2973 |
using Corth unfolding pairwise_def inj_on_def |
2974 |
by (blast intro: orthogonal_clauses) |
|
69620 | 2975 |
obtain f where "linear f" and ffb: "\<And>x. x \<in> B \<Longrightarrow> f x = fb x" |
2976 |
using linear_independent_extend \<open>independent B\<close> by fastforce |
|
2977 |
have "span (f ` B) \<subseteq> span C" |
|
2978 |
by (metis \<open>fb ` B \<subseteq> C\<close> ffb image_cong span_mono) |
|
2979 |
then have "f ` S \<subseteq> T" |
|
2980 |
unfolding \<open>span B = S\<close> \<open>span C = T\<close> span_linear_image[OF \<open>linear f\<close>] . |
|
2981 |
have [simp]: "\<And>x. x \<in> B \<Longrightarrow> norm (fb x) = norm x" |
|
2982 |
using B1 C1 \<open>fb ` B \<subseteq> C\<close> by auto |
|
2983 |
have "norm (f x) = norm x" if "x \<in> S" for x |
|
2984 |
proof - |
|
2985 |
interpret linear f by fact |
|
2986 |
obtain a where x: "x = (\<Sum>v \<in> B. a v *\<^sub>R v)" |
|
2987 |
using \<open>finite B\<close> \<open>span B = S\<close> \<open>x \<in> S\<close> span_finite by fastforce |
|
2988 |
have "norm (f x)^2 = norm (\<Sum>v\<in>B. a v *\<^sub>R fb v)^2" by (simp add: sum scale ffb x) |
|
2989 |
also have "\<dots> = (\<Sum>v\<in>B. norm ((a v *\<^sub>R fb v))^2)" |
|
71769 | 2990 |
proof (rule norm_sum_Pythagorean [OF \<open>finite B\<close>]) |
2991 |
show "pairwise (\<lambda>v j. orthogonal (a v *\<^sub>R fb v) (a j *\<^sub>R fb j)) B" |
|
2992 |
by (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) |
|
2993 |
qed |
|
69620 | 2994 |
also have "\<dots> = norm x ^2" |
2995 |
by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \<open>finite B\<close>]) |
|
2996 |
finally show ?thesis |
|
2997 |
by (simp add: norm_eq_sqrt_inner) |
|
2998 |
qed |
|
2999 |
then show ?thesis |
|
78457 | 3000 |
by (meson \<open>f ` S \<subseteq> T\<close> \<open>linear f\<close> image_subset_iff_funcset that) |
69620 | 3001 |
qed |
3002 |
||
3003 |
proposition isometries_subspaces: |
|
3004 |
fixes S :: "'a::euclidean_space set" |
|
3005 |
and T :: "'b::euclidean_space set" |
|
3006 |
assumes S: "subspace S" |
|
3007 |
and T: "subspace T" |
|
3008 |
and d: "dim S = dim T" |
|
3009 |
obtains f g where "linear f" "linear g" "f ` S = T" "g ` T = S" |
|
3010 |
"\<And>x. x \<in> S \<Longrightarrow> norm(f x) = norm x" |
|
3011 |
"\<And>x. x \<in> T \<Longrightarrow> norm(g x) = norm x" |
|
3012 |
"\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" |
|
3013 |
"\<And>x. x \<in> T \<Longrightarrow> f(g x) = x" |
|
3014 |
proof - |
|
3015 |
obtain B where "B \<subseteq> S" and Borth: "pairwise orthogonal B" |
|
3016 |
and B1: "\<And>x. x \<in> B \<Longrightarrow> norm x = 1" |
|
3017 |
and "independent B" "finite B" "card B = dim S" "span B = S" |
|
78516
56a408fa2440
substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents:
78474
diff
changeset
|
3018 |
by (metis orthonormal_basis_subspace [OF S] independent_imp_finite) |
69620 | 3019 |
obtain C where "C \<subseteq> T" and Corth: "pairwise orthogonal C" |
3020 |
and C1:"\<And>x. x \<in> C \<Longrightarrow> norm x = 1" |
|
3021 |
and "independent C" "finite C" "card C = dim T" "span C = T" |
|
78516
56a408fa2440
substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents:
78474
diff
changeset
|
3022 |
by (metis orthonormal_basis_subspace [OF T] independent_imp_finite) |
69620 | 3023 |
obtain fb where "bij_betw fb B C" |
3024 |
by (metis \<open>finite B\<close> \<open>finite C\<close> bij_betw_iff_card \<open>card B = dim S\<close> \<open>card C = dim T\<close> d) |
|
3025 |
then have pairwise_orth_fb: "pairwise (\<lambda>v j. orthogonal (fb v) (fb j)) B" |
|
72372 | 3026 |
using Corth unfolding pairwise_def inj_on_def bij_betw_def |
3027 |
by (blast intro: orthogonal_clauses) |
|
69620 | 3028 |
obtain f where "linear f" and ffb: "\<And>x. x \<in> B \<Longrightarrow> f x = fb x" |
3029 |
using linear_independent_extend \<open>independent B\<close> by fastforce |
|
3030 |
interpret f: linear f by fact |
|
3031 |
define gb where "gb \<equiv> inv_into B fb" |
|
3032 |
then have pairwise_orth_gb: "pairwise (\<lambda>v j. orthogonal (gb v) (gb j)) C" |
|
72372 | 3033 |
using Borth \<open>bij_betw fb B C\<close> unfolding pairwise_def bij_betw_def by force |
69620 | 3034 |
obtain g where "linear g" and ggb: "\<And>x. x \<in> C \<Longrightarrow> g x = gb x" |
3035 |
using linear_independent_extend \<open>independent C\<close> by fastforce |
|
3036 |
interpret g: linear g by fact |
|
3037 |
have "span (f ` B) \<subseteq> span C" |
|
3038 |
by (metis \<open>bij_betw fb B C\<close> bij_betw_imp_surj_on eq_iff ffb image_cong) |
|
3039 |
then have "f ` S \<subseteq> T" |
|
72372 | 3040 |
unfolding \<open>span B = S\<close> \<open>span C = T\<close> span_linear_image[OF \<open>linear f\<close>] . |
69620 | 3041 |
have [simp]: "\<And>x. x \<in> B \<Longrightarrow> norm (fb x) = norm x" |
3042 |
using B1 C1 \<open>bij_betw fb B C\<close> bij_betw_imp_surj_on by fastforce |
|
3043 |
have f [simp]: "norm (f x) = norm x" "g (f x) = x" if "x \<in> S" for x |
|
3044 |
proof - |
|
3045 |
obtain a where x: "x = (\<Sum>v \<in> B. a v *\<^sub>R v)" |
|
3046 |
using \<open>finite B\<close> \<open>span B = S\<close> \<open>x \<in> S\<close> span_finite by fastforce |
|
3047 |
have "f x = (\<Sum>v \<in> B. f (a v *\<^sub>R v))" |
|
3048 |
using linear_sum [OF \<open>linear f\<close>] x by auto |
|
3049 |
also have "\<dots> = (\<Sum>v \<in> B. a v *\<^sub>R f v)" |
|
3050 |
by (simp add: f.sum f.scale) |
|
3051 |
also have "\<dots> = (\<Sum>v \<in> B. a v *\<^sub>R fb v)" |
|
3052 |
by (simp add: ffb cong: sum.cong) |
|
3053 |
finally have *: "f x = (\<Sum>v\<in>B. a v *\<^sub>R fb v)" . |
|
3054 |
then have "(norm (f x))\<^sup>2 = (norm (\<Sum>v\<in>B. a v *\<^sub>R fb v))\<^sup>2" by simp |
|
3055 |
also have "\<dots> = (\<Sum>v\<in>B. norm ((a v *\<^sub>R fb v))^2)" |
|
71769 | 3056 |
proof (rule norm_sum_Pythagorean [OF \<open>finite B\<close>]) |
3057 |
show "pairwise (\<lambda>v j. orthogonal (a v *\<^sub>R fb v) (a j *\<^sub>R fb j)) B" |
|
3058 |
by (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) |
|
3059 |
qed |
|
69620 | 3060 |
also have "\<dots> = (norm x)\<^sup>2" |
3061 |
by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \<open>finite B\<close>]) |
|
3062 |
finally show "norm (f x) = norm x" |
|
3063 |
by (simp add: norm_eq_sqrt_inner) |
|
3064 |
have "g (f x) = g (\<Sum>v\<in>B. a v *\<^sub>R fb v)" by (simp add: *) |
|
3065 |
also have "\<dots> = (\<Sum>v\<in>B. g (a v *\<^sub>R fb v))" |
|
3066 |
by (simp add: g.sum g.scale) |
|
3067 |
also have "\<dots> = (\<Sum>v\<in>B. a v *\<^sub>R g (fb v))" |
|
3068 |
by (simp add: g.scale) |
|
3069 |
also have "\<dots> = (\<Sum>v\<in>B. a v *\<^sub>R v)" |
|
71769 | 3070 |
proof (rule sum.cong [OF refl]) |
3071 |
show "a x *\<^sub>R g (fb x) = a x *\<^sub>R x" if "x \<in> B" for x |
|
3072 |
using that \<open>bij_betw fb B C\<close> bij_betwE bij_betw_inv_into_left gb_def ggb by fastforce |
|
3073 |
qed |
|
69620 | 3074 |
also have "\<dots> = x" |
3075 |
using x by blast |
|
3076 |
finally show "g (f x) = x" . |
|
3077 |
qed |
|
3078 |
have [simp]: "\<And>x. x \<in> C \<Longrightarrow> norm (gb x) = norm x" |
|
3079 |
by (metis B1 C1 \<open>bij_betw fb B C\<close> bij_betw_imp_surj_on gb_def inv_into_into) |
|
3080 |
have g [simp]: "f (g x) = x" if "x \<in> T" for x |
|
3081 |
proof - |
|
3082 |
obtain a where x: "x = (\<Sum>v \<in> C. a v *\<^sub>R v)" |
|
3083 |
using \<open>finite C\<close> \<open>span C = T\<close> \<open>x \<in> T\<close> span_finite by fastforce |
|
3084 |
have "g x = (\<Sum>v \<in> C. g (a v *\<^sub>R v))" |
|
3085 |
by (simp add: x g.sum) |
|
3086 |
also have "\<dots> = (\<Sum>v \<in> C. a v *\<^sub>R g v)" |
|
3087 |
by (simp add: g.scale) |
|
3088 |
also have "\<dots> = (\<Sum>v \<in> C. a v *\<^sub>R gb v)" |
|
3089 |
by (simp add: ggb cong: sum.cong) |
|
3090 |
finally have "f (g x) = f (\<Sum>v\<in>C. a v *\<^sub>R gb v)" by simp |
|
3091 |
also have "\<dots> = (\<Sum>v\<in>C. f (a v *\<^sub>R gb v))" |
|
3092 |
by (simp add: f.scale f.sum) |
|
3093 |
also have "\<dots> = (\<Sum>v\<in>C. a v *\<^sub>R f (gb v))" |
|
3094 |
by (simp add: f.scale f.sum) |
|
3095 |
also have "\<dots> = (\<Sum>v\<in>C. a v *\<^sub>R v)" |
|
3096 |
using \<open>bij_betw fb B C\<close> |
|
3097 |
by (simp add: bij_betw_def gb_def bij_betw_inv_into_right ffb inv_into_into) |
|
3098 |
also have "\<dots> = x" |
|
3099 |
using x by blast |
|
3100 |
finally show "f (g x) = x" . |
|
3101 |
qed |
|
3102 |
have gim: "g ` T = S" |
|
3103 |
by (metis (full_types) S T \<open>f ` S \<subseteq> T\<close> d dim_eq_span dim_image_le f(2) g.linear_axioms |
|
3104 |
image_iff linear_subspace_image span_eq_iff subset_iff) |
|
3105 |
have fim: "f ` S = T" |
|
3106 |
using \<open>g ` T = S\<close> image_iff by fastforce |
|
3107 |
have [simp]: "norm (g x) = norm x" if "x \<in> T" for x |
|
3108 |
using fim that by auto |
|
3109 |
show ?thesis |
|
71769 | 3110 |
by (rule that [OF \<open>linear f\<close> \<open>linear g\<close>]) (simp_all add: fim gim) |
69620 | 3111 |
qed |
3112 |
||
3113 |
corollary isometry_subspaces: |
|
3114 |
fixes S :: "'a::euclidean_space set" |
|
3115 |
and T :: "'b::euclidean_space set" |
|
3116 |
assumes S: "subspace S" |
|
3117 |
and T: "subspace T" |
|
3118 |
and d: "dim S = dim T" |
|
3119 |
obtains f where "linear f" "f ` S = T" "\<And>x. x \<in> S \<Longrightarrow> norm(f x) = norm x" |
|
3120 |
using isometries_subspaces [OF assms] |
|
3121 |
by metis |
|
3122 |
||
3123 |
corollary isomorphisms_UNIV_UNIV: |
|
3124 |
assumes "DIM('M) = DIM('N)" |
|
3125 |
obtains f::"'M::euclidean_space \<Rightarrow>'N::euclidean_space" and g |
|
3126 |
where "linear f" "linear g" |
|
3127 |
"\<And>x. norm(f x) = norm x" "\<And>y. norm(g y) = norm y" |
|
3128 |
"\<And>x. g (f x) = x" "\<And>y. f(g y) = y" |
|
3129 |
using assms by (auto intro: isometries_subspaces [of "UNIV::'M set" "UNIV::'N set"]) |
|
3130 |
||
3131 |
lemma homeomorphic_subspaces: |
|
3132 |
fixes S :: "'a::euclidean_space set" |
|
3133 |
and T :: "'b::euclidean_space set" |
|
3134 |
assumes S: "subspace S" |
|
3135 |
and T: "subspace T" |
|
3136 |
and d: "dim S = dim T" |
|
3137 |
shows "S homeomorphic T" |
|
3138 |
proof - |
|
3139 |
obtain f g where "linear f" "linear g" "f ` S = T" "g ` T = S" |
|
3140 |
"\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" "\<And>x. x \<in> T \<Longrightarrow> f(g x) = x" |
|
3141 |
by (blast intro: isometries_subspaces [OF assms]) |
|
3142 |
then show ?thesis |
|
72372 | 3143 |
unfolding homeomorphic_def homeomorphism_def |
3144 |
apply (rule_tac x=f in exI, rule_tac x=g in exI) |
|
69620 | 3145 |
apply (auto simp: linear_continuous_on linear_conv_bounded_linear) |
3146 |
done |
|
3147 |
qed |
|
3148 |
||
3149 |
lemma homeomorphic_affine_sets: |
|
3150 |
assumes "affine S" "affine T" "aff_dim S = aff_dim T" |
|
3151 |
shows "S homeomorphic T" |
|
3152 |
proof (cases "S = {} \<or> T = {}") |
|
3153 |
case True with assms aff_dim_empty homeomorphic_empty show ?thesis |
|
3154 |
by metis |
|
3155 |
next |
|
3156 |
case False |
|
3157 |
then obtain a b where ab: "a \<in> S" "b \<in> T" by auto |
|
3158 |
then have ss: "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)" |
|
3159 |
using affine_diffs_subspace assms by blast+ |
|
3160 |
have dd: "dim ((+) (- a) ` S) = dim ((+) (- b) ` T)" |
|
3161 |
using assms ab by (simp add: aff_dim_eq_dim [OF hull_inc] image_def) |
|
3162 |
have "S homeomorphic ((+) (- a) ` S)" |
|
69768 | 3163 |
by (fact homeomorphic_translation) |
69620 | 3164 |
also have "\<dots> homeomorphic ((+) (- b) ` T)" |
3165 |
by (rule homeomorphic_subspaces [OF ss dd]) |
|
3166 |
also have "\<dots> homeomorphic T" |
|
69768 | 3167 |
using homeomorphic_translation [of T "- b"] by (simp add: homeomorphic_sym [of T]) |
69620 | 3168 |
finally show ?thesis . |
3169 |
qed |
|
3170 |
||
3171 |
||
3172 |
subsection\<open>Retracts, in a general sense, preserve (co)homotopic triviality)\<close> |
|
3173 |
||
70136 | 3174 |
locale\<^marker>\<open>tag important\<close> Retracts = |
77684 | 3175 |
fixes S h t k |
3176 |
assumes conth: "continuous_on S h" |
|
3177 |
and imh: "h ` S = t" |
|
69620 | 3178 |
and contk: "continuous_on t k" |
78457 | 3179 |
and imk: "k \<in> t \<rightarrow> S" |
69620 | 3180 |
and idhk: "\<And>y. y \<in> t \<Longrightarrow> h(k y) = y" |
3181 |
||
3182 |
begin |
|
3183 |
||
3184 |
lemma homotopically_trivial_retraction_gen: |
|
78457 | 3185 |
assumes P: "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> t; Q f\<rbrakk> \<Longrightarrow> P(k \<circ> f)" |
3186 |
and Q: "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S; P f\<rbrakk> \<Longrightarrow> Q(h \<circ> f)" |
|
71769 | 3187 |
and Qeq: "\<And>h k. (\<And>x. x \<in> U \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" |
78457 | 3188 |
and hom: "\<And>f g. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S; P f; |
3189 |
continuous_on U g; g \<in> U \<rightarrow> S; P g\<rbrakk> |
|
77684 | 3190 |
\<Longrightarrow> homotopic_with_canon P U S f g" |
78457 | 3191 |
and contf: "continuous_on U f" and imf: "f \<in> U \<rightarrow> t" and Qf: "Q f" |
3192 |
and contg: "continuous_on U g" and img: "g \<in> U \<rightarrow> t" and Qg: "Q g" |
|
71769 | 3193 |
shows "homotopic_with_canon Q U t f g" |
69620 | 3194 |
proof - |
71769 | 3195 |
have "continuous_on U (k \<circ> f)" |
78457 | 3196 |
by (meson contf continuous_on_compose continuous_on_subset contk funcset_image imf) |
77684 | 3197 |
moreover have "(k \<circ> f) ` U \<subseteq> S" |
69620 | 3198 |
using imf imk by fastforce |
3199 |
moreover have "P (k \<circ> f)" |
|
3200 |
by (simp add: P Qf contf imf) |
|
71769 | 3201 |
moreover have "continuous_on U (k \<circ> g)" |
78457 | 3202 |
by (meson contg continuous_on_compose continuous_on_subset contk funcset_image img) |
77684 | 3203 |
moreover have "(k \<circ> g) ` U \<subseteq> S" |
69620 | 3204 |
using img imk by fastforce |
3205 |
moreover have "P (k \<circ> g)" |
|
3206 |
by (simp add: P Qg contg img) |
|
77684 | 3207 |
ultimately have "homotopic_with_canon P U S (k \<circ> f) (k \<circ> g)" |
78457 | 3208 |
by (simp add: hom image_subset_iff) |
71769 | 3209 |
then have "homotopic_with_canon Q U t (h \<circ> (k \<circ> f)) (h \<circ> (k \<circ> g))" |
69620 | 3210 |
apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3211 |
using Q conth imh by force+ |
69620 | 3212 |
then show ?thesis |
72372 | 3213 |
proof (rule homotopic_with_eq; simp) |
3214 |
show "\<And>h k. (\<And>x. x \<in> U \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" |
|
3215 |
using Qeq topspace_euclidean_subtopology by blast |
|
3216 |
show "\<And>x. x \<in> U \<Longrightarrow> f x = h (k (f x))" "\<And>x. x \<in> U \<Longrightarrow> g x = h (k (g x))" |
|
78457 | 3217 |
using idhk imf img by fastforce+ |
72372 | 3218 |
qed |
69620 | 3219 |
qed |
3220 |
||
3221 |
lemma homotopically_trivial_retraction_null_gen: |
|
78457 | 3222 |
assumes P: "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> t; Q f\<rbrakk> \<Longrightarrow> P(k \<circ> f)" |
3223 |
and Q: "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S; P f\<rbrakk> \<Longrightarrow> Q(h \<circ> f)" |
|
71769 | 3224 |
and Qeq: "\<And>h k. (\<And>x. x \<in> U \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" |
78457 | 3225 |
and hom: "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S; P f\<rbrakk> |
77684 | 3226 |
\<Longrightarrow> \<exists>c. homotopic_with_canon P U S f (\<lambda>x. c)" |
78457 | 3227 |
and contf: "continuous_on U f" and imf:"f \<in> U \<rightarrow> t" and Qf: "Q f" |
71769 | 3228 |
obtains c where "homotopic_with_canon Q U t f (\<lambda>x. c)" |
69620 | 3229 |
proof - |
71769 | 3230 |
have feq: "\<And>x. x \<in> U \<Longrightarrow> (h \<circ> (k \<circ> f)) x = f x" using idhk imf by auto |
3231 |
have "continuous_on U (k \<circ> f)" |
|
78457 | 3232 |
by (meson contf continuous_on_compose continuous_on_subset contk funcset_image imf) |
3233 |
moreover have "(k \<circ> f) \<in> U \<rightarrow> S" |
|
69620 | 3234 |
using imf imk by fastforce |
3235 |
moreover have "P (k \<circ> f)" |
|
3236 |
by (simp add: P Qf contf imf) |
|
77684 | 3237 |
ultimately obtain c where "homotopic_with_canon P U S (k \<circ> f) (\<lambda>x. c)" |
69620 | 3238 |
by (metis hom) |
71769 | 3239 |
then have "homotopic_with_canon Q U t (h \<circ> (k \<circ> f)) (h \<circ> (\<lambda>x. c))" |
69620 | 3240 |
apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3241 |
using Q conth imh by force+ |
71769 | 3242 |
then have "homotopic_with_canon Q U t f (\<lambda>x. h c)" |
3243 |
proof (rule homotopic_with_eq) |
|
3244 |
show "\<And>x. x \<in> topspace (top_of_set U) \<Longrightarrow> f x = (h \<circ> (k \<circ> f)) x" |
|
3245 |
using feq by auto |
|
3246 |
show "\<And>h k. (\<And>x. x \<in> topspace (top_of_set U) \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" |
|
3247 |
using Qeq topspace_euclidean_subtopology by blast |
|
3248 |
qed auto |
|
69620 | 3249 |
then show ?thesis |
71769 | 3250 |
using that by blast |
69620 | 3251 |
qed |
3252 |
||
3253 |
lemma cohomotopically_trivial_retraction_gen: |
|
78457 | 3254 |
assumes P: "\<And>f. \<lbrakk>continuous_on t f; f \<in> t \<rightarrow> U; Q f\<rbrakk> \<Longrightarrow> P(f \<circ> h)" |
3255 |
and Q: "\<And>f. \<lbrakk>continuous_on S f; f \<in> S \<rightarrow> U; P f\<rbrakk> \<Longrightarrow> Q(f \<circ> k)" |
|
69620 | 3256 |
and Qeq: "\<And>h k. (\<And>x. x \<in> t \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" |
78457 | 3257 |
and hom: "\<And>f g. \<lbrakk>continuous_on S f; f \<in> S \<rightarrow> U; P f; |
3258 |
continuous_on S g; g \<in> S \<rightarrow> U; P g\<rbrakk> |
|
77684 | 3259 |
\<Longrightarrow> homotopic_with_canon P S U f g" |
78457 | 3260 |
and contf: "continuous_on t f" and imf: "f \<in> t \<rightarrow> U" and Qf: "Q f" |
3261 |
and contg: "continuous_on t g" and img: "g \<in> t \<rightarrow> U" and Qg: "Q g" |
|
71769 | 3262 |
shows "homotopic_with_canon Q t U f g" |
69620 | 3263 |
proof - |
3264 |
have feq: "\<And>x. x \<in> t \<Longrightarrow> (f \<circ> h \<circ> k) x = f x" using idhk imf by auto |
|
3265 |
have geq: "\<And>x. x \<in> t \<Longrightarrow> (g \<circ> h \<circ> k) x = g x" using idhk img by auto |
|
77684 | 3266 |
have "continuous_on S (f \<circ> h)" |
69620 | 3267 |
using contf conth continuous_on_compose imh by blast |
78457 | 3268 |
moreover have "(f \<circ> h) \<in> S \<rightarrow> U" |
69620 | 3269 |
using imf imh by fastforce |
3270 |
moreover have "P (f \<circ> h)" |
|
3271 |
by (simp add: P Qf contf imf) |
|
77684 | 3272 |
moreover have "continuous_on S (g \<circ> h)" |
69620 | 3273 |
using contg continuous_on_compose continuous_on_subset conth imh by blast |
78457 | 3274 |
moreover have "(g \<circ> h) \<in> S \<rightarrow> U" |
69620 | 3275 |
using img imh by fastforce |
3276 |
moreover have "P (g \<circ> h)" |
|
3277 |
by (simp add: P Qg contg img) |
|
77684 | 3278 |
ultimately have "homotopic_with_canon P S U (f \<circ> h) (g \<circ> h)" |
78457 | 3279 |
by (simp add: hom) |
71769 | 3280 |
then have "homotopic_with_canon Q t U (f \<circ> h \<circ> k) (g \<circ> h \<circ> k)" |
69620 | 3281 |
apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3282 |
using Q contk imk by force+ |
69620 | 3283 |
then show ?thesis |
71746 | 3284 |
proof (rule homotopic_with_eq) |
3285 |
show "f x = (f \<circ> h \<circ> k) x" "g x = (g \<circ> h \<circ> k) x" |
|
3286 |
if "x \<in> topspace (top_of_set t)" for x |
|
3287 |
using feq geq that by force+ |
|
3288 |
qed (use Qeq topspace_euclidean_subtopology in blast) |
|
69620 | 3289 |
qed |
3290 |
||
3291 |
lemma cohomotopically_trivial_retraction_null_gen: |
|
78457 | 3292 |
assumes P: "\<And>f. \<lbrakk>continuous_on t f; f \<in> t \<rightarrow> U; Q f\<rbrakk> \<Longrightarrow> P(f \<circ> h)" |
3293 |
and Q: "\<And>f. \<lbrakk>continuous_on S f; f \<in> S \<rightarrow> U; P f\<rbrakk> \<Longrightarrow> Q(f \<circ> k)" |
|
69620 | 3294 |
and Qeq: "\<And>h k. (\<And>x. x \<in> t \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" |
78457 | 3295 |
and hom: "\<And>f g. \<lbrakk>continuous_on S f; f \<in> S \<rightarrow> U; P f\<rbrakk> |
77684 | 3296 |
\<Longrightarrow> \<exists>c. homotopic_with_canon P S U f (\<lambda>x. c)" |
78457 | 3297 |
and contf: "continuous_on t f" and imf: "f \<in> t \<rightarrow> U" and Qf: "Q f" |
71769 | 3298 |
obtains c where "homotopic_with_canon Q t U f (\<lambda>x. c)" |
69620 | 3299 |
proof - |
3300 |
have feq: "\<And>x. x \<in> t \<Longrightarrow> (f \<circ> h \<circ> k) x = f x" using idhk imf by auto |
|
77684 | 3301 |
have "continuous_on S (f \<circ> h)" |
69620 | 3302 |
using contf conth continuous_on_compose imh by blast |
78457 | 3303 |
moreover have "(f \<circ> h) \<in> S \<rightarrow> U" |
69620 | 3304 |
using imf imh by fastforce |
3305 |
moreover have "P (f \<circ> h)" |
|
3306 |
by (simp add: P Qf contf imf) |
|
77684 | 3307 |
ultimately obtain c where "homotopic_with_canon P S U (f \<circ> h) (\<lambda>x. c)" |
69620 | 3308 |
by (metis hom) |
71769 | 3309 |
then have \<section>: "homotopic_with_canon Q t U (f \<circ> h \<circ> k) ((\<lambda>x. c) \<circ> k)" |
71746 | 3310 |
proof (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) |
77684 | 3311 |
show "\<And>h. \<lbrakk>continuous_map (top_of_set S) (top_of_set U) h; P h\<rbrakk> \<Longrightarrow> Q (h \<circ> k)" |
71746 | 3312 |
using Q by auto |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3313 |
qed (use contk imk in force)+ |
71769 | 3314 |
moreover have "homotopic_with_canon Q t U f (\<lambda>x. c)" |
71746 | 3315 |
using homotopic_with_eq [OF \<section>] feq Qeq by fastforce |
3316 |
ultimately show ?thesis |
|
3317 |
using that by blast |
|
69620 | 3318 |
qed |
3319 |
||
3320 |
end |
|
3321 |
||
3322 |
lemma simply_connected_retraction_gen: |
|
3323 |
shows "\<lbrakk>simply_connected S; continuous_on S h; h ` S = T; |
|
78457 | 3324 |
continuous_on T k; k \<in> T \<rightarrow> S; \<And>y. y \<in> T \<Longrightarrow> h(k y) = y\<rbrakk> |
69620 | 3325 |
\<Longrightarrow> simply_connected T" |
3326 |
apply (simp add: simply_connected_def path_def path_image_def homotopic_loops_def, clarify) |
|
3327 |
apply (rule Retracts.homotopically_trivial_retraction_gen |
|
3328 |
[of S h _ k _ "\<lambda>p. pathfinish p = pathstart p" "\<lambda>p. pathfinish p = pathstart p"]) |
|
78457 | 3329 |
apply (simp_all add: Retracts_def pathfinish_def pathstart_def image_subset_iff_funcset) |
69620 | 3330 |
done |
3331 |
||
3332 |
lemma homeomorphic_simply_connected: |
|
3333 |
"\<lbrakk>S homeomorphic T; simply_connected S\<rbrakk> \<Longrightarrow> simply_connected T" |
|
3334 |
by (auto simp: homeomorphic_def homeomorphism_def intro: simply_connected_retraction_gen) |
|
3335 |
||
3336 |
lemma homeomorphic_simply_connected_eq: |
|
3337 |
"S homeomorphic T \<Longrightarrow> (simply_connected S \<longleftrightarrow> simply_connected T)" |
|
3338 |
by (metis homeomorphic_simply_connected homeomorphic_sym) |
|
3339 |
||
3340 |
||
3341 |
subsection\<open>Homotopy equivalence\<close> |
|
3342 |
||
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3343 |
subsection\<open>Homotopy equivalence of topological spaces.\<close> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3344 |
|
70136 | 3345 |
definition\<^marker>\<open>tag important\<close> homotopy_equivalent_space |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80052
diff
changeset
|
3346 |
(infix \<open>homotopy'_equivalent'_space\<close> 50) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3347 |
where "X homotopy_equivalent_space Y \<equiv> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3348 |
(\<exists>f g. continuous_map X Y f \<and> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3349 |
continuous_map Y X g \<and> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3350 |
homotopic_with (\<lambda>x. True) X X (g \<circ> f) id \<and> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3351 |
homotopic_with (\<lambda>x. True) Y Y (f \<circ> g) id)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3352 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3353 |
lemma homeomorphic_imp_homotopy_equivalent_space: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3354 |
"X homeomorphic_space Y \<Longrightarrow> X homotopy_equivalent_space Y" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3355 |
unfolding homeomorphic_space_def homotopy_equivalent_space_def |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3356 |
apply (erule ex_forward)+ |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3357 |
by (simp add: homotopic_with_equal homotopic_with_sym homeomorphic_maps_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3358 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3359 |
lemma homotopy_equivalent_space_refl: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3360 |
"X homotopy_equivalent_space X" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3361 |
by (simp add: homeomorphic_imp_homotopy_equivalent_space homeomorphic_space_refl) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3362 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3363 |
lemma homotopy_equivalent_space_sym: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3364 |
"X homotopy_equivalent_space Y \<longleftrightarrow> Y homotopy_equivalent_space X" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3365 |
by (meson homotopy_equivalent_space_def) |
69620 | 3366 |
|
3367 |
lemma homotopy_eqv_trans [trans]: |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3368 |
assumes 1: "X homotopy_equivalent_space Y" and 2: "Y homotopy_equivalent_space U" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3369 |
shows "X homotopy_equivalent_space U" |
69620 | 3370 |
proof - |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3371 |
obtain f1 g1 where f1: "continuous_map X Y f1" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3372 |
and g1: "continuous_map Y X g1" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3373 |
and hom1: "homotopic_with (\<lambda>x. True) X X (g1 \<circ> f1) id" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3374 |
"homotopic_with (\<lambda>x. True) Y Y (f1 \<circ> g1) id" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3375 |
using 1 by (auto simp: homotopy_equivalent_space_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3376 |
obtain f2 g2 where f2: "continuous_map Y U f2" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3377 |
and g2: "continuous_map U Y g2" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3378 |
and hom2: "homotopic_with (\<lambda>x. True) Y Y (g2 \<circ> f2) id" |
69620 | 3379 |
"homotopic_with (\<lambda>x. True) U U (f2 \<circ> g2) id" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3380 |
using 2 by (auto simp: homotopy_equivalent_space_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3381 |
have "homotopic_with (\<lambda>f. True) X Y (g2 \<circ> f2 \<circ> f1) (id \<circ> f1)" |
71745 | 3382 |
using f1 hom2(1) homotopic_with_compose_continuous_map_right by metis |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3383 |
then have "homotopic_with (\<lambda>f. True) X Y (g2 \<circ> (f2 \<circ> f1)) (id \<circ> f1)" |
69620 | 3384 |
by (simp add: o_assoc) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3385 |
then have "homotopic_with (\<lambda>x. True) X X |
69620 | 3386 |
(g1 \<circ> (g2 \<circ> (f2 \<circ> f1))) (g1 \<circ> (id \<circ> f1))" |
71745 | 3387 |
by (simp add: g1 homotopic_with_compose_continuous_map_left) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3388 |
moreover have "homotopic_with (\<lambda>x. True) X X (g1 \<circ> id \<circ> f1) id" |
69620 | 3389 |
using hom1 by simp |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3390 |
ultimately have SS: "homotopic_with (\<lambda>x. True) X X (g1 \<circ> g2 \<circ> (f2 \<circ> f1)) id" |
71769 | 3391 |
by (metis comp_assoc homotopic_with_trans id_comp) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3392 |
have "homotopic_with (\<lambda>f. True) U Y (f1 \<circ> g1 \<circ> g2) (id \<circ> g2)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3393 |
using g2 hom1(2) homotopic_with_compose_continuous_map_right by fastforce |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3394 |
then have "homotopic_with (\<lambda>f. True) U Y (f1 \<circ> (g1 \<circ> g2)) (id \<circ> g2)" |
69620 | 3395 |
by (simp add: o_assoc) |
3396 |
then have "homotopic_with (\<lambda>x. True) U U |
|
3397 |
(f2 \<circ> (f1 \<circ> (g1 \<circ> g2))) (f2 \<circ> (id \<circ> g2))" |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3398 |
by (simp add: f2 homotopic_with_compose_continuous_map_left) |
69620 | 3399 |
moreover have "homotopic_with (\<lambda>x. True) U U (f2 \<circ> id \<circ> g2) id" |
3400 |
using hom2 by simp |
|
3401 |
ultimately have UU: "homotopic_with (\<lambda>x. True) U U (f2 \<circ> f1 \<circ> (g1 \<circ> g2)) id" |
|
71769 | 3402 |
by (simp add: fun.map_comp hom2(2) homotopic_with_trans) |
69620 | 3403 |
show ?thesis |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3404 |
unfolding homotopy_equivalent_space_def |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3405 |
by (blast intro: f1 f2 g1 g2 continuous_map_compose SS UU) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3406 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3407 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3408 |
lemma deformation_retraction_imp_homotopy_equivalent_space: |
77684 | 3409 |
"\<lbrakk>homotopic_with (\<lambda>x. True) X X (S \<circ> r) id; retraction_maps X Y r S\<rbrakk> |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3410 |
\<Longrightarrow> X homotopy_equivalent_space Y" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3411 |
unfolding homotopy_equivalent_space_def retraction_maps_def |
71769 | 3412 |
using homotopic_with_id2 by fastforce |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3413 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3414 |
lemma deformation_retract_imp_homotopy_equivalent_space: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3415 |
"\<lbrakk>homotopic_with (\<lambda>x. True) X X r id; retraction_maps X Y r id\<rbrakk> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3416 |
\<Longrightarrow> X homotopy_equivalent_space Y" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3417 |
using deformation_retraction_imp_homotopy_equivalent_space by force |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3418 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3419 |
lemma deformation_retract_of_space: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3420 |
"S \<subseteq> topspace X \<and> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3421 |
(\<exists>r. homotopic_with (\<lambda>x. True) X X id r \<and> retraction_maps X (subtopology X S) r id) \<longleftrightarrow> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3422 |
S retract_of_space X \<and> (\<exists>f. homotopic_with (\<lambda>x. True) X X id f \<and> f ` (topspace X) \<subseteq> S)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3423 |
proof (cases "S \<subseteq> topspace X") |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3424 |
case True |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3425 |
moreover have "(\<exists>r. homotopic_with (\<lambda>x. True) X X id r \<and> retraction_maps X (subtopology X S) r id) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3426 |
\<longleftrightarrow> (S retract_of_space X \<and> (\<exists>f. homotopic_with (\<lambda>x. True) X X id f \<and> f ` topspace X \<subseteq> S))" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3427 |
unfolding retract_of_space_def |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3428 |
proof safe |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3429 |
fix f r |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3430 |
assume f: "homotopic_with (\<lambda>x. True) X X id f" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3431 |
and fS: "f ` topspace X \<subseteq> S" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3432 |
and r: "continuous_map X (subtopology X S) r" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3433 |
and req: "\<forall>x\<in>S. r x = x" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3434 |
show "\<exists>r. homotopic_with (\<lambda>x. True) X X id r \<and> retraction_maps X (subtopology X S) r id" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3435 |
proof (intro exI conjI) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3436 |
have "homotopic_with (\<lambda>x. True) X X f r" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3437 |
proof (rule homotopic_with_eq) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3438 |
show "homotopic_with (\<lambda>x. True) X X (r \<circ> f) (r \<circ> id)" |
71745 | 3439 |
by (metis continuous_map_into_fulltopology f homotopic_with_compose_continuous_map_left homotopic_with_symD r) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3440 |
show "f x = (r \<circ> f) x" if "x \<in> topspace X" for x |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3441 |
using that fS req by auto |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3442 |
qed auto |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3443 |
then show "homotopic_with (\<lambda>x. True) X X id r" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3444 |
by (rule homotopic_with_trans [OF f]) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3445 |
next |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3446 |
show "retraction_maps X (subtopology X S) r id" |
71172 | 3447 |
by (simp add: r req retraction_maps_def) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3448 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3449 |
qed (use True in \<open>auto simp: retraction_maps_def topspace_subtopology_subset continuous_map_in_subtopology\<close>) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3450 |
ultimately show ?thesis by simp |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3451 |
qed (auto simp: retract_of_space_def retraction_maps_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3452 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3453 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3454 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3455 |
subsection\<open>Contractible spaces\<close> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3456 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3457 |
text\<open>The definition (which agrees with "contractible" on subsets of Euclidean space) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3458 |
is a little cryptic because we don't in fact assume that the constant "a" is in the space. |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3459 |
This forces the convention that the empty space / set is contractible, avoiding some special cases. \<close> |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3460 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3461 |
definition contractible_space where |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3462 |
"contractible_space X \<equiv> \<exists>a. homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3463 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3464 |
lemma contractible_space_top_of_set [simp]:"contractible_space (top_of_set S) \<longleftrightarrow> contractible S" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3465 |
by (auto simp: contractible_space_def contractible_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3466 |
|
78336 | 3467 |
lemma contractible_space_empty [simp]: |
3468 |
"contractible_space trivial_topology" |
|
71745 | 3469 |
unfolding contractible_space_def homotopic_with_def |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3470 |
apply (rule_tac x=undefined in exI) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3471 |
apply (rule_tac x="\<lambda>(t,x). if t = 0 then x else undefined" in exI) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3472 |
apply (auto simp: continuous_map_on_empty) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3473 |
done |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3474 |
|
78336 | 3475 |
lemma contractible_space_singleton [simp]: |
3476 |
"contractible_space (discrete_topology{a})" |
|
71745 | 3477 |
unfolding contractible_space_def homotopic_with_def |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3478 |
apply (rule_tac x=a in exI) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3479 |
apply (rule_tac x="\<lambda>(t,x). if t = 0 then x else a" in exI) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3480 |
apply (auto intro: continuous_map_eq [where f = "\<lambda>z. a"]) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3481 |
done |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3482 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3483 |
lemma contractible_space_subset_singleton: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3484 |
"topspace X \<subseteq> {a} \<Longrightarrow> contractible_space X" |
78336 | 3485 |
by (metis contractible_space_empty contractible_space_singleton null_topspace_iff_trivial subset_singletonD subtopology_eq_discrete_topology_sing) |
3486 |
||
3487 |
lemma contractible_space_subtopology_singleton [simp]: |
|
3488 |
"contractible_space (subtopology X {a})" |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3489 |
by (meson contractible_space_subset_singleton insert_subset path_connectedin_singleton path_connectedin_subtopology subsetI) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3490 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3491 |
lemma contractible_space: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3492 |
"contractible_space X \<longleftrightarrow> |
78336 | 3493 |
X = trivial_topology \<or> |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3494 |
(\<exists>a \<in> topspace X. homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a))" |
78336 | 3495 |
proof (cases "X = trivial_topology") |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3496 |
case False |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3497 |
then show ?thesis |
71745 | 3498 |
using homotopic_with_imp_continuous_maps by (fastforce simp: contractible_space_def) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3499 |
qed (simp add: contractible_space_empty) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3500 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3501 |
lemma contractible_imp_path_connected_space: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3502 |
assumes "contractible_space X" shows "path_connected_space X" |
78336 | 3503 |
proof (cases "X = trivial_topology") |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3504 |
case False |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3505 |
have *: "path_connected_space X" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3506 |
if "a \<in> topspace X" and conth: "continuous_map (prod_topology (top_of_set {0..1}) X) X h" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3507 |
and h: "\<forall>x. h (0, x) = x" "\<forall>x. h (1, x) = a" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3508 |
for a and h :: "real \<times> 'a \<Rightarrow> 'a" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3509 |
proof - |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3510 |
have "path_component_of X b a" if "b \<in> topspace X" for b |
71745 | 3511 |
unfolding path_component_of_def |
3512 |
proof (intro exI conjI) |
|
3513 |
let ?g = "h \<circ> (\<lambda>x. (x,b))" |
|
3514 |
show "pathin X ?g" |
|
3515 |
unfolding pathin_def |
|
71769 | 3516 |
proof (rule continuous_map_compose [OF _ conth]) |
3517 |
show "continuous_map (top_of_set {0..1}) (prod_topology (top_of_set {0..1}) X) (\<lambda>x. (x, b))" |
|
3518 |
using that by (auto intro!: continuous_intros) |
|
3519 |
qed |
|
71745 | 3520 |
qed (use h in auto) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3521 |
then show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3522 |
by (metis path_component_of_equiv path_connected_space_iff_path_component) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3523 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3524 |
show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3525 |
using assms False by (auto simp: contractible_space homotopic_with_def *) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3526 |
qed (simp add: path_connected_space_topspace_empty) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3527 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3528 |
lemma contractible_imp_connected_space: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3529 |
"contractible_space X \<Longrightarrow> connected_space X" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3530 |
by (simp add: contractible_imp_path_connected_space path_connected_imp_connected_space) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3531 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3532 |
lemma contractible_space_alt: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3533 |
"contractible_space X \<longleftrightarrow> (\<forall>a \<in> topspace X. homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a))" (is "?lhs = ?rhs") |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3534 |
proof |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3535 |
assume X: ?lhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3536 |
then obtain a where a: "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3537 |
by (auto simp: contractible_space_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3538 |
show ?rhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3539 |
proof |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3540 |
show "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. b)" if "b \<in> topspace X" for b |
71745 | 3541 |
proof (rule homotopic_with_trans [OF a]) |
3542 |
show "homotopic_with (\<lambda>x. True) X X (\<lambda>x. a) (\<lambda>x. b)" |
|
3543 |
using homotopic_constant_maps path_connected_space_imp_path_component_of |
|
78336 | 3544 |
by (metis X a contractible_imp_path_connected_space homotopic_with_sym homotopic_with_trans path_component_of_equiv that) |
71745 | 3545 |
qed |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3546 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3547 |
next |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3548 |
assume R: ?rhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3549 |
then show ?lhs |
78336 | 3550 |
using contractible_space_def by fastforce |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3551 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3552 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3553 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3554 |
lemma compose_const [simp]: "f \<circ> (\<lambda>x. a) = (\<lambda>x. f a)" "(\<lambda>x. a) \<circ> g = (\<lambda>x. a)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3555 |
by (simp_all add: o_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3556 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3557 |
lemma nullhomotopic_through_contractible_space: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3558 |
assumes f: "continuous_map X Y f" and g: "continuous_map Y Z g" and Y: "contractible_space Y" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3559 |
obtains c where "homotopic_with (\<lambda>h. True) X Z (g \<circ> f) (\<lambda>x. c)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3560 |
proof - |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3561 |
obtain b where b: "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. b)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3562 |
using Y by (auto simp: contractible_space_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3563 |
show thesis |
71745 | 3564 |
using homotopic_with_compose_continuous_map_right |
3565 |
[OF homotopic_with_compose_continuous_map_left [OF b g] f] |
|
78474 | 3566 |
by (force simp: that) |
69620 | 3567 |
qed |
3568 |
||
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3569 |
lemma nullhomotopic_into_contractible_space: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3570 |
assumes f: "continuous_map X Y f" and Y: "contractible_space Y" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3571 |
obtains c where "homotopic_with (\<lambda>h. True) X Y f (\<lambda>x. c)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3572 |
using nullhomotopic_through_contractible_space [OF f _ Y] |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3573 |
by (metis continuous_map_id id_comp) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3574 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3575 |
lemma nullhomotopic_from_contractible_space: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3576 |
assumes f: "continuous_map X Y f" and X: "contractible_space X" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3577 |
obtains c where "homotopic_with (\<lambda>h. True) X Y f (\<lambda>x. c)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3578 |
using nullhomotopic_through_contractible_space [OF _ f X] |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3579 |
by (metis comp_id continuous_map_id) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3580 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3581 |
lemma homotopy_dominated_contractibility: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3582 |
assumes f: "continuous_map X Y f" and g: "continuous_map Y X g" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3583 |
and hom: "homotopic_with (\<lambda>x. True) Y Y (f \<circ> g) id" and X: "contractible_space X" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3584 |
shows "contractible_space Y" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3585 |
proof - |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3586 |
obtain c where c: "homotopic_with (\<lambda>h. True) X Y f (\<lambda>x. c)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3587 |
using nullhomotopic_from_contractible_space [OF f X] . |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3588 |
have "homotopic_with (\<lambda>x. True) Y Y (f \<circ> g) (\<lambda>x. c)" |
71745 | 3589 |
using homotopic_with_compose_continuous_map_right [OF c g] by fastforce |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3590 |
then have "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. c)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3591 |
using homotopic_with_trans [OF _ hom] homotopic_with_symD by blast |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3592 |
then show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3593 |
unfolding contractible_space_def .. |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3594 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3595 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3596 |
lemma homotopy_equivalent_space_contractibility: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3597 |
"X homotopy_equivalent_space Y \<Longrightarrow> (contractible_space X \<longleftrightarrow> contractible_space Y)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3598 |
unfolding homotopy_equivalent_space_def |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3599 |
by (blast intro: homotopy_dominated_contractibility) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3600 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3601 |
lemma homeomorphic_space_contractibility: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3602 |
"X homeomorphic_space Y |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3603 |
\<Longrightarrow> (contractible_space X \<longleftrightarrow> contractible_space Y)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3604 |
by (simp add: homeomorphic_imp_homotopy_equivalent_space homotopy_equivalent_space_contractibility) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3605 |
|
78127
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3606 |
lemma homotopic_through_contractible_space: |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3607 |
"continuous_map X Y f \<and> |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3608 |
continuous_map X Y f' \<and> |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3609 |
continuous_map Y Z g \<and> |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3610 |
continuous_map Y Z g' \<and> |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3611 |
contractible_space Y \<and> path_connected_space Z |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3612 |
\<Longrightarrow> homotopic_with (\<lambda>h. True) X Z (g \<circ> f) (g' \<circ> f')" |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3613 |
using nullhomotopic_through_contractible_space [of X Y f Z g] |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3614 |
using nullhomotopic_through_contractible_space [of X Y f' Z g'] |
78336 | 3615 |
by (smt (verit) continuous_map_const homotopic_constant_maps homotopic_with_imp_continuous_maps |
3616 |
homotopic_with_symD homotopic_with_trans path_connected_space_imp_path_component_of) |
|
78127
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3617 |
|
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3618 |
lemma homotopic_from_contractible_space: |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3619 |
"continuous_map X Y f \<and> continuous_map X Y g \<and> |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3620 |
contractible_space X \<and> path_connected_space Y |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3621 |
\<Longrightarrow> homotopic_with (\<lambda>x. True) X Y f g" |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3622 |
by (metis comp_id continuous_map_id homotopic_through_contractible_space) |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3623 |
|
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3624 |
lemma homotopic_into_contractible_space: |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3625 |
"continuous_map X Y f \<and> continuous_map X Y g \<and> |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3626 |
contractible_space Y |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3627 |
\<Longrightarrow> homotopic_with (\<lambda>x. True) X Y f g" |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3628 |
by (metis continuous_map_id contractible_imp_path_connected_space homotopic_through_contractible_space id_comp) |
24b70433c2e8
New HOL Light material on metric spaces and topological spaces
paulson <lp15@cam.ac.uk>
parents:
77929
diff
changeset
|
3629 |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3630 |
lemma contractible_eq_homotopy_equivalent_singleton_subtopology: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3631 |
"contractible_space X \<longleftrightarrow> |
78336 | 3632 |
X = trivial_topology \<or> (\<exists>a \<in> topspace X. X homotopy_equivalent_space (subtopology X {a}))"(is "?lhs = ?rhs") |
3633 |
proof (cases "X = trivial_topology") |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3634 |
case False |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3635 |
show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3636 |
proof |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3637 |
assume ?lhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3638 |
then obtain a where a: "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3639 |
by (auto simp: contractible_space_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3640 |
then have "a \<in> topspace X" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3641 |
by (metis False continuous_map_const homotopic_with_imp_continuous_maps) |
72372 | 3642 |
then have "homotopic_with (\<lambda>x. True) (subtopology X {a}) (subtopology X {a}) id (\<lambda>x. a)" |
3643 |
using connectedin_absolute connectedin_sing contractible_space_alt contractible_space_subtopology_singleton by fastforce |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3644 |
then have "X homotopy_equivalent_space subtopology X {a}" |
72372 | 3645 |
unfolding homotopy_equivalent_space_def using \<open>a \<in> topspace X\<close> |
3646 |
by (metis (full_types) a comp_id continuous_map_const continuous_map_id_subt empty_subsetI homotopic_with_symD |
|
3647 |
id_comp insertI1 insert_subset topspace_subtopology_subset) |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3648 |
with \<open>a \<in> topspace X\<close> show ?rhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3649 |
by blast |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3650 |
next |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3651 |
assume ?rhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3652 |
then show ?lhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3653 |
by (meson False contractible_space_subtopology_singleton homotopy_equivalent_space_contractibility) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3654 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3655 |
qed (simp add: contractible_space_empty) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3656 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3657 |
lemma contractible_space_retraction_map_image: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3658 |
assumes "retraction_map X Y f" and X: "contractible_space X" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3659 |
shows "contractible_space Y" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3660 |
proof - |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3661 |
obtain g where f: "continuous_map X Y f" and g: "continuous_map Y X g" and fg: "\<forall>y \<in> topspace Y. f(g y) = y" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3662 |
using assms by (auto simp: retraction_map_def retraction_maps_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3663 |
obtain a where a: "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3664 |
using X by (auto simp: contractible_space_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3665 |
have "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. f a)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3666 |
proof (rule homotopic_with_eq) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3667 |
show "homotopic_with (\<lambda>x. True) Y Y (f \<circ> id \<circ> g) (f \<circ> (\<lambda>x. a) \<circ> g)" |
71745 | 3668 |
using f g a homotopic_with_compose_continuous_map_left homotopic_with_compose_continuous_map_right by metis |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3669 |
qed (use fg in auto) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3670 |
then show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3671 |
unfolding contractible_space_def by blast |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3672 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3673 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3674 |
lemma contractible_space_prod_topology: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3675 |
"contractible_space(prod_topology X Y) \<longleftrightarrow> |
78336 | 3676 |
X = trivial_topology \<or> Y = trivial_topology \<or> contractible_space X \<and> contractible_space Y" |
3677 |
proof (cases "X = trivial_topology \<or> Y = trivial_topology") |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3678 |
case True |
78336 | 3679 |
then have "(prod_topology X Y) = trivial_topology" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3680 |
by simp |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3681 |
then show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3682 |
by (auto simp: contractible_space_empty) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3683 |
next |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3684 |
case False |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3685 |
have "contractible_space(prod_topology X Y) \<longleftrightarrow> contractible_space X \<and> contractible_space Y" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3686 |
proof safe |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3687 |
assume XY: "contractible_space (prod_topology X Y)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3688 |
with False have "retraction_map (prod_topology X Y) X fst" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3689 |
by (auto simp: contractible_space False retraction_map_fst) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3690 |
then show "contractible_space X" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3691 |
by (rule contractible_space_retraction_map_image [OF _ XY]) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3692 |
have "retraction_map (prod_topology X Y) Y snd" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3693 |
using False XY by (auto simp: contractible_space False retraction_map_snd) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3694 |
then show "contractible_space Y" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3695 |
by (rule contractible_space_retraction_map_image [OF _ XY]) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3696 |
next |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3697 |
assume "contractible_space X" and "contractible_space Y" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3698 |
with False obtain a b |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3699 |
where "a \<in> topspace X" and a: "homotopic_with (\<lambda>x. True) X X id (\<lambda>x. a)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3700 |
and "b \<in> topspace Y" and b: "homotopic_with (\<lambda>x. True) Y Y id (\<lambda>x. b)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3701 |
by (auto simp: contractible_space) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3702 |
with False show "contractible_space (prod_topology X Y)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3703 |
apply (simp add: contractible_space) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3704 |
apply (rule_tac x=a in bexI) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3705 |
apply (rule_tac x=b in bexI) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3706 |
using homotopic_with_prod_topology [OF a b] |
70033 | 3707 |
apply (metis (no_types, lifting) case_prod_Pair case_prod_beta' eq_id_iff) |
3708 |
apply auto |
|
3709 |
done |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3710 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3711 |
with False show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3712 |
by auto |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3713 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3714 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3715 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3716 |
lemma contractible_space_product_topology: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3717 |
"contractible_space(product_topology X I) \<longleftrightarrow> |
78336 | 3718 |
(product_topology X I) = trivial_topology \<or> (\<forall>i \<in> I. contractible_space(X i))" |
3719 |
proof (cases "(product_topology X I) = trivial_topology") |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3720 |
case False |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3721 |
have 1: "contractible_space (X i)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3722 |
if XI: "contractible_space (product_topology X I)" and "i \<in> I" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3723 |
for i |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3724 |
proof (rule contractible_space_retraction_map_image [OF _ XI]) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3725 |
show "retraction_map (product_topology X I) (X i) (\<lambda>x. x i)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3726 |
using False by (simp add: retraction_map_product_projection \<open>i \<in> I\<close>) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3727 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3728 |
have 2: "contractible_space (product_topology X I)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3729 |
if "x \<in> topspace (product_topology X I)" and cs: "\<forall>i\<in>I. contractible_space (X i)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3730 |
for x :: "'a \<Rightarrow> 'b" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3731 |
proof - |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3732 |
obtain f where f: "\<And>i. i\<in>I \<Longrightarrow> homotopic_with (\<lambda>x. True) (X i) (X i) id (\<lambda>x. f i)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3733 |
using cs unfolding contractible_space_def by metis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3734 |
have "homotopic_with (\<lambda>x. True) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3735 |
(product_topology X I) (product_topology X I) id (\<lambda>x. restrict f I)" |
71172 | 3736 |
by (rule homotopic_with_eq [OF homotopic_with_product_topology [OF f]]) (auto) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3737 |
then show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3738 |
by (auto simp: contractible_space_def) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3739 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3740 |
show ?thesis |
78336 | 3741 |
using False 1 2 by (meson equals0I subtopology_eq_discrete_topology_empty) |
3742 |
qed auto |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3743 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3744 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3745 |
lemma contractible_space_subtopology_euclideanreal [simp]: |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3746 |
"contractible_space(subtopology euclideanreal S) \<longleftrightarrow> is_interval S" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3747 |
(is "?lhs = ?rhs") |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3748 |
proof |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3749 |
assume ?lhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3750 |
then have "path_connectedin (subtopology euclideanreal S) S" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3751 |
using contractible_imp_path_connected_space path_connectedin_topspace path_connectedin_absolute |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3752 |
by (simp add: contractible_imp_path_connected) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3753 |
then show ?rhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3754 |
by (simp add: is_interval_path_connected_1) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3755 |
next |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3756 |
assume ?rhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3757 |
then have "convex S" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3758 |
by (simp add: is_interval_convex_1) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3759 |
show ?lhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3760 |
proof (cases "S = {}") |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3761 |
case False |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3762 |
then obtain z where "z \<in> S" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3763 |
by blast |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3764 |
show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3765 |
unfolding contractible_space_def homotopic_with_def |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3766 |
proof (intro exI conjI allI) |
71745 | 3767 |
note \<section> = convexD [OF \<open>convex S\<close>, simplified] |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3768 |
show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set S)) (top_of_set S) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3769 |
(\<lambda>(t,x). (1 - t) * x + t * z)" |
71745 | 3770 |
using \<open>z \<in> S\<close> |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
3771 |
by (auto simp: case_prod_unfold intro!: continuous_intros \<section>) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3772 |
qed auto |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3773 |
qed (simp add: contractible_space_empty) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3774 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3775 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3776 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3777 |
corollary contractible_space_euclideanreal: "contractible_space euclideanreal" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3778 |
proof - |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3779 |
have "contractible_space (subtopology euclideanreal UNIV)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3780 |
using contractible_space_subtopology_euclideanreal by blast |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3781 |
then show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3782 |
by simp |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3783 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3784 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3785 |
|
70136 | 3786 |
abbreviation\<^marker>\<open>tag important\<close> homotopy_eqv :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool" |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80052
diff
changeset
|
3787 |
(infix \<open>homotopy'_eqv\<close> 50) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3788 |
where "S homotopy_eqv T \<equiv> top_of_set S homotopy_equivalent_space top_of_set T" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3789 |
|
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3790 |
lemma homeomorphic_imp_homotopy_eqv: "S homeomorphic T \<Longrightarrow> S homotopy_eqv T" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3791 |
unfolding homeomorphic_def homeomorphism_def homotopy_equivalent_space_def |
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
3792 |
apply (erule ex_forward)+ |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
3793 |
by (metis continuous_map_subtopology_eu homotopic_with_id2 openin_imp_subset openin_subtopology_self topspace_euclidean_subtopology |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
3794 |
image_subset_iff_funcset) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3795 |
|
69620 | 3796 |
lemma homotopy_eqv_inj_linear_image: |
3797 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
3798 |
assumes "linear f" "inj f" |
|
3799 |
shows "(f ` S) homotopy_eqv S" |
|
71745 | 3800 |
by (metis assms homeomorphic_sym homeomorphic_imp_homotopy_eqv linear_homeomorphic_image) |
69620 | 3801 |
|
3802 |
lemma homotopy_eqv_translation: |
|
3803 |
fixes S :: "'a::real_normed_vector set" |
|
3804 |
shows "(+) a ` S homotopy_eqv S" |
|
71745 | 3805 |
using homeomorphic_imp_homotopy_eqv homeomorphic_translation homeomorphic_sym by blast |
69620 | 3806 |
|
3807 |
lemma homotopy_eqv_homotopic_triviality_imp: |
|
3808 |
fixes S :: "'a::real_normed_vector set" |
|
3809 |
and T :: "'b::real_normed_vector set" |
|
3810 |
and U :: "'c::real_normed_vector set" |
|
3811 |
assumes "S homotopy_eqv T" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3812 |
and f: "continuous_on U f" "f \<in> U \<rightarrow> T" |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3813 |
and g: "continuous_on U g" "g \<in> U \<rightarrow> T" |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3814 |
and homUS: "\<And>f g. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S; |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3815 |
continuous_on U g; g \<in> U \<rightarrow> S\<rbrakk> |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3816 |
\<Longrightarrow> homotopic_with_canon (\<lambda>x. True) U S f g" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3817 |
shows "homotopic_with_canon (\<lambda>x. True) U T f g" |
69620 | 3818 |
proof - |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3819 |
obtain h k where h: "continuous_on S h" "h \<in> S \<rightarrow> T" |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3820 |
and k: "continuous_on T k" "k \<in> T \<rightarrow> S" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3821 |
and hom: "homotopic_with_canon (\<lambda>x. True) S S (k \<circ> h) id" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3822 |
"homotopic_with_canon (\<lambda>x. True) T T (h \<circ> k) id" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3823 |
using assms by (force simp: homotopy_equivalent_space_def image_subset_iff_funcset) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3824 |
have "homotopic_with_canon (\<lambda>f. True) U S (k \<circ> f) (k \<circ> g)" |
71745 | 3825 |
proof (rule homUS) |
3826 |
show "continuous_on U (k \<circ> f)" "continuous_on U (k \<circ> g)" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3827 |
using continuous_on_compose continuous_on_subset f g k by (metis funcset_image)+ |
71745 | 3828 |
qed (use f g k in \<open>(force simp: o_def)+\<close> ) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3829 |
then have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> (k \<circ> f)) (h \<circ> (k \<circ> g))" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3830 |
by (simp add: h homotopic_with_compose_continuous_map_left image_subset_iff_funcset) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3831 |
moreover have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> k \<circ> f) (id \<circ> f)" |
71745 | 3832 |
by (rule homotopic_with_compose_continuous_right [where X=T and Y=T]; simp add: hom f) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3833 |
moreover have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> k \<circ> g) (id \<circ> g)" |
71745 | 3834 |
by (rule homotopic_with_compose_continuous_right [where X=T and Y=T]; simp add: hom g) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3835 |
ultimately show "homotopic_with_canon (\<lambda>x. True) U T f g" |
71745 | 3836 |
unfolding o_assoc |
3837 |
by (metis homotopic_with_trans homotopic_with_sym id_comp) |
|
69620 | 3838 |
qed |
3839 |
||
3840 |
lemma homotopy_eqv_homotopic_triviality: |
|
3841 |
fixes S :: "'a::real_normed_vector set" |
|
3842 |
and T :: "'b::real_normed_vector set" |
|
3843 |
and U :: "'c::real_normed_vector set" |
|
3844 |
assumes "S homotopy_eqv T" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3845 |
shows "(\<forall>f g. continuous_on U f \<and> f \<in> U \<rightarrow> S \<and> |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3846 |
continuous_on U g \<and> g \<in> U \<rightarrow> S |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3847 |
\<longrightarrow> homotopic_with_canon (\<lambda>x. True) U S f g) \<longleftrightarrow> |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3848 |
(\<forall>f g. continuous_on U f \<and> f \<in> U \<rightarrow> T \<and> |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3849 |
continuous_on U g \<and> g \<in> U \<rightarrow> T |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3850 |
\<longrightarrow> homotopic_with_canon (\<lambda>x. True) U T f g)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3851 |
(is "?lhs = ?rhs") |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3852 |
proof |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3853 |
assume ?lhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3854 |
then show ?rhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3855 |
by (metis assms homotopy_eqv_homotopic_triviality_imp) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3856 |
next |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3857 |
assume ?rhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3858 |
moreover |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3859 |
have "T homotopy_eqv S" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3860 |
using assms homotopy_equivalent_space_sym by blast |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3861 |
ultimately show ?lhs |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3862 |
by (blast intro: homotopy_eqv_homotopic_triviality_imp) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3863 |
qed |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3864 |
|
69620 | 3865 |
|
3866 |
lemma homotopy_eqv_cohomotopic_triviality_null_imp: |
|
3867 |
fixes S :: "'a::real_normed_vector set" |
|
3868 |
and T :: "'b::real_normed_vector set" |
|
3869 |
and U :: "'c::real_normed_vector set" |
|
3870 |
assumes "S homotopy_eqv T" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3871 |
and f: "continuous_on T f" "f \<in> T \<rightarrow> U" |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3872 |
and homSU: "\<And>f. \<lbrakk>continuous_on S f; f \<in> S \<rightarrow> U\<rbrakk> |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3873 |
\<Longrightarrow> \<exists>c. homotopic_with_canon (\<lambda>x. True) S U f (\<lambda>x. c)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3874 |
obtains c where "homotopic_with_canon (\<lambda>x. True) T U f (\<lambda>x. c)" |
69620 | 3875 |
proof - |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3876 |
obtain h k where h: "continuous_on S h" "h \<in> S \<rightarrow> T" |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3877 |
and k: "continuous_on T k" "k \<in> T \<rightarrow> S" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3878 |
and hom: "homotopic_with_canon (\<lambda>x. True) S S (k \<circ> h) id" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3879 |
"homotopic_with_canon (\<lambda>x. True) T T (h \<circ> k) id" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3880 |
using assms by (force simp: homotopy_equivalent_space_def image_subset_iff_funcset) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3881 |
obtain c where "homotopic_with_canon (\<lambda>x. True) S U (f \<circ> h) (\<lambda>x. c)" |
71745 | 3882 |
proof (rule exE [OF homSU]) |
3883 |
show "continuous_on S (f \<circ> h)" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3884 |
by (metis continuous_on_compose continuous_on_subset f h funcset_image) |
71745 | 3885 |
qed (use f h in force) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3886 |
then have "homotopic_with_canon (\<lambda>x. True) T U ((f \<circ> h) \<circ> k) ((\<lambda>x. c) \<circ> k)" |
71745 | 3887 |
by (rule homotopic_with_compose_continuous_right [where X=S]) (use k in auto) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3888 |
moreover have "homotopic_with_canon (\<lambda>x. True) T U (f \<circ> id) (f \<circ> (h \<circ> k))" |
71745 | 3889 |
by (rule homotopic_with_compose_continuous_left [where Y=T]) |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
3890 |
(use f in \<open>auto simp: hom homotopic_with_symD\<close>) |
69620 | 3891 |
ultimately show ?thesis |
78474 | 3892 |
using that homotopic_with_trans by (fastforce simp: o_def) |
69620 | 3893 |
qed |
3894 |
||
3895 |
lemma homotopy_eqv_cohomotopic_triviality_null: |
|
3896 |
fixes S :: "'a::real_normed_vector set" |
|
3897 |
and T :: "'b::real_normed_vector set" |
|
3898 |
and U :: "'c::real_normed_vector set" |
|
3899 |
assumes "S homotopy_eqv T" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3900 |
shows "(\<forall>f. continuous_on S f \<and> f \<in> S \<rightarrow> U |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3901 |
\<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) S U f (\<lambda>x. c))) \<longleftrightarrow> |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3902 |
(\<forall>f. continuous_on T f \<and> f \<in> T \<rightarrow> U |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3903 |
\<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) T U f (\<lambda>x. c)))" |
71745 | 3904 |
by (rule iffI; metis assms homotopy_eqv_cohomotopic_triviality_null_imp homotopy_equivalent_space_sym) |
3905 |
||
3906 |
text \<open>Similar to the proof above\<close> |
|
69620 | 3907 |
lemma homotopy_eqv_homotopic_triviality_null_imp: |
3908 |
fixes S :: "'a::real_normed_vector set" |
|
3909 |
and T :: "'b::real_normed_vector set" |
|
3910 |
and U :: "'c::real_normed_vector set" |
|
3911 |
assumes "S homotopy_eqv T" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3912 |
and f: "continuous_on U f" "f \<in> U \<rightarrow> T" |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3913 |
and homSU: "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S\<rbrakk> |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3914 |
\<Longrightarrow> \<exists>c. homotopic_with_canon (\<lambda>x. True) U S f (\<lambda>x. c)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3915 |
shows "\<exists>c. homotopic_with_canon (\<lambda>x. True) U T f (\<lambda>x. c)" |
69620 | 3916 |
proof - |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3917 |
obtain h k where h: "continuous_on S h" "h \<in> S \<rightarrow> T" |
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3918 |
and k: "continuous_on T k" "k \<in> T \<rightarrow> S" |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3919 |
and hom: "homotopic_with_canon (\<lambda>x. True) S S (k \<circ> h) id" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3920 |
"homotopic_with_canon (\<lambda>x. True) T T (h \<circ> k) id" |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3921 |
using assms by (force simp: homotopy_equivalent_space_def image_subset_iff_funcset) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3922 |
obtain c::'a where "homotopic_with_canon (\<lambda>x. True) U S (k \<circ> f) (\<lambda>x. c)" |
71745 | 3923 |
proof (rule exE [OF homSU [of "k \<circ> f"]]) |
3924 |
show "continuous_on U (k \<circ> f)" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3925 |
using continuous_on_compose continuous_on_subset f k by (metis funcset_image) |
71745 | 3926 |
qed (use f k in force) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3927 |
then have "homotopic_with_canon (\<lambda>x. True) U T (h \<circ> (k \<circ> f)) (h \<circ> (\<lambda>x. c))" |
71745 | 3928 |
by (rule homotopic_with_compose_continuous_left [where Y=S]) (use h in auto) |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3929 |
moreover have "homotopic_with_canon (\<lambda>x. True) U T (id \<circ> f) ((h \<circ> k) \<circ> f)" |
71745 | 3930 |
by (rule homotopic_with_compose_continuous_right [where X=T]) |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
3931 |
(use f in \<open>auto simp: hom homotopic_with_symD\<close>) |
69620 | 3932 |
ultimately show ?thesis |
78474 | 3933 |
using homotopic_with_trans by (fastforce simp: o_def) |
69620 | 3934 |
qed |
3935 |
||
3936 |
lemma homotopy_eqv_homotopic_triviality_null: |
|
3937 |
fixes S :: "'a::real_normed_vector set" |
|
3938 |
and T :: "'b::real_normed_vector set" |
|
3939 |
and U :: "'c::real_normed_vector set" |
|
3940 |
assumes "S homotopy_eqv T" |
|
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3941 |
shows "(\<forall>f. continuous_on U f \<and> f \<in> U \<rightarrow> S |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3942 |
\<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) U S f (\<lambda>x. c))) \<longleftrightarrow> |
78248
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
paulson <lp15@cam.ac.uk>
parents:
78127
diff
changeset
|
3943 |
(\<forall>f. continuous_on U f \<and> f \<in> U \<rightarrow> T |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3944 |
\<longrightarrow> (\<exists>c. homotopic_with_canon (\<lambda>x. True) U T f (\<lambda>x. c)))" |
71745 | 3945 |
by (rule iffI; metis assms homotopy_eqv_homotopic_triviality_null_imp homotopy_equivalent_space_sym) |
69620 | 3946 |
|
3947 |
lemma homotopy_eqv_contractible_sets: |
|
3948 |
fixes S :: "'a::real_normed_vector set" |
|
3949 |
and T :: "'b::real_normed_vector set" |
|
3950 |
assumes "contractible S" "contractible T" "S = {} \<longleftrightarrow> T = {}" |
|
3951 |
shows "S homotopy_eqv T" |
|
3952 |
proof (cases "S = {}") |
|
3953 |
case True with assms show ?thesis |
|
78336 | 3954 |
using homeomorphic_imp_homotopy_eqv by fastforce |
69620 | 3955 |
next |
3956 |
case False |
|
3957 |
with assms obtain a b where "a \<in> S" "b \<in> T" |
|
3958 |
by auto |
|
3959 |
then show ?thesis |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3960 |
unfolding homotopy_equivalent_space_def |
72372 | 3961 |
apply (rule_tac x="\<lambda>x. b" in exI, rule_tac x="\<lambda>x. a" in exI) |
3962 |
apply (intro assms conjI continuous_on_id' homotopic_into_contractible; force) |
|
69620 | 3963 |
done |
3964 |
qed |
|
3965 |
||
3966 |
lemma homotopy_eqv_empty1 [simp]: |
|
3967 |
fixes S :: "'a::real_normed_vector set" |
|
71769 | 3968 |
shows "S homotopy_eqv ({}::'b::real_normed_vector set) \<longleftrightarrow> S = {}" (is "?lhs = ?rhs") |
3969 |
proof |
|
3970 |
assume ?lhs then show ?rhs |
|
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
3971 |
by (meson continuous_map_subtopology_eu equals0D equals0I funcset_mem |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
3972 |
homotopy_equivalent_space_def) |
78336 | 3973 |
qed (use homeomorphic_imp_homotopy_eqv in force) |
69620 | 3974 |
|
3975 |
lemma homotopy_eqv_empty2 [simp]: |
|
3976 |
fixes S :: "'a::real_normed_vector set" |
|
3977 |
shows "({}::'b::real_normed_vector set) homotopy_eqv S \<longleftrightarrow> S = {}" |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3978 |
using homotopy_equivalent_space_sym homotopy_eqv_empty1 by blast |
69620 | 3979 |
|
3980 |
lemma homotopy_eqv_contractibility: |
|
3981 |
fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" |
|
3982 |
shows "S homotopy_eqv T \<Longrightarrow> (contractible S \<longleftrightarrow> contractible T)" |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
3983 |
by (meson contractible_space_top_of_set homotopy_equivalent_space_contractibility) |
69620 | 3984 |
|
3985 |
lemma homotopy_eqv_sing: |
|
3986 |
fixes S :: "'a::real_normed_vector set" and a :: "'b::real_normed_vector" |
|
3987 |
shows "S homotopy_eqv {a} \<longleftrightarrow> S \<noteq> {} \<and> contractible S" |
|
78336 | 3988 |
by (metis contractible_sing empty_not_insert homotopy_eqv_contractibility homotopy_eqv_contractible_sets homotopy_eqv_empty2) |
69620 | 3989 |
|
3990 |
lemma homeomorphic_contractible_eq: |
|
3991 |
fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" |
|
3992 |
shows "S homeomorphic T \<Longrightarrow> (contractible S \<longleftrightarrow> contractible T)" |
|
3993 |
by (simp add: homeomorphic_imp_homotopy_eqv homotopy_eqv_contractibility) |
|
3994 |
||
3995 |
lemma homeomorphic_contractible: |
|
3996 |
fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" |
|
3997 |
shows "\<lbrakk>contractible S; S homeomorphic T\<rbrakk> \<Longrightarrow> contractible T" |
|
3998 |
by (metis homeomorphic_contractible_eq) |
|
3999 |
||
4000 |
||
70136 | 4001 |
subsection\<^marker>\<open>tag unimportant\<close>\<open>Misc other results\<close> |
69620 | 4002 |
|
4003 |
lemma bounded_connected_Compl_real: |
|
4004 |
fixes S :: "real set" |
|
4005 |
assumes "bounded S" and conn: "connected(- S)" |
|
4006 |
shows "S = {}" |
|
4007 |
proof - |
|
4008 |
obtain a b where "S \<subseteq> box a b" |
|
4009 |
by (meson assms bounded_subset_box_symmetric) |
|
4010 |
then have "a \<notin> S" "b \<notin> S" |
|
4011 |
by auto |
|
4012 |
then have "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> x \<in> - S" |
|
4013 |
by (meson Compl_iff conn connected_iff_interval) |
|
4014 |
then show ?thesis |
|
4015 |
using \<open>S \<subseteq> box a b\<close> by auto |
|
4016 |
qed |
|
4017 |
||
69918
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents:
69768
diff
changeset
|
4018 |
corollary bounded_path_connected_Compl_real: |
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents:
69768
diff
changeset
|
4019 |
fixes S :: "real set" |
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents:
69768
diff
changeset
|
4020 |
assumes "bounded S" "path_connected(- S)" shows "S = {}" |
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents:
69768
diff
changeset
|
4021 |
by (simp add: assms bounded_connected_Compl_real path_connected_imp_connected) |
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents:
69768
diff
changeset
|
4022 |
|
69620 | 4023 |
lemma bounded_connected_Compl_1: |
4024 |
fixes S :: "'a::{euclidean_space} set" |
|
4025 |
assumes "bounded S" and conn: "connected(- S)" and 1: "DIM('a) = 1" |
|
4026 |
shows "S = {}" |
|
4027 |
proof - |
|
4028 |
have "DIM('a) = DIM(real)" |
|
4029 |
by (simp add: "1") |
|
4030 |
then obtain f::"'a \<Rightarrow> real" and g |
|
71769 | 4031 |
where "linear f" "\<And>x. norm(f x) = norm x" and fg: "\<And>x. g(f x) = x" "\<And>y. f(g y) = y" |
69620 | 4032 |
by (rule isomorphisms_UNIV_UNIV) blast |
4033 |
with \<open>bounded S\<close> have "bounded (f ` S)" |
|
4034 |
using bounded_linear_image linear_linear by blast |
|
71769 | 4035 |
have "bij f" by (metis fg bijI') |
69620 | 4036 |
have "connected (f ` (-S))" |
4037 |
using connected_linear_image assms \<open>linear f\<close> by blast |
|
4038 |
moreover have "f ` (-S) = - (f ` S)" |
|
71769 | 4039 |
by (simp add: \<open>bij f\<close> bij_image_Compl_eq) |
69620 | 4040 |
finally have "connected (- (f ` S))" |
4041 |
by simp |
|
4042 |
then have "f ` S = {}" |
|
4043 |
using \<open>bounded (f ` S)\<close> bounded_connected_Compl_real by blast |
|
4044 |
then show ?thesis |
|
4045 |
by blast |
|
4046 |
qed |
|
4047 |
||
77200
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
76874
diff
changeset
|
4048 |
lemma connected_card_eq_iff_nontrivial: |
69620 | 4049 |
fixes S :: "'a::metric_space set" |
77200
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
76874
diff
changeset
|
4050 |
shows "connected S \<Longrightarrow> uncountable S \<longleftrightarrow> \<not>(\<exists>a. S \<subseteq> {a})" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
76874
diff
changeset
|
4051 |
by (metis connected_uncountable finite.emptyI finite.insertI rev_finite_subset singleton_iff subsetI uncountable_infinite) |
69620 | 4052 |
|
4053 |
lemma connected_finite_iff_sing: |
|
4054 |
fixes S :: "'a::metric_space set" |
|
4055 |
assumes "connected S" |
|
77200
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
76874
diff
changeset
|
4056 |
shows "finite S \<longleftrightarrow> S = {} \<or> (\<exists>a. S = {a})" |
8f2e6186408f
Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents:
76874
diff
changeset
|
4057 |
using assms connected_uncountable countable_finite by blast |
69620 | 4058 |
|
70136 | 4059 |
subsection\<^marker>\<open>tag unimportant\<close>\<open> Some simple positive connection theorems\<close> |
69620 | 4060 |
|
4061 |
proposition path_connected_convex_diff_countable: |
|
4062 |
fixes U :: "'a::euclidean_space set" |
|
4063 |
assumes "convex U" "\<not> collinear U" "countable S" |
|
4064 |
shows "path_connected(U - S)" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
4065 |
proof (clarsimp simp: path_connected_def) |
69620 | 4066 |
fix a b |
4067 |
assume "a \<in> U" "a \<notin> S" "b \<in> U" "b \<notin> S" |
|
4068 |
let ?m = "midpoint a b" |
|
4069 |
show "\<exists>g. path g \<and> path_image g \<subseteq> U - S \<and> pathstart g = a \<and> pathfinish g = b" |
|
4070 |
proof (cases "a = b") |
|
4071 |
case True |
|
4072 |
then show ?thesis |
|
4073 |
by (metis DiffI \<open>a \<in> U\<close> \<open>a \<notin> S\<close> path_component_def path_component_refl) |
|
4074 |
next |
|
4075 |
case False |
|
4076 |
then have "a \<noteq> ?m" "b \<noteq> ?m" |
|
4077 |
using midpoint_eq_endpoint by fastforce+ |
|
4078 |
have "?m \<in> U" |
|
4079 |
using \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>convex U\<close> convex_contains_segment by force |
|
4080 |
obtain c where "c \<in> U" and nc_abc: "\<not> collinear {a,b,c}" |
|
4081 |
by (metis False \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>\<not> collinear U\<close> collinear_triples insert_absorb) |
|
4082 |
have ncoll_mca: "\<not> collinear {?m,c,a}" |
|
4083 |
by (metis (full_types) \<open>a \<noteq> ?m\<close> collinear_3_trans collinear_midpoint insert_commute nc_abc) |
|
4084 |
have ncoll_mcb: "\<not> collinear {?m,c,b}" |
|
4085 |
by (metis (full_types) \<open>b \<noteq> ?m\<close> collinear_3_trans collinear_midpoint insert_commute nc_abc) |
|
4086 |
have "c \<noteq> ?m" |
|
4087 |
by (metis collinear_midpoint insert_commute nc_abc) |
|
4088 |
then have "closed_segment ?m c \<subseteq> U" |
|
4089 |
by (simp add: \<open>c \<in> U\<close> \<open>?m \<in> U\<close> \<open>convex U\<close> closed_segment_subset) |
|
4090 |
then obtain z where z: "z \<in> closed_segment ?m c" |
|
4091 |
and disjS: "(closed_segment a z \<union> closed_segment z b) \<inter> S = {}" |
|
4092 |
proof - |
|
4093 |
have False if "closed_segment ?m c \<subseteq> {z. (closed_segment a z \<union> closed_segment z b) \<inter> S \<noteq> {}}" |
|
4094 |
proof - |
|
4095 |
have closb: "closed_segment ?m c \<subseteq> |
|
4096 |
{z \<in> closed_segment ?m c. closed_segment a z \<inter> S \<noteq> {}} \<union> {z \<in> closed_segment ?m c. closed_segment z b \<inter> S \<noteq> {}}" |
|
4097 |
using that by blast |
|
4098 |
have *: "countable {z \<in> closed_segment ?m c. closed_segment z u \<inter> S \<noteq> {}}" |
|
4099 |
if "u \<in> U" "u \<notin> S" and ncoll: "\<not> collinear {?m, c, u}" for u |
|
4100 |
proof - |
|
4101 |
have **: False if x1: "x1 \<in> closed_segment ?m c" and x2: "x2 \<in> closed_segment ?m c" |
|
4102 |
and "x1 \<noteq> x2" "x1 \<noteq> u" |
|
4103 |
and w: "w \<in> closed_segment x1 u" "w \<in> closed_segment x2 u" |
|
4104 |
and "w \<in> S" for x1 x2 w |
|
4105 |
proof - |
|
4106 |
have "x1 \<in> affine hull {?m,c}" "x2 \<in> affine hull {?m,c}" |
|
4107 |
using segment_as_ball x1 x2 by auto |
|
4108 |
then have coll_x1: "collinear {x1, ?m, c}" and coll_x2: "collinear {?m, c, x2}" |
|
4109 |
by (simp_all add: affine_hull_3_imp_collinear) (metis affine_hull_3_imp_collinear insert_commute) |
|
4110 |
have "\<not> collinear {x1, u, x2}" |
|
4111 |
proof |
|
4112 |
assume "collinear {x1, u, x2}" |
|
4113 |
then have "collinear {?m, c, u}" |
|
4114 |
by (metis (full_types) \<open>c \<noteq> ?m\<close> coll_x1 coll_x2 collinear_3_trans insert_commute ncoll \<open>x1 \<noteq> x2\<close>) |
|
4115 |
with ncoll show False .. |
|
4116 |
qed |
|
4117 |
then have "closed_segment x1 u \<inter> closed_segment u x2 = {u}" |
|
4118 |
by (blast intro!: Int_closed_segment) |
|
4119 |
then have "w = u" |
|
4120 |
using closed_segment_commute w by auto |
|
4121 |
show ?thesis |
|
4122 |
using \<open>u \<notin> S\<close> \<open>w = u\<close> that(7) by auto |
|
4123 |
qed |
|
4124 |
then have disj: "disjoint ((\<Union>z\<in>closed_segment ?m c. {closed_segment z u \<inter> S}))" |
|
4125 |
by (fastforce simp: pairwise_def disjnt_def) |
|
4126 |
have cou: "countable ((\<Union>z \<in> closed_segment ?m c. {closed_segment z u \<inter> S}) - {{}})" |
|
4127 |
apply (rule pairwise_disjnt_countable_Union [OF _ pairwise_subset [OF disj]]) |
|
4128 |
apply (rule countable_subset [OF _ \<open>countable S\<close>], auto) |
|
4129 |
done |
|
4130 |
define f where "f \<equiv> \<lambda>X. (THE z. z \<in> closed_segment ?m c \<and> X = closed_segment z u \<inter> S)" |
|
4131 |
show ?thesis |
|
4132 |
proof (rule countable_subset [OF _ countable_image [OF cou, where f=f]], clarify) |
|
4133 |
fix x |
|
4134 |
assume x: "x \<in> closed_segment ?m c" "closed_segment x u \<inter> S \<noteq> {}" |
|
4135 |
show "x \<in> f ` ((\<Union>z\<in>closed_segment ?m c. {closed_segment z u \<inter> S}) - {{}})" |
|
4136 |
proof (rule_tac x="closed_segment x u \<inter> S" in image_eqI) |
|
4137 |
show "x = f (closed_segment x u \<inter> S)" |
|
71769 | 4138 |
unfolding f_def |
4139 |
by (rule the_equality [symmetric]) (use x in \<open>auto dest: **\<close>) |
|
69620 | 4140 |
qed (use x in auto) |
4141 |
qed |
|
4142 |
qed |
|
4143 |
have "uncountable (closed_segment ?m c)" |
|
4144 |
by (metis \<open>c \<noteq> ?m\<close> uncountable_closed_segment) |
|
4145 |
then show False |
|
4146 |
using closb * [OF \<open>a \<in> U\<close> \<open>a \<notin> S\<close> ncoll_mca] * [OF \<open>b \<in> U\<close> \<open>b \<notin> S\<close> ncoll_mcb] |
|
72372 | 4147 |
by (simp add: closed_segment_commute countable_subset) |
69620 | 4148 |
qed |
4149 |
then show ?thesis |
|
4150 |
by (force intro: that) |
|
4151 |
qed |
|
4152 |
show ?thesis |
|
4153 |
proof (intro exI conjI) |
|
4154 |
have "path_image (linepath a z +++ linepath z b) \<subseteq> U" |
|
4155 |
by (metis \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>closed_segment ?m c \<subseteq> U\<close> z \<open>convex U\<close> closed_segment_subset contra_subsetD path_image_linepath subset_path_image_join) |
|
4156 |
with disjS show "path_image (linepath a z +++ linepath z b) \<subseteq> U - S" |
|
4157 |
by (force simp: path_image_join) |
|
4158 |
qed auto |
|
4159 |
qed |
|
4160 |
qed |
|
4161 |
||
4162 |
||
4163 |
corollary connected_convex_diff_countable: |
|
4164 |
fixes U :: "'a::euclidean_space set" |
|
4165 |
assumes "convex U" "\<not> collinear U" "countable S" |
|
4166 |
shows "connected(U - S)" |
|
4167 |
by (simp add: assms path_connected_convex_diff_countable path_connected_imp_connected) |
|
4168 |
||
4169 |
lemma path_connected_punctured_convex: |
|
4170 |
assumes "convex S" and aff: "aff_dim S \<noteq> 1" |
|
4171 |
shows "path_connected(S - {a})" |
|
4172 |
proof - |
|
4173 |
consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S \<ge> 2" |
|
4174 |
using assms aff_dim_geq [of S] by linarith |
|
4175 |
then show ?thesis |
|
4176 |
proof cases |
|
4177 |
assume "aff_dim S = -1" |
|
4178 |
then show ?thesis |
|
4179 |
by (metis aff_dim_empty empty_Diff path_connected_empty) |
|
4180 |
next |
|
4181 |
assume "aff_dim S = 0" |
|
4182 |
then show ?thesis |
|
4183 |
by (metis aff_dim_eq_0 Diff_cancel Diff_empty Diff_insert0 convex_empty convex_imp_path_connected path_connected_singleton singletonD) |
|
4184 |
next |
|
4185 |
assume ge2: "aff_dim S \<ge> 2" |
|
4186 |
then have "\<not> collinear S" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
4187 |
proof (clarsimp simp: collinear_affine_hull) |
69620 | 4188 |
fix u v |
4189 |
assume "S \<subseteq> affine hull {u, v}" |
|
4190 |
then have "aff_dim S \<le> aff_dim {u, v}" |
|
4191 |
by (metis (no_types) aff_dim_affine_hull aff_dim_subset) |
|
4192 |
with ge2 show False |
|
4193 |
by (metis (no_types) aff_dim_2 antisym aff not_numeral_le_zero one_le_numeral order_trans) |
|
4194 |
qed |
|
71769 | 4195 |
moreover have "countable {a}" |
69620 | 4196 |
by simp |
71769 | 4197 |
ultimately show ?thesis |
4198 |
by (metis path_connected_convex_diff_countable [OF \<open>convex S\<close>]) |
|
69620 | 4199 |
qed |
4200 |
qed |
|
4201 |
||
4202 |
lemma connected_punctured_convex: |
|
4203 |
shows "\<lbrakk>convex S; aff_dim S \<noteq> 1\<rbrakk> \<Longrightarrow> connected(S - {a})" |
|
4204 |
using path_connected_imp_connected path_connected_punctured_convex by blast |
|
4205 |
||
4206 |
lemma path_connected_complement_countable: |
|
4207 |
fixes S :: "'a::euclidean_space set" |
|
4208 |
assumes "2 \<le> DIM('a)" "countable S" |
|
4209 |
shows "path_connected(- S)" |
|
4210 |
proof - |
|
71769 | 4211 |
have "\<not> collinear (UNIV::'a set)" |
69620 | 4212 |
using assms by (auto simp: collinear_aff_dim [of "UNIV :: 'a set"]) |
71769 | 4213 |
then have "path_connected(UNIV - S)" |
4214 |
by (simp add: \<open>countable S\<close> path_connected_convex_diff_countable) |
|
69620 | 4215 |
then show ?thesis |
4216 |
by (simp add: Compl_eq_Diff_UNIV) |
|
4217 |
qed |
|
4218 |
||
4219 |
proposition path_connected_openin_diff_countable: |
|
4220 |
fixes S :: "'a::euclidean_space set" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
4221 |
assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S" |
69620 | 4222 |
and "\<not> collinear S" "countable T" |
4223 |
shows "path_connected(S - T)" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
4224 |
proof (clarsimp simp: path_connected_component) |
69620 | 4225 |
fix x y |
4226 |
assume xy: "x \<in> S" "x \<notin> T" "y \<in> S" "y \<notin> T" |
|
4227 |
show "path_component (S - T) x y" |
|
4228 |
proof (rule connected_equivalence_relation_gen [OF \<open>connected S\<close>, where P = "\<lambda>x. x \<notin> T"]) |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
4229 |
show "\<exists>z. z \<in> U \<and> z \<notin> T" if opeU: "openin (top_of_set S) U" and "x \<in> U" for U x |
69620 | 4230 |
proof - |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
4231 |
have "openin (top_of_set (affine hull S)) U" |
69620 | 4232 |
using opeU ope openin_trans by blast |
4233 |
with \<open>x \<in> U\<close> obtain r where Usub: "U \<subseteq> affine hull S" and "r > 0" |
|
4234 |
and subU: "ball x r \<inter> affine hull S \<subseteq> U" |
|
4235 |
by (auto simp: openin_contains_ball) |
|
4236 |
with \<open>x \<in> U\<close> have x: "x \<in> ball x r \<inter> affine hull S" |
|
4237 |
by auto |
|
4238 |
have "\<not> S \<subseteq> {x}" |
|
4239 |
using \<open>\<not> collinear S\<close> collinear_subset by blast |
|
4240 |
then obtain x' where "x' \<noteq> x" "x' \<in> S" |
|
4241 |
by blast |
|
4242 |
obtain y where y: "y \<noteq> x" "y \<in> ball x r \<inter> affine hull S" |
|
4243 |
proof |
|
4244 |
show "x + (r / 2 / norm(x' - x)) *\<^sub>R (x' - x) \<noteq> x" |
|
4245 |
using \<open>x' \<noteq> x\<close> \<open>r > 0\<close> by auto |
|
4246 |
show "x + (r / 2 / norm (x' - x)) *\<^sub>R (x' - x) \<in> ball x r \<inter> affine hull S" |
|
4247 |
using \<open>x' \<noteq> x\<close> \<open>r > 0\<close> \<open>x' \<in> S\<close> x |
|
4248 |
by (simp add: dist_norm mem_affine_3_minus hull_inc) |
|
4249 |
qed |
|
4250 |
have "convex (ball x r \<inter> affine hull S)" |
|
4251 |
by (simp add: affine_imp_convex convex_Int) |
|
4252 |
with x y subU have "uncountable U" |
|
4253 |
by (meson countable_subset uncountable_convex) |
|
4254 |
then have "\<not> U \<subseteq> T" |
|
4255 |
using \<open>countable T\<close> countable_subset by blast |
|
4256 |
then show ?thesis by blast |
|
4257 |
qed |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
4258 |
show "\<exists>U. openin (top_of_set S) U \<and> x \<in> U \<and> |
69620 | 4259 |
(\<forall>x\<in>U. \<forall>y\<in>U. x \<notin> T \<and> y \<notin> T \<longrightarrow> path_component (S - T) x y)" |
4260 |
if "x \<in> S" for x |
|
4261 |
proof - |
|
4262 |
obtain r where Ssub: "S \<subseteq> affine hull S" and "r > 0" |
|
4263 |
and subS: "ball x r \<inter> affine hull S \<subseteq> S" |
|
4264 |
using ope \<open>x \<in> S\<close> by (auto simp: openin_contains_ball) |
|
4265 |
then have conv: "convex (ball x r \<inter> affine hull S)" |
|
4266 |
by (simp add: affine_imp_convex convex_Int) |
|
4267 |
have "\<not> aff_dim (affine hull S) \<le> 1" |
|
4268 |
using \<open>\<not> collinear S\<close> collinear_aff_dim by auto |
|
72372 | 4269 |
then have "\<not> aff_dim (ball x r \<inter> affine hull S) \<le> 1" |
73932
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
72372
diff
changeset
|
4270 |
by (metis (no_types, opaque_lifting) aff_dim_convex_Int_open IntI open_ball \<open>0 < r\<close> aff_dim_affine_hull affine_affine_hull affine_imp_convex centre_in_ball empty_iff hull_subset inf_commute subsetCE that) |
69620 | 4271 |
then have "\<not> collinear (ball x r \<inter> affine hull S)" |
72372 | 4272 |
by (simp add: collinear_aff_dim) |
69620 | 4273 |
then have *: "path_connected ((ball x r \<inter> affine hull S) - T)" |
4274 |
by (rule path_connected_convex_diff_countable [OF conv _ \<open>countable T\<close>]) |
|
4275 |
have ST: "ball x r \<inter> affine hull S - T \<subseteq> S - T" |
|
4276 |
using subS by auto |
|
4277 |
show ?thesis |
|
4278 |
proof (intro exI conjI) |
|
4279 |
show "x \<in> ball x r \<inter> affine hull S" |
|
4280 |
using \<open>x \<in> S\<close> \<open>r > 0\<close> by (simp add: hull_inc) |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
4281 |
have "openin (top_of_set (affine hull S)) (ball x r \<inter> affine hull S)" |
69620 | 4282 |
by (subst inf.commute) (simp add: openin_Int_open) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
4283 |
then show "openin (top_of_set S) (ball x r \<inter> affine hull S)" |
69620 | 4284 |
by (rule openin_subset_trans [OF _ subS Ssub]) |
4285 |
qed (use * path_component_trans in \<open>auto simp: path_connected_component path_component_of_subset [OF ST]\<close>) |
|
4286 |
qed |
|
4287 |
qed (use xy path_component_trans in auto) |
|
4288 |
qed |
|
4289 |
||
4290 |
corollary connected_openin_diff_countable: |
|
4291 |
fixes S :: "'a::euclidean_space set" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
4292 |
assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S" |
69620 | 4293 |
and "\<not> collinear S" "countable T" |
4294 |
shows "connected(S - T)" |
|
4295 |
by (metis path_connected_imp_connected path_connected_openin_diff_countable [OF assms]) |
|
4296 |
||
4297 |
corollary path_connected_open_diff_countable: |
|
4298 |
fixes S :: "'a::euclidean_space set" |
|
4299 |
assumes "2 \<le> DIM('a)" "open S" "connected S" "countable T" |
|
4300 |
shows "path_connected(S - T)" |
|
4301 |
proof (cases "S = {}") |
|
4302 |
case True |
|
4303 |
then show ?thesis |
|
71172 | 4304 |
by (simp) |
69620 | 4305 |
next |
4306 |
case False |
|
4307 |
show ?thesis |
|
4308 |
proof (rule path_connected_openin_diff_countable) |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
4309 |
show "openin (top_of_set (affine hull S)) S" |
69620 | 4310 |
by (simp add: assms hull_subset open_subset) |
4311 |
show "\<not> collinear S" |
|
4312 |
using assms False by (simp add: collinear_aff_dim aff_dim_open) |
|
4313 |
qed (simp_all add: assms) |
|
4314 |
qed |
|
4315 |
||
4316 |
corollary connected_open_diff_countable: |
|
4317 |
fixes S :: "'a::euclidean_space set" |
|
4318 |
assumes "2 \<le> DIM('a)" "open S" "connected S" "countable T" |
|
4319 |
shows "connected(S - T)" |
|
4320 |
by (simp add: assms path_connected_imp_connected path_connected_open_diff_countable) |
|
4321 |
||
4322 |
||
4323 |
||
70136 | 4324 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Self-homeomorphisms shuffling points about\<close> |
4325 |
||
4326 |
subsubsection\<^marker>\<open>tag unimportant\<close>\<open>The theorem \<open>homeomorphism_moving_points_exists\<close>\<close> |
|
69620 | 4327 |
|
4328 |
lemma homeomorphism_moving_point_1: |
|
4329 |
fixes a :: "'a::euclidean_space" |
|
4330 |
assumes "affine T" "a \<in> T" and u: "u \<in> ball a r \<inter> T" |
|
4331 |
obtains f g where "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g" |
|
4332 |
"f a = u" "\<And>x. x \<in> sphere a r \<Longrightarrow> f x = x" |
|
4333 |
proof - |
|
4334 |
have nou: "norm (u - a) < r" and "u \<in> T" |
|
4335 |
using u by (auto simp: dist_norm norm_minus_commute) |
|
4336 |
then have "0 < r" |
|
4337 |
by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) |
|
4338 |
define f where "f \<equiv> \<lambda>x. (1 - norm(x - a) / r) *\<^sub>R (u - a) + x" |
|
4339 |
have *: "False" if eq: "x + (norm y / r) *\<^sub>R u = y + (norm x / r) *\<^sub>R u" |
|
4340 |
and nou: "norm u < r" and yx: "norm y < norm x" for x y and u::'a |
|
4341 |
proof - |
|
4342 |
have "x = y + (norm x / r - (norm y / r)) *\<^sub>R u" |
|
4343 |
using eq by (simp add: algebra_simps) |
|
4344 |
then have "norm x = norm (y + ((norm x - norm y) / r) *\<^sub>R u)" |
|
4345 |
by (metis diff_divide_distrib) |
|
4346 |
also have "\<dots> \<le> norm y + norm(((norm x - norm y) / r) *\<^sub>R u)" |
|
4347 |
using norm_triangle_ineq by blast |
|
4348 |
also have "\<dots> = norm y + (norm x - norm y) * (norm u / r)" |
|
4349 |
using yx \<open>r > 0\<close> |
|
70817
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents:
70802
diff
changeset
|
4350 |
by (simp add: field_split_simps) |
69620 | 4351 |
also have "\<dots> < norm y + (norm x - norm y) * 1" |
71769 | 4352 |
proof (subst add_less_cancel_left) |
4353 |
show "(norm x - norm y) * (norm u / r) < (norm x - norm y) * 1" |
|
4354 |
proof (rule mult_strict_left_mono) |
|
4355 |
show "norm u / r < 1" |
|
4356 |
using \<open>0 < r\<close> divide_less_eq_1_pos nou by blast |
|
4357 |
qed (simp add: yx) |
|
4358 |
qed |
|
69620 | 4359 |
also have "\<dots> = norm x" |
4360 |
by simp |
|
4361 |
finally show False by simp |
|
4362 |
qed |
|
4363 |
have "inj f" |
|
4364 |
unfolding f_def |
|
4365 |
proof (clarsimp simp: inj_on_def) |
|
4366 |
fix x y |
|
4367 |
assume "(1 - norm (x - a) / r) *\<^sub>R (u - a) + x = |
|
4368 |
(1 - norm (y - a) / r) *\<^sub>R (u - a) + y" |
|
4369 |
then have eq: "(x - a) + (norm (y - a) / r) *\<^sub>R (u - a) = (y - a) + (norm (x - a) / r) *\<^sub>R (u - a)" |
|
4370 |
by (auto simp: algebra_simps) |
|
4371 |
show "x=y" |
|
4372 |
proof (cases "norm (x - a) = norm (y - a)") |
|
4373 |
case True |
|
4374 |
then show ?thesis |
|
4375 |
using eq by auto |
|
4376 |
next |
|
4377 |
case False |
|
4378 |
then consider "norm (x - a) < norm (y - a)" | "norm (x - a) > norm (y - a)" |
|
4379 |
by linarith |
|
4380 |
then have "False" |
|
4381 |
proof cases |
|
4382 |
case 1 show False |
|
4383 |
using * [OF _ nou 1] eq by simp |
|
4384 |
next |
|
4385 |
case 2 with * [OF eq nou] show False |
|
4386 |
by auto |
|
4387 |
qed |
|
4388 |
then show "x=y" .. |
|
4389 |
qed |
|
4390 |
qed |
|
4391 |
then have inj_onf: "inj_on f (cball a r \<inter> T)" |
|
4392 |
using inj_on_Int by fastforce |
|
4393 |
have contf: "continuous_on (cball a r \<inter> T) f" |
|
4394 |
unfolding f_def using \<open>0 < r\<close> by (intro continuous_intros) blast |
|
4395 |
have fim: "f ` (cball a r \<inter> T) = cball a r \<inter> T" |
|
4396 |
proof |
|
4397 |
have *: "norm (y + (1 - norm y / r) *\<^sub>R u) \<le> r" if "norm y \<le> r" "norm u < r" for y u::'a |
|
4398 |
proof - |
|
4399 |
have "norm (y + (1 - norm y / r) *\<^sub>R u) \<le> norm y + norm((1 - norm y / r) *\<^sub>R u)" |
|
4400 |
using norm_triangle_ineq by blast |
|
4401 |
also have "\<dots> = norm y + abs(1 - norm y / r) * norm u" |
|
4402 |
by simp |
|
4403 |
also have "\<dots> \<le> r" |
|
4404 |
proof - |
|
4405 |
have "(r - norm u) * (r - norm y) \<ge> 0" |
|
4406 |
using that by auto |
|
4407 |
then have "r * norm u + r * norm y \<le> r * r + norm u * norm y" |
|
4408 |
by (simp add: algebra_simps) |
|
4409 |
then show ?thesis |
|
4410 |
using that \<open>0 < r\<close> by (simp add: abs_if field_simps) |
|
4411 |
qed |
|
4412 |
finally show ?thesis . |
|
4413 |
qed |
|
4414 |
have "f ` (cball a r) \<subseteq> cball a r" |
|
72372 | 4415 |
using * nou |
4416 |
apply (clarsimp simp: dist_norm norm_minus_commute f_def) |
|
4417 |
by (metis diff_add_eq diff_diff_add diff_diff_eq2 norm_minus_commute) |
|
69620 | 4418 |
moreover have "f ` T \<subseteq> T" |
4419 |
unfolding f_def using \<open>affine T\<close> \<open>a \<in> T\<close> \<open>u \<in> T\<close> |
|
4420 |
by (force simp: add.commute mem_affine_3_minus) |
|
4421 |
ultimately show "f ` (cball a r \<inter> T) \<subseteq> cball a r \<inter> T" |
|
4422 |
by blast |
|
4423 |
next |
|
4424 |
show "cball a r \<inter> T \<subseteq> f ` (cball a r \<inter> T)" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
4425 |
proof (clarsimp simp: dist_norm norm_minus_commute) |
69620 | 4426 |
fix x |
4427 |
assume x: "norm (x - a) \<le> r" and "x \<in> T" |
|
4428 |
have "\<exists>v \<in> {0..1}. ((1 - v) * r - norm ((x - a) - v *\<^sub>R (u - a))) \<bullet> 1 = 0" |
|
4429 |
by (rule ivt_decreasing_component_on_1) (auto simp: x continuous_intros) |
|
70802
160eaf566bcb
formally augmented corresponding rules for field_simps
haftmann
parents:
70196
diff
changeset
|
4430 |
then obtain v where "0 \<le> v" "v \<le> 1" |
160eaf566bcb
formally augmented corresponding rules for field_simps
haftmann
parents:
70196
diff
changeset
|
4431 |
and v: "(1 - v) * r = norm ((x - a) - v *\<^sub>R (u - a))" |
69620 | 4432 |
by auto |
70802
160eaf566bcb
formally augmented corresponding rules for field_simps
haftmann
parents:
70196
diff
changeset
|
4433 |
then have n: "norm (a - (x - v *\<^sub>R (u - a))) = r - r * v" |
160eaf566bcb
formally augmented corresponding rules for field_simps
haftmann
parents:
70196
diff
changeset
|
4434 |
by (simp add: field_simps norm_minus_commute) |
69620 | 4435 |
show "x \<in> f ` (cball a r \<inter> T)" |
4436 |
proof (rule image_eqI) |
|
4437 |
show "x = f (x - v *\<^sub>R (u - a))" |
|
70802
160eaf566bcb
formally augmented corresponding rules for field_simps
haftmann
parents:
70196
diff
changeset
|
4438 |
using \<open>r > 0\<close> v by (simp add: f_def) (simp add: field_simps) |
69620 | 4439 |
have "x - v *\<^sub>R (u - a) \<in> cball a r" |
70802
160eaf566bcb
formally augmented corresponding rules for field_simps
haftmann
parents:
70196
diff
changeset
|
4440 |
using \<open>r > 0\<close>\<open>0 \<le> v\<close> |
160eaf566bcb
formally augmented corresponding rules for field_simps
haftmann
parents:
70196
diff
changeset
|
4441 |
by (simp add: dist_norm n) |
69620 | 4442 |
moreover have "x - v *\<^sub>R (u - a) \<in> T" |
71172 | 4443 |
by (simp add: f_def \<open>u \<in> T\<close> \<open>x \<in> T\<close> assms mem_affine_3_minus2) |
69620 | 4444 |
ultimately show "x - v *\<^sub>R (u - a) \<in> cball a r \<inter> T" |
4445 |
by blast |
|
4446 |
qed |
|
4447 |
qed |
|
4448 |
qed |
|
71769 | 4449 |
have "compact (cball a r \<inter> T)" |
4450 |
by (simp add: affine_closed compact_Int_closed \<open>affine T\<close>) |
|
4451 |
then obtain g where "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g" |
|
4452 |
by (metis homeomorphism_compact [OF _ contf fim inj_onf]) |
|
4453 |
then show thesis |
|
4454 |
apply (rule_tac f=f in that) |
|
4455 |
using \<open>r > 0\<close> by (simp_all add: f_def dist_norm norm_minus_commute) |
|
69620 | 4456 |
qed |
4457 |
||
70136 | 4458 |
corollary\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_point_2: |
69620 | 4459 |
fixes a :: "'a::euclidean_space" |
4460 |
assumes "affine T" "a \<in> T" and u: "u \<in> ball a r \<inter> T" and v: "v \<in> ball a r \<inter> T" |
|
4461 |
obtains f g where "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g" |
|
4462 |
"f u = v" "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> f x = x" |
|
4463 |
proof - |
|
4464 |
have "0 < r" |
|
4465 |
by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) |
|
4466 |
obtain f1 g1 where hom1: "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f1 g1" |
|
4467 |
and "f1 a = u" and f1: "\<And>x. x \<in> sphere a r \<Longrightarrow> f1 x = x" |
|
4468 |
using homeomorphism_moving_point_1 [OF \<open>affine T\<close> \<open>a \<in> T\<close> u] by blast |
|
4469 |
obtain f2 g2 where hom2: "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f2 g2" |
|
4470 |
and "f2 a = v" and f2: "\<And>x. x \<in> sphere a r \<Longrightarrow> f2 x = x" |
|
4471 |
using homeomorphism_moving_point_1 [OF \<open>affine T\<close> \<open>a \<in> T\<close> v] by blast |
|
4472 |
show ?thesis |
|
4473 |
proof |
|
4474 |
show "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) (f2 \<circ> g1) (f1 \<circ> g2)" |
|
4475 |
by (metis homeomorphism_compose homeomorphism_symD hom1 hom2) |
|
4476 |
have "g1 u = a" |
|
4477 |
using \<open>0 < r\<close> \<open>f1 a = u\<close> assms hom1 homeomorphism_apply1 by fastforce |
|
4478 |
then show "(f2 \<circ> g1) u = v" |
|
4479 |
by (simp add: \<open>f2 a = v\<close>) |
|
4480 |
show "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> (f2 \<circ> g1) x = x" |
|
4481 |
using f1 f2 hom1 homeomorphism_apply1 by fastforce |
|
4482 |
qed |
|
4483 |
qed |
|
4484 |
||
4485 |
||
70136 | 4486 |
corollary\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_point_3: |
69620 | 4487 |
fixes a :: "'a::euclidean_space" |
4488 |
assumes "affine T" "a \<in> T" and ST: "ball a r \<inter> T \<subseteq> S" "S \<subseteq> T" |
|
4489 |
and u: "u \<in> ball a r \<inter> T" and v: "v \<in> ball a r \<inter> T" |
|
4490 |
obtains f g where "homeomorphism S S f g" |
|
4491 |
"f u = v" "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> ball a r \<inter> T" |
|
4492 |
proof - |
|
4493 |
obtain f g where hom: "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g" |
|
4494 |
and "f u = v" and fid: "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> f x = x" |
|
4495 |
using homeomorphism_moving_point_2 [OF \<open>affine T\<close> \<open>a \<in> T\<close> u v] by blast |
|
4496 |
have gid: "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> g x = x" |
|
4497 |
using fid hom homeomorphism_apply1 by fastforce |
|
4498 |
define ff where "ff \<equiv> \<lambda>x. if x \<in> ball a r \<inter> T then f x else x" |
|
4499 |
define gg where "gg \<equiv> \<lambda>x. if x \<in> ball a r \<inter> T then g x else x" |
|
4500 |
show ?thesis |
|
4501 |
proof |
|
4502 |
show "homeomorphism S S ff gg" |
|
4503 |
proof (rule homeomorphismI) |
|
4504 |
have "continuous_on ((cball a r \<inter> T) \<union> (T - ball a r)) ff" |
|
71769 | 4505 |
unfolding ff_def |
72372 | 4506 |
using homeomorphism_cont1 [OF hom] |
4507 |
by (intro continuous_on_cases) (auto simp: affine_closed \<open>affine T\<close> fid) |
|
69620 | 4508 |
then show "continuous_on S ff" |
71769 | 4509 |
by (rule continuous_on_subset) (use ST in auto) |
69620 | 4510 |
have "continuous_on ((cball a r \<inter> T) \<union> (T - ball a r)) gg" |
71769 | 4511 |
unfolding gg_def |
72372 | 4512 |
using homeomorphism_cont2 [OF hom] |
4513 |
by (intro continuous_on_cases) (auto simp: affine_closed \<open>affine T\<close> gid) |
|
69620 | 4514 |
then show "continuous_on S gg" |
71769 | 4515 |
by (rule continuous_on_subset) (use ST in auto) |
69620 | 4516 |
show "ff ` S \<subseteq> S" |
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
4517 |
proof (clarsimp simp: ff_def) |
69620 | 4518 |
fix x |
4519 |
assume "x \<in> S" and x: "dist a x < r" and "x \<in> T" |
|
4520 |
then have "f x \<in> cball a r \<inter> T" |
|
4521 |
using homeomorphism_image1 [OF hom] by force |
|
4522 |
then show "f x \<in> S" |
|
4523 |
using ST(1) \<open>x \<in> T\<close> gid hom homeomorphism_def x by fastforce |
|
4524 |
qed |
|
4525 |
show "gg ` S \<subseteq> S" |
|
77929
48aa9928f090
Numerous significant simplifications
paulson <lp15@cam.ac.uk>
parents:
77684
diff
changeset
|
4526 |
proof (clarsimp simp: gg_def) |
69620 | 4527 |
fix x |
4528 |
assume "x \<in> S" and x: "dist a x < r" and "x \<in> T" |
|
4529 |
then have "g x \<in> cball a r \<inter> T" |
|
4530 |
using homeomorphism_image2 [OF hom] by force |
|
4531 |
then have "g x \<in> ball a r" |
|
4532 |
using homeomorphism_apply2 [OF hom] |
|
4533 |
by (metis Diff_Diff_Int Diff_iff \<open>x \<in> T\<close> cball_def fid le_less mem_Collect_eq mem_ball mem_sphere x) |
|
4534 |
then show "g x \<in> S" |
|
4535 |
using ST(1) \<open>g x \<in> cball a r \<inter> T\<close> by force |
|
4536 |
qed |
|
4537 |
show "\<And>x. x \<in> S \<Longrightarrow> gg (ff x) = x" |
|
71769 | 4538 |
unfolding ff_def gg_def |
69620 | 4539 |
using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] |
72372 | 4540 |
by simp (metis Int_iff homeomorphism_apply1 [OF hom] fid image_eqI less_eq_real_def mem_cball mem_sphere) |
69620 | 4541 |
show "\<And>x. x \<in> S \<Longrightarrow> ff (gg x) = x" |
71769 | 4542 |
unfolding ff_def gg_def |
69620 | 4543 |
using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] |
72372 | 4544 |
by simp (metis Int_iff fid image_eqI less_eq_real_def mem_cball mem_sphere) |
69620 | 4545 |
qed |
4546 |
show "ff u = v" |
|
4547 |
using u by (auto simp: ff_def \<open>f u = v\<close>) |
|
4548 |
show "{x. \<not> (ff x = x \<and> gg x = x)} \<subseteq> ball a r \<inter> T" |
|
4549 |
by (auto simp: ff_def gg_def) |
|
4550 |
qed |
|
4551 |
qed |
|
4552 |
||
4553 |
||
70136 | 4554 |
proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_point: |
69620 | 4555 |
fixes a :: "'a::euclidean_space" |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
4556 |
assumes ope: "openin (top_of_set (affine hull S)) S" |
69620 | 4557 |
and "S \<subseteq> T" |
4558 |
and TS: "T \<subseteq> affine hull S" |
|
4559 |
and S: "connected S" "a \<in> S" "b \<in> S" |
|
4560 |
obtains f g where "homeomorphism T T f g" "f a = b" |
|
4561 |
"{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" |
|
4562 |
"bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
4563 |
proof - |
|
4564 |
have 1: "\<exists>h k. homeomorphism T T h k \<and> h (f d) = d \<and> |
|
4565 |
{x. \<not> (h x = x \<and> k x = x)} \<subseteq> S \<and> bounded {x. \<not> (h x = x \<and> k x = x)}" |
|
4566 |
if "d \<in> S" "f d \<in> S" and homfg: "homeomorphism T T f g" |
|
4567 |
and S: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" |
|
4568 |
and bo: "bounded {x. \<not> (f x = x \<and> g x = x)}" for d f g |
|
4569 |
proof (intro exI conjI) |
|
4570 |
show homgf: "homeomorphism T T g f" |
|
4571 |
by (metis homeomorphism_symD homfg) |
|
4572 |
then show "g (f d) = d" |
|
4573 |
by (meson \<open>S \<subseteq> T\<close> homeomorphism_def subsetD \<open>d \<in> S\<close>) |
|
4574 |
show "{x. \<not> (g x = x \<and> f x = x)} \<subseteq> S" |
|
4575 |
using S by blast |
|
4576 |
show "bounded {x. \<not> (g x = x \<and> f x = x)}" |
|
4577 |
using bo by (simp add: conj_commute) |
|
4578 |
qed |
|
4579 |
have 2: "\<exists>f g. homeomorphism T T f g \<and> f x = f2 (f1 x) \<and> |
|
4580 |
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
4581 |
if "x \<in> S" "f1 x \<in> S" "f2 (f1 x) \<in> S" |
|
4582 |
and hom: "homeomorphism T T f1 g1" "homeomorphism T T f2 g2" |
|
4583 |
and sub: "{x. \<not> (f1 x = x \<and> g1 x = x)} \<subseteq> S" "{x. \<not> (f2 x = x \<and> g2 x = x)} \<subseteq> S" |
|
4584 |
and bo: "bounded {x. \<not> (f1 x = x \<and> g1 x = x)}" "bounded {x. \<not> (f2 x = x \<and> g2 x = x)}" |
|
4585 |
for x f1 f2 g1 g2 |
|
4586 |
proof (intro exI conjI) |
|
4587 |
show homgf: "homeomorphism T T (f2 \<circ> f1) (g1 \<circ> g2)" |
|
4588 |
by (metis homeomorphism_compose hom) |
|
4589 |
then show "(f2 \<circ> f1) x = f2 (f1 x)" |
|
4590 |
by force |
|
4591 |
show "{x. \<not> ((f2 \<circ> f1) x = x \<and> (g1 \<circ> g2) x = x)} \<subseteq> S" |
|
4592 |
using sub by force |
|
4593 |
have "bounded ({x. \<not>(f1 x = x \<and> g1 x = x)} \<union> {x. \<not>(f2 x = x \<and> g2 x = x)})" |
|
4594 |
using bo by simp |
|
4595 |
then show "bounded {x. \<not> ((f2 \<circ> f1) x = x \<and> (g1 \<circ> g2) x = x)}" |
|
4596 |
by (rule bounded_subset) auto |
|
4597 |
qed |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
4598 |
have 3: "\<exists>U. openin (top_of_set S) U \<and> |
69620 | 4599 |
d \<in> U \<and> |
4600 |
(\<forall>x\<in>U. |
|
4601 |
\<exists>f g. homeomorphism T T f g \<and> f d = x \<and> |
|
4602 |
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> |
|
4603 |
bounded {x. \<not> (f x = x \<and> g x = x)})" |
|
4604 |
if "d \<in> S" for d |
|
4605 |
proof - |
|
4606 |
obtain r where "r > 0" and r: "ball d r \<inter> affine hull S \<subseteq> S" |
|
4607 |
by (metis \<open>d \<in> S\<close> ope openin_contains_ball) |
|
4608 |
have *: "\<exists>f g. homeomorphism T T f g \<and> f d = e \<and> |
|
4609 |
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> |
|
4610 |
bounded {x. \<not> (f x = x \<and> g x = x)}" if "e \<in> S" "e \<in> ball d r" for e |
|
4611 |
apply (rule homeomorphism_moving_point_3 [of "affine hull S" d r T d e]) |
|
72372 | 4612 |
using r \<open>S \<subseteq> T\<close> TS that |
69620 | 4613 |
apply (auto simp: \<open>d \<in> S\<close> \<open>0 < r\<close> hull_inc) |
4614 |
using bounded_subset by blast |
|
4615 |
show ?thesis |
|
71769 | 4616 |
by (rule_tac x="S \<inter> ball d r" in exI) (fastforce simp: openin_open_Int \<open>0 < r\<close> that intro: *) |
69620 | 4617 |
qed |
4618 |
have "\<exists>f g. homeomorphism T T f g \<and> f a = b \<and> |
|
4619 |
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
71769 | 4620 |
by (rule connected_equivalence_relation [OF S]; blast intro: 1 2 3) |
69620 | 4621 |
then show ?thesis |
4622 |
using that by auto |
|
4623 |
qed |
|
4624 |
||
4625 |
||
4626 |
lemma homeomorphism_moving_points_exists_gen: |
|
4627 |
assumes K: "finite K" "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S" |
|
4628 |
"pairwise (\<lambda>i j. (x i \<noteq> x j) \<and> (y i \<noteq> y j)) K" |
|
4629 |
and "2 \<le> aff_dim S" |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
4630 |
and ope: "openin (top_of_set (affine hull S)) S" |
69620 | 4631 |
and "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S" |
4632 |
shows "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(x i) = y i) \<and> |
|
4633 |
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
4634 |
using assms |
|
4635 |
proof (induction K) |
|
4636 |
case empty |
|
4637 |
then show ?case |
|
4638 |
by (force simp: homeomorphism_ident) |
|
4639 |
next |
|
4640 |
case (insert i K) |
|
4641 |
then have xney: "\<And>j. \<lbrakk>j \<in> K; j \<noteq> i\<rbrakk> \<Longrightarrow> x i \<noteq> x j \<and> y i \<noteq> y j" |
|
4642 |
and pw: "pairwise (\<lambda>i j. x i \<noteq> x j \<and> y i \<noteq> y j) K" |
|
4643 |
and "x i \<in> S" "y i \<in> S" |
|
4644 |
and xyS: "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S" |
|
4645 |
by (simp_all add: pairwise_insert) |
|
4646 |
obtain f g where homfg: "homeomorphism T T f g" and feq: "\<And>i. i \<in> K \<Longrightarrow> f(x i) = y i" |
|
4647 |
and fg_sub: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" |
|
4648 |
and bo_fg: "bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
4649 |
using insert.IH [OF xyS pw] insert.prems by (blast intro: that) |
|
4650 |
then have "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(x i) = y i) \<and> |
|
4651 |
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
4652 |
using insert by blast |
|
4653 |
have aff_eq: "affine hull (S - y ` K) = affine hull S" |
|
71769 | 4654 |
proof (rule affine_hull_Diff [OF ope]) |
4655 |
show "finite (y ` K)" |
|
4656 |
by (simp add: insert.hyps(1)) |
|
4657 |
show "y ` K \<subset> S" |
|
4658 |
using \<open>y i \<in> S\<close> insert.hyps(2) xney xyS by fastforce |
|
4659 |
qed |
|
69620 | 4660 |
have f_in_S: "f x \<in> S" if "x \<in> S" for x |
4661 |
using homfg fg_sub homeomorphism_apply1 \<open>S \<subseteq> T\<close> |
|
4662 |
proof - |
|
4663 |
have "(f (f x) \<noteq> f x \<or> g (f x) \<noteq> f x) \<or> f x \<in> S" |
|
4664 |
by (metis \<open>S \<subseteq> T\<close> homfg subsetD homeomorphism_apply1 that) |
|
4665 |
then show ?thesis |
|
4666 |
using fg_sub by force |
|
4667 |
qed |
|
4668 |
obtain h k where homhk: "homeomorphism T T h k" and heq: "h (f (x i)) = y i" |
|
4669 |
and hk_sub: "{x. \<not> (h x = x \<and> k x = x)} \<subseteq> S - y ` K" |
|
4670 |
and bo_hk: "bounded {x. \<not> (h x = x \<and> k x = x)}" |
|
4671 |
proof (rule homeomorphism_moving_point [of "S - y`K" T "f(x i)" "y i"]) |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
4672 |
show "openin (top_of_set (affine hull (S - y ` K))) (S - y ` K)" |
69620 | 4673 |
by (simp add: aff_eq openin_diff finite_imp_closedin image_subset_iff hull_inc insert xyS) |
4674 |
show "S - y ` K \<subseteq> T" |
|
4675 |
using \<open>S \<subseteq> T\<close> by auto |
|
4676 |
show "T \<subseteq> affine hull (S - y ` K)" |
|
4677 |
using insert by (simp add: aff_eq) |
|
4678 |
show "connected (S - y ` K)" |
|
4679 |
proof (rule connected_openin_diff_countable [OF \<open>connected S\<close> ope]) |
|
4680 |
show "\<not> collinear S" |
|
4681 |
using collinear_aff_dim \<open>2 \<le> aff_dim S\<close> by force |
|
4682 |
show "countable (y ` K)" |
|
4683 |
using countable_finite insert.hyps(1) by blast |
|
4684 |
qed |
|
71769 | 4685 |
have "\<And>k. \<lbrakk>f (x i) = y k; k \<in> K\<rbrakk> \<Longrightarrow> False" |
69620 | 4686 |
by (metis feq homfg \<open>x i \<in> S\<close> homeomorphism_def \<open>S \<subseteq> T\<close> \<open>i \<notin> K\<close> subsetCE xney xyS) |
71769 | 4687 |
then show "f (x i) \<in> S - y ` K" |
4688 |
by (auto simp: f_in_S \<open>x i \<in> S\<close>) |
|
69620 | 4689 |
show "y i \<in> S - y ` K" |
4690 |
using insert.hyps xney by (auto simp: \<open>y i \<in> S\<close>) |
|
4691 |
qed blast |
|
4692 |
show ?case |
|
4693 |
proof (intro exI conjI) |
|
4694 |
show "homeomorphism T T (h \<circ> f) (g \<circ> k)" |
|
4695 |
using homfg homhk homeomorphism_compose by blast |
|
4696 |
show "\<forall>i \<in> insert i K. (h \<circ> f) (x i) = y i" |
|
4697 |
using feq hk_sub by (auto simp: heq) |
|
4698 |
show "{x. \<not> ((h \<circ> f) x = x \<and> (g \<circ> k) x = x)} \<subseteq> S" |
|
4699 |
using fg_sub hk_sub by force |
|
4700 |
have "bounded ({x. \<not>(f x = x \<and> g x = x)} \<union> {x. \<not>(h x = x \<and> k x = x)})" |
|
4701 |
using bo_fg bo_hk bounded_Un by blast |
|
4702 |
then show "bounded {x. \<not> ((h \<circ> f) x = x \<and> (g \<circ> k) x = x)}" |
|
4703 |
by (rule bounded_subset) auto |
|
4704 |
qed |
|
4705 |
qed |
|
4706 |
||
70136 | 4707 |
proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_points_exists: |
69620 | 4708 |
fixes S :: "'a::euclidean_space set" |
4709 |
assumes 2: "2 \<le> DIM('a)" "open S" "connected S" "S \<subseteq> T" "finite K" |
|
4710 |
and KS: "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S" |
|
4711 |
and pw: "pairwise (\<lambda>i j. (x i \<noteq> x j) \<and> (y i \<noteq> y j)) K" |
|
4712 |
and S: "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S" |
|
4713 |
obtains f g where "homeomorphism T T f g" "\<And>i. i \<in> K \<Longrightarrow> f(x i) = y i" |
|
4714 |
"{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" "bounded {x. (\<not> (f x = x \<and> g x = x))}" |
|
4715 |
proof (cases "S = {}") |
|
4716 |
case True |
|
4717 |
then show ?thesis |
|
4718 |
using KS homeomorphism_ident that by fastforce |
|
4719 |
next |
|
4720 |
case False |
|
4721 |
then have affS: "affine hull S = UNIV" |
|
4722 |
by (simp add: affine_hull_open \<open>open S\<close>) |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
4723 |
then have ope: "openin (top_of_set (affine hull S)) S" |
69620 | 4724 |
using \<open>open S\<close> open_openin by auto |
4725 |
have "2 \<le> DIM('a)" by (rule 2) |
|
4726 |
also have "\<dots> = aff_dim (UNIV :: 'a set)" |
|
4727 |
by simp |
|
4728 |
also have "\<dots> \<le> aff_dim S" |
|
4729 |
by (metis aff_dim_UNIV aff_dim_affine_hull aff_dim_le_DIM affS) |
|
4730 |
finally have "2 \<le> aff_dim S" |
|
4731 |
by linarith |
|
4732 |
then show ?thesis |
|
4733 |
using homeomorphism_moving_points_exists_gen [OF \<open>finite K\<close> KS pw _ ope S] that by fastforce |
|
4734 |
qed |
|
4735 |
||
70136 | 4736 |
subsubsection\<^marker>\<open>tag unimportant\<close>\<open>The theorem \<open>homeomorphism_grouping_points_exists\<close>\<close> |
69620 | 4737 |
|
4738 |
lemma homeomorphism_grouping_point_1: |
|
4739 |
fixes a::real and c::real |
|
4740 |
assumes "a < b" "c < d" |
|
4741 |
obtains f g where "homeomorphism (cbox a b) (cbox c d) f g" "f a = c" "f b = d" |
|
4742 |
proof - |
|
4743 |
define f where "f \<equiv> \<lambda>x. ((d - c) / (b - a)) * x + (c - a * ((d - c) / (b - a)))" |
|
4744 |
have "\<exists>g. homeomorphism (cbox a b) (cbox c d) f g" |
|
4745 |
proof (rule homeomorphism_compact) |
|
4746 |
show "continuous_on (cbox a b) f" |
|
71769 | 4747 |
unfolding f_def by (intro continuous_intros) |
69620 | 4748 |
have "f ` {a..b} = {c..d}" |
4749 |
unfolding f_def image_affinity_atLeastAtMost |
|
71172 | 4750 |
using assms sum_sqs_eq by (auto simp: field_split_simps) |
69620 | 4751 |
then show "f ` cbox a b = cbox c d" |
4752 |
by auto |
|
4753 |
show "inj_on f (cbox a b)" |
|
4754 |
unfolding f_def inj_on_def using assms by auto |
|
4755 |
qed auto |
|
4756 |
then obtain g where "homeomorphism (cbox a b) (cbox c d) f g" .. |
|
4757 |
then show ?thesis |
|
4758 |
proof |
|
4759 |
show "f a = c" |
|
4760 |
by (simp add: f_def) |
|
4761 |
show "f b = d" |
|
71172 | 4762 |
using assms sum_sqs_eq [of a b] by (auto simp: f_def field_split_simps) |
69620 | 4763 |
qed |
4764 |
qed |
|
4765 |
||
4766 |
lemma homeomorphism_grouping_point_2: |
|
4767 |
fixes a::real and w::real |
|
4768 |
assumes hom_ab: "homeomorphism (cbox a b) (cbox u v) f1 g1" |
|
4769 |
and hom_bc: "homeomorphism (cbox b c) (cbox v w) f2 g2" |
|
4770 |
and "b \<in> cbox a c" "v \<in> cbox u w" |
|
4771 |
and eq: "f1 a = u" "f1 b = v" "f2 b = v" "f2 c = w" |
|
4772 |
obtains f g where "homeomorphism (cbox a c) (cbox u w) f g" "f a = u" "f c = w" |
|
4773 |
"\<And>x. x \<in> cbox a b \<Longrightarrow> f x = f1 x" "\<And>x. x \<in> cbox b c \<Longrightarrow> f x = f2 x" |
|
4774 |
proof - |
|
4775 |
have le: "a \<le> b" "b \<le> c" "u \<le> v" "v \<le> w" |
|
4776 |
using assms by simp_all |
|
4777 |
then have ac: "cbox a c = cbox a b \<union> cbox b c" and uw: "cbox u w = cbox u v \<union> cbox v w" |
|
4778 |
by auto |
|
4779 |
define f where "f \<equiv> \<lambda>x. if x \<le> b then f1 x else f2 x" |
|
4780 |
have "\<exists>g. homeomorphism (cbox a c) (cbox u w) f g" |
|
4781 |
proof (rule homeomorphism_compact) |
|
4782 |
have cf1: "continuous_on (cbox a b) f1" |
|
4783 |
using hom_ab homeomorphism_cont1 by blast |
|
4784 |
have cf2: "continuous_on (cbox b c) f2" |
|
4785 |
using hom_bc homeomorphism_cont1 by blast |
|
4786 |
show "continuous_on (cbox a c) f" |
|
72372 | 4787 |
unfolding f_def using le eq |
4788 |
by (force intro: continuous_on_cases_le [OF continuous_on_subset [OF cf1] continuous_on_subset [OF cf2]]) |
|
69620 | 4789 |
have "f ` cbox a b = f1 ` cbox a b" "f ` cbox b c = f2 ` cbox b c" |
4790 |
unfolding f_def using eq by force+ |
|
4791 |
then show "f ` cbox a c = cbox u w" |
|
71769 | 4792 |
unfolding ac uw image_Un by (metis hom_ab hom_bc homeomorphism_def) |
69620 | 4793 |
have neq12: "f1 x \<noteq> f2 y" if x: "a \<le> x" "x \<le> b" and y: "b < y" "y \<le> c" for x y |
4794 |
proof - |
|
4795 |
have "f1 x \<in> cbox u v" |
|
4796 |
by (metis hom_ab homeomorphism_def image_eqI mem_box_real(2) x) |
|
4797 |
moreover have "f2 y \<in> cbox v w" |
|
4798 |
by (metis (full_types) hom_bc homeomorphism_def image_subset_iff mem_box_real(2) not_le not_less_iff_gr_or_eq order_refl y) |
|
4799 |
moreover have "f2 y \<noteq> f2 b" |
|
4800 |
by (metis cancel_comm_monoid_add_class.diff_cancel diff_gt_0_iff_gt hom_bc homeomorphism_def le(2) less_imp_le less_numeral_extra(3) mem_box_real(2) order_refl y) |
|
4801 |
ultimately show ?thesis |
|
4802 |
using le eq by simp |
|
4803 |
qed |
|
4804 |
have "inj_on f1 (cbox a b)" |
|
4805 |
by (metis (full_types) hom_ab homeomorphism_def inj_onI) |
|
4806 |
moreover have "inj_on f2 (cbox b c)" |
|
4807 |
by (metis (full_types) hom_bc homeomorphism_def inj_onI) |
|
4808 |
ultimately show "inj_on f (cbox a c)" |
|
4809 |
apply (simp (no_asm) add: inj_on_def) |
|
4810 |
apply (simp add: f_def inj_on_eq_iff) |
|
71769 | 4811 |
using neq12 by force |
69620 | 4812 |
qed auto |
4813 |
then obtain g where "homeomorphism (cbox a c) (cbox u w) f g" .. |
|
4814 |
then show ?thesis |
|
71769 | 4815 |
using eq f_def le that by force |
69620 | 4816 |
qed |
4817 |
||
4818 |
lemma homeomorphism_grouping_point_3: |
|
4819 |
fixes a::real |
|
4820 |
assumes cbox_sub: "cbox c d \<subseteq> box a b" "cbox u v \<subseteq> box a b" |
|
4821 |
and box_ne: "box c d \<noteq> {}" "box u v \<noteq> {}" |
|
4822 |
obtains f g where "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" |
|
4823 |
"\<And>x. x \<in> cbox c d \<Longrightarrow> f x \<in> cbox u v" |
|
4824 |
proof - |
|
4825 |
have less: "a < c" "a < u" "d < b" "v < b" "c < d" "u < v" "cbox c d \<noteq> {}" |
|
4826 |
using assms |
|
4827 |
by (simp_all add: cbox_sub subset_eq) |
|
4828 |
obtain f1 g1 where 1: "homeomorphism (cbox a c) (cbox a u) f1 g1" |
|
4829 |
and f1_eq: "f1 a = a" "f1 c = u" |
|
4830 |
using homeomorphism_grouping_point_1 [OF \<open>a < c\<close> \<open>a < u\<close>] . |
|
4831 |
obtain f2 g2 where 2: "homeomorphism (cbox c d) (cbox u v) f2 g2" |
|
4832 |
and f2_eq: "f2 c = u" "f2 d = v" |
|
4833 |
using homeomorphism_grouping_point_1 [OF \<open>c < d\<close> \<open>u < v\<close>] . |
|
4834 |
obtain f3 g3 where 3: "homeomorphism (cbox d b) (cbox v b) f3 g3" |
|
4835 |
and f3_eq: "f3 d = v" "f3 b = b" |
|
4836 |
using homeomorphism_grouping_point_1 [OF \<open>d < b\<close> \<open>v < b\<close>] . |
|
4837 |
obtain f4 g4 where 4: "homeomorphism (cbox a d) (cbox a v) f4 g4" and "f4 a = a" "f4 d = v" |
|
4838 |
and f4_eq: "\<And>x. x \<in> cbox a c \<Longrightarrow> f4 x = f1 x" "\<And>x. x \<in> cbox c d \<Longrightarrow> f4 x = f2 x" |
|
4839 |
using homeomorphism_grouping_point_2 [OF 1 2] less by (auto simp: f1_eq f2_eq) |
|
4840 |
obtain f g where fg: "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" |
|
4841 |
and f_eq: "\<And>x. x \<in> cbox a d \<Longrightarrow> f x = f4 x" "\<And>x. x \<in> cbox d b \<Longrightarrow> f x = f3 x" |
|
4842 |
using homeomorphism_grouping_point_2 [OF 4 3] less by (auto simp: f4_eq f3_eq f2_eq f1_eq) |
|
4843 |
show ?thesis |
|
71769 | 4844 |
proof (rule that [OF fg]) |
4845 |
show "f x \<in> cbox u v" if "x \<in> cbox c d" for x |
|
4846 |
using that f4_eq f_eq homeomorphism_image1 [OF 2] |
|
4847 |
by (metis atLeastAtMost_iff box_real(2) image_eqI less(1) less_eq_real_def order_trans) |
|
4848 |
qed |
|
69620 | 4849 |
qed |
4850 |
||
4851 |
||
4852 |
lemma homeomorphism_grouping_point_4: |
|
4853 |
fixes T :: "real set" |
|
4854 |
assumes "open U" "open S" "connected S" "U \<noteq> {}" "finite K" "K \<subseteq> S" "U \<subseteq> S" "S \<subseteq> T" |
|
4855 |
obtains f g where "homeomorphism T T f g" |
|
4856 |
"\<And>x. x \<in> K \<Longrightarrow> f x \<in> U" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S" |
|
4857 |
"bounded {x. (\<not> (f x = x \<and> g x = x))}" |
|
4858 |
proof - |
|
4859 |
obtain c d where "box c d \<noteq> {}" "cbox c d \<subseteq> U" |
|
4860 |
proof - |
|
4861 |
obtain u where "u \<in> U" |
|
4862 |
using \<open>U \<noteq> {}\<close> by blast |
|
4863 |
then obtain e where "e > 0" "cball u e \<subseteq> U" |
|
4864 |
using \<open>open U\<close> open_contains_cball by blast |
|
4865 |
then show ?thesis |
|
4866 |
by (rule_tac c=u and d="u+e" in that) (auto simp: dist_norm subset_iff) |
|
4867 |
qed |
|
4868 |
have "compact K" |
|
4869 |
by (simp add: \<open>finite K\<close> finite_imp_compact) |
|
4870 |
obtain a b where "box a b \<noteq> {}" "K \<subseteq> cbox a b" "cbox a b \<subseteq> S" |
|
4871 |
proof (cases "K = {}") |
|
4872 |
case True then show ?thesis |
|
4873 |
using \<open>box c d \<noteq> {}\<close> \<open>cbox c d \<subseteq> U\<close> \<open>U \<subseteq> S\<close> that by blast |
|
4874 |
next |
|
4875 |
case False |
|
4876 |
then obtain a b where "a \<in> K" "b \<in> K" |
|
4877 |
and a: "\<And>x. x \<in> K \<Longrightarrow> a \<le> x" and b: "\<And>x. x \<in> K \<Longrightarrow> x \<le> b" |
|
4878 |
using compact_attains_inf compact_attains_sup by (metis \<open>compact K\<close>)+ |
|
4879 |
obtain e where "e > 0" "cball b e \<subseteq> S" |
|
4880 |
using \<open>open S\<close> open_contains_cball |
|
4881 |
by (metis \<open>b \<in> K\<close> \<open>K \<subseteq> S\<close> subsetD) |
|
4882 |
show ?thesis |
|
4883 |
proof |
|
4884 |
show "box a (b + e) \<noteq> {}" |
|
4885 |
using \<open>0 < e\<close> \<open>b \<in> K\<close> a by force |
|
4886 |
show "K \<subseteq> cbox a (b + e)" |
|
4887 |
using \<open>0 < e\<close> a b by fastforce |
|
4888 |
have "a \<in> S" |
|
4889 |
using \<open>a \<in> K\<close> assms(6) by blast |
|
4890 |
have "b + e \<in> S" |
|
4891 |
using \<open>0 < e\<close> \<open>cball b e \<subseteq> S\<close> by (force simp: dist_norm) |
|
4892 |
show "cbox a (b + e) \<subseteq> S" |
|
4893 |
using \<open>a \<in> S\<close> \<open>b + e \<in> S\<close> \<open>connected S\<close> connected_contains_Icc by auto |
|
4894 |
qed |
|
4895 |
qed |
|
4896 |
obtain w z where "cbox w z \<subseteq> S" and sub_wz: "cbox a b \<union> cbox c d \<subseteq> box w z" |
|
4897 |
proof - |
|
4898 |
have "a \<in> S" "b \<in> S" |
|
4899 |
using \<open>box a b \<noteq> {}\<close> \<open>cbox a b \<subseteq> S\<close> by auto |
|
4900 |
moreover have "c \<in> S" "d \<in> S" |
|
4901 |
using \<open>box c d \<noteq> {}\<close> \<open>cbox c d \<subseteq> U\<close> \<open>U \<subseteq> S\<close> by force+ |
|
4902 |
ultimately have "min a c \<in> S" "max b d \<in> S" |
|
4903 |
by linarith+ |
|
4904 |
then obtain e1 e2 where "e1 > 0" "cball (min a c) e1 \<subseteq> S" "e2 > 0" "cball (max b d) e2 \<subseteq> S" |
|
4905 |
using \<open>open S\<close> open_contains_cball by metis |
|
4906 |
then have *: "min a c - e1 \<in> S" "max b d + e2 \<in> S" |
|
4907 |
by (auto simp: dist_norm) |
|
4908 |
show ?thesis |
|
4909 |
proof |
|
4910 |
show "cbox (min a c - e1) (max b d+ e2) \<subseteq> S" |
|
4911 |
using * \<open>connected S\<close> connected_contains_Icc by auto |
|
4912 |
show "cbox a b \<union> cbox c d \<subseteq> box (min a c - e1) (max b d + e2)" |
|
4913 |
using \<open>0 < e1\<close> \<open>0 < e2\<close> by auto |
|
4914 |
qed |
|
4915 |
qed |
|
4916 |
then |
|
4917 |
obtain f g where hom: "homeomorphism (cbox w z) (cbox w z) f g" |
|
4918 |
and "f w = w" "f z = z" |
|
4919 |
and fin: "\<And>x. x \<in> cbox a b \<Longrightarrow> f x \<in> cbox c d" |
|
4920 |
using homeomorphism_grouping_point_3 [of a b w z c d] |
|
4921 |
using \<open>box a b \<noteq> {}\<close> \<open>box c d \<noteq> {}\<close> by blast |
|
4922 |
have contfg: "continuous_on (cbox w z) f" "continuous_on (cbox w z) g" |
|
4923 |
using hom homeomorphism_def by blast+ |
|
4924 |
define f' where "f' \<equiv> \<lambda>x. if x \<in> cbox w z then f x else x" |
|
4925 |
define g' where "g' \<equiv> \<lambda>x. if x \<in> cbox w z then g x else x" |
|
4926 |
show ?thesis |
|
4927 |
proof |
|
4928 |
have T: "cbox w z \<union> (T - box w z) = T" |
|
4929 |
using \<open>cbox w z \<subseteq> S\<close> \<open>S \<subseteq> T\<close> by auto |
|
4930 |
show "homeomorphism T T f' g'" |
|
4931 |
proof |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
4932 |
have clo: "closedin (top_of_set (cbox w z \<union> (T - box w z))) (T - box w z)" |
69620 | 4933 |
by (metis Diff_Diff_Int Diff_subset T closedin_def open_box openin_open_Int topspace_euclidean_subtopology) |
72372 | 4934 |
have "\<And>x. \<lbrakk>w \<le> x \<and> x \<le> z; w < x \<longrightarrow> \<not> x < z\<rbrakk> \<Longrightarrow> f x = x" |
4935 |
using \<open>f w = w\<close> \<open>f z = z\<close> by auto |
|
4936 |
moreover have "\<And>x. \<lbrakk>w \<le> x \<and> x \<le> z; w < x \<longrightarrow> \<not> x < z\<rbrakk> \<Longrightarrow> g x = x" |
|
4937 |
using \<open>f w = w\<close> \<open>f z = z\<close> hom homeomorphism_apply1 by fastforce |
|
4938 |
ultimately |
|
69620 | 4939 |
have "continuous_on (cbox w z \<union> (T - box w z)) f'" "continuous_on (cbox w z \<union> (T - box w z)) g'" |
4940 |
unfolding f'_def g'_def |
|
72372 | 4941 |
by (intro continuous_on_cases_local contfg continuous_on_id clo; auto simp: closed_subset)+ |
69620 | 4942 |
then show "continuous_on T f'" "continuous_on T g'" |
4943 |
by (simp_all only: T) |
|
4944 |
show "f' ` T \<subseteq> T" |
|
4945 |
unfolding f'_def |
|
4946 |
by clarsimp (metis \<open>cbox w z \<subseteq> S\<close> \<open>S \<subseteq> T\<close> subsetD hom homeomorphism_def imageI mem_box_real(2)) |
|
4947 |
show "g' ` T \<subseteq> T" |
|
4948 |
unfolding g'_def |
|
4949 |
by clarsimp (metis \<open>cbox w z \<subseteq> S\<close> \<open>S \<subseteq> T\<close> subsetD hom homeomorphism_def imageI mem_box_real(2)) |
|
4950 |
show "\<And>x. x \<in> T \<Longrightarrow> g' (f' x) = x" |
|
4951 |
unfolding f'_def g'_def |
|
4952 |
using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] by fastforce |
|
4953 |
show "\<And>y. y \<in> T \<Longrightarrow> f' (g' y) = y" |
|
4954 |
unfolding f'_def g'_def |
|
4955 |
using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] by fastforce |
|
4956 |
qed |
|
4957 |
show "\<And>x. x \<in> K \<Longrightarrow> f' x \<in> U" |
|
4958 |
using fin sub_wz \<open>K \<subseteq> cbox a b\<close> \<open>cbox c d \<subseteq> U\<close> by (force simp: f'_def) |
|
4959 |
show "{x. \<not> (f' x = x \<and> g' x = x)} \<subseteq> S" |
|
4960 |
using \<open>cbox w z \<subseteq> S\<close> by (auto simp: f'_def g'_def) |
|
4961 |
show "bounded {x. \<not> (f' x = x \<and> g' x = x)}" |
|
71769 | 4962 |
proof (rule bounded_subset [of "cbox w z"]) |
4963 |
show "bounded (cbox w z)" |
|
4964 |
using bounded_cbox by blast |
|
4965 |
show "{x. \<not> (f' x = x \<and> g' x = x)} \<subseteq> cbox w z" |
|
4966 |
by (auto simp: f'_def g'_def) |
|
4967 |
qed |
|
69620 | 4968 |
qed |
4969 |
qed |
|
4970 |
||
70136 | 4971 |
proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_grouping_points_exists: |
69620 | 4972 |
fixes S :: "'a::euclidean_space set" |
4973 |
assumes "open U" "open S" "connected S" "U \<noteq> {}" "finite K" "K \<subseteq> S" "U \<subseteq> S" "S \<subseteq> T" |
|
4974 |
obtains f g where "homeomorphism T T f g" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S" |
|
4975 |
"bounded {x. (\<not> (f x = x \<and> g x = x))}" "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U" |
|
4976 |
proof (cases "2 \<le> DIM('a)") |
|
4977 |
case True |
|
4978 |
have TS: "T \<subseteq> affine hull S" |
|
4979 |
using affine_hull_open assms by blast |
|
4980 |
have "infinite U" |
|
4981 |
using \<open>open U\<close> \<open>U \<noteq> {}\<close> finite_imp_not_open by blast |
|
4982 |
then obtain P where "P \<subseteq> U" "finite P" "card K = card P" |
|
4983 |
using infinite_arbitrarily_large by metis |
|
4984 |
then obtain \<gamma> where \<gamma>: "bij_betw \<gamma> K P" |
|
4985 |
using \<open>finite K\<close> finite_same_card_bij by blast |
|
4986 |
obtain f g where "homeomorphism T T f g" "\<And>i. i \<in> K \<Longrightarrow> f (id i) = \<gamma> i" "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" "bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
4987 |
proof (rule homeomorphism_moving_points_exists [OF True \<open>open S\<close> \<open>connected S\<close> \<open>S \<subseteq> T\<close> \<open>finite K\<close>]) |
|
4988 |
show "\<And>i. i \<in> K \<Longrightarrow> id i \<in> S \<and> \<gamma> i \<in> S" |
|
4989 |
using \<open>P \<subseteq> U\<close> \<open>bij_betw \<gamma> K P\<close> \<open>K \<subseteq> S\<close> \<open>U \<subseteq> S\<close> bij_betwE by blast |
|
4990 |
show "pairwise (\<lambda>i j. id i \<noteq> id j \<and> \<gamma> i \<noteq> \<gamma> j) K" |
|
4991 |
using \<gamma> by (auto simp: pairwise_def bij_betw_def inj_on_def) |
|
4992 |
qed (use affine_hull_open assms that in auto) |
|
4993 |
then show ?thesis |
|
78474 | 4994 |
using \<gamma> \<open>P \<subseteq> U\<close> bij_betwE by (fastforce simp: intro!: that) |
69620 | 4995 |
next |
4996 |
case False |
|
4997 |
with DIM_positive have "DIM('a) = 1" |
|
4998 |
by (simp add: dual_order.antisym) |
|
4999 |
then obtain h::"'a \<Rightarrow>real" and j |
|
5000 |
where "linear h" "linear j" |
|
5001 |
and noh: "\<And>x. norm(h x) = norm x" and noj: "\<And>y. norm(j y) = norm y" |
|
5002 |
and hj: "\<And>x. j(h x) = x" "\<And>y. h(j y) = y" |
|
5003 |
and ranh: "surj h" |
|
5004 |
using isomorphisms_UNIV_UNIV |
|
73932
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
72372
diff
changeset
|
5005 |
by (metis (mono_tags, opaque_lifting) DIM_real UNIV_eq_I range_eqI) |
69620 | 5006 |
obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" |
5007 |
and f: "\<And>x. x \<in> h ` K \<Longrightarrow> f x \<in> h ` U" |
|
5008 |
and sub: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> h ` S" |
|
5009 |
and bou: "bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
5010 |
apply (rule homeomorphism_grouping_point_4 [of "h ` U" "h ` S" "h ` K" "h ` T"]) |
|
5011 |
by (simp_all add: assms image_mono \<open>linear h\<close> open_surjective_linear_image connected_linear_image ranh) |
|
5012 |
have jf: "j (f (h x)) = x \<longleftrightarrow> f (h x) = h x" for x |
|
5013 |
by (metis hj) |
|
5014 |
have jg: "j (g (h x)) = x \<longleftrightarrow> g (h x) = h x" for x |
|
5015 |
by (metis hj) |
|
5016 |
have cont_hj: "continuous_on X h" "continuous_on Y j" for X Y |
|
5017 |
by (simp_all add: \<open>linear h\<close> \<open>linear j\<close> linear_linear linear_continuous_on) |
|
5018 |
show ?thesis |
|
5019 |
proof |
|
5020 |
show "homeomorphism T T (j \<circ> f \<circ> h) (j \<circ> g \<circ> h)" |
|
5021 |
proof |
|
5022 |
show "continuous_on T (j \<circ> f \<circ> h)" "continuous_on T (j \<circ> g \<circ> h)" |
|
5023 |
using hom homeomorphism_def |
|
5024 |
by (blast intro: continuous_on_compose cont_hj)+ |
|
5025 |
show "(j \<circ> f \<circ> h) ` T \<subseteq> T" "(j \<circ> g \<circ> h) ` T \<subseteq> T" |
|
73932
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
72372
diff
changeset
|
5026 |
by auto (metis (mono_tags, opaque_lifting) hj(1) hom homeomorphism_def imageE imageI)+ |
69620 | 5027 |
show "\<And>x. x \<in> T \<Longrightarrow> (j \<circ> g \<circ> h) ((j \<circ> f \<circ> h) x) = x" |
5028 |
using hj hom homeomorphism_apply1 by fastforce |
|
5029 |
show "\<And>y. y \<in> T \<Longrightarrow> (j \<circ> f \<circ> h) ((j \<circ> g \<circ> h) y) = y" |
|
5030 |
using hj hom homeomorphism_apply2 by fastforce |
|
5031 |
qed |
|
5032 |
show "{x. \<not> ((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> h) x = x)} \<subseteq> S" |
|
72372 | 5033 |
proof (clarsimp simp: jf jg hj) |
5034 |
show "f (h x) = h x \<longrightarrow> g (h x) \<noteq> h x \<Longrightarrow> x \<in> S" for x |
|
5035 |
using sub [THEN subsetD, of "h x"] hj by simp (metis imageE) |
|
5036 |
qed |
|
69620 | 5037 |
have "bounded (j ` {x. (\<not> (f x = x \<and> g x = x))})" |
5038 |
by (rule bounded_linear_image [OF bou]) (use \<open>linear j\<close> linear_conv_bounded_linear in auto) |
|
5039 |
moreover |
|
5040 |
have *: "{x. \<not>((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> h) x = x)} = j ` {x. (\<not> (f x = x \<and> g x = x))}" |
|
5041 |
using hj by (auto simp: jf jg image_iff, metis+) |
|
5042 |
ultimately show "bounded {x. \<not> ((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> h) x = x)}" |
|
5043 |
by metis |
|
5044 |
show "\<And>x. x \<in> K \<Longrightarrow> (j \<circ> f \<circ> h) x \<in> U" |
|
5045 |
using f hj by fastforce |
|
5046 |
qed |
|
5047 |
qed |
|
5048 |
||
5049 |
||
70136 | 5050 |
proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_grouping_points_exists_gen: |
69620 | 5051 |
fixes S :: "'a::euclidean_space set" |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
5052 |
assumes opeU: "openin (top_of_set S) U" |
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
5053 |
and opeS: "openin (top_of_set (affine hull S)) S" |
69620 | 5054 |
and "U \<noteq> {}" "finite K" "K \<subseteq> S" and S: "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S" |
5055 |
obtains f g where "homeomorphism T T f g" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S" |
|
5056 |
"bounded {x. (\<not> (f x = x \<and> g x = x))}" "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U" |
|
5057 |
proof (cases "2 \<le> aff_dim S") |
|
5058 |
case True |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
5059 |
have opeU': "openin (top_of_set (affine hull S)) U" |
69620 | 5060 |
using opeS opeU openin_trans by blast |
5061 |
obtain u where "u \<in> U" "u \<in> S" |
|
5062 |
using \<open>U \<noteq> {}\<close> opeU openin_imp_subset by fastforce+ |
|
5063 |
have "infinite U" |
|
72372 | 5064 |
proof (rule infinite_openin [OF opeU \<open>u \<in> U\<close>]) |
5065 |
show "u islimpt S" |
|
5066 |
using True \<open>u \<in> S\<close> assms(8) connected_imp_perfect_aff_dim by fastforce |
|
5067 |
qed |
|
69620 | 5068 |
then obtain P where "P \<subseteq> U" "finite P" "card K = card P" |
5069 |
using infinite_arbitrarily_large by metis |
|
5070 |
then obtain \<gamma> where \<gamma>: "bij_betw \<gamma> K P" |
|
5071 |
using \<open>finite K\<close> finite_same_card_bij by blast |
|
5072 |
have "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(id i) = \<gamma> i) \<and> |
|
5073 |
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
5074 |
proof (rule homeomorphism_moving_points_exists_gen [OF \<open>finite K\<close> _ _ True opeS S]) |
|
5075 |
show "\<And>i. i \<in> K \<Longrightarrow> id i \<in> S \<and> \<gamma> i \<in> S" |
|
5076 |
by (metis id_apply opeU openin_contains_cball subsetCE \<open>P \<subseteq> U\<close> \<open>bij_betw \<gamma> K P\<close> \<open>K \<subseteq> S\<close> bij_betwE) |
|
5077 |
show "pairwise (\<lambda>i j. id i \<noteq> id j \<and> \<gamma> i \<noteq> \<gamma> j) K" |
|
5078 |
using \<gamma> by (auto simp: pairwise_def bij_betw_def inj_on_def) |
|
5079 |
qed |
|
5080 |
then show ?thesis |
|
78474 | 5081 |
using \<gamma> \<open>P \<subseteq> U\<close> bij_betwE by (fastforce simp: intro!: that) |
69620 | 5082 |
next |
5083 |
case False |
|
5084 |
with aff_dim_geq [of S] consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S = 1" by linarith |
|
5085 |
then show ?thesis |
|
5086 |
proof cases |
|
5087 |
assume "aff_dim S = -1" |
|
5088 |
then have "S = {}" |
|
5089 |
using aff_dim_empty by blast |
|
5090 |
then have "False" |
|
5091 |
using \<open>U \<noteq> {}\<close> \<open>K \<subseteq> S\<close> openin_imp_subset [OF opeU] by blast |
|
5092 |
then show ?thesis .. |
|
5093 |
next |
|
5094 |
assume "aff_dim S = 0" |
|
5095 |
then obtain a where "S = {a}" |
|
5096 |
using aff_dim_eq_0 by blast |
|
5097 |
then have "K \<subseteq> U" |
|
5098 |
using \<open>U \<noteq> {}\<close> \<open>K \<subseteq> S\<close> openin_imp_subset [OF opeU] by blast |
|
5099 |
show ?thesis |
|
72372 | 5100 |
using \<open>K \<subseteq> U\<close> by (intro that [of id id]) (auto intro: homeomorphismI) |
69620 | 5101 |
next |
5102 |
assume "aff_dim S = 1" |
|
5103 |
then have "affine hull S homeomorphic (UNIV :: real set)" |
|
5104 |
by (auto simp: homeomorphic_affine_sets) |
|
5105 |
then obtain h::"'a\<Rightarrow>real" and j where homhj: "homeomorphism (affine hull S) UNIV h j" |
|
5106 |
using homeomorphic_def by blast |
|
5107 |
then have h: "\<And>x. x \<in> affine hull S \<Longrightarrow> j(h(x)) = x" and j: "\<And>y. j y \<in> affine hull S \<and> h(j y) = y" |
|
5108 |
by (auto simp: homeomorphism_def) |
|
5109 |
have connh: "connected (h ` S)" |
|
5110 |
by (meson Topological_Spaces.connected_continuous_image \<open>connected S\<close> homeomorphism_cont1 homeomorphism_of_subsets homhj hull_subset top_greatest) |
|
5111 |
have hUS: "h ` U \<subseteq> h ` S" |
|
5112 |
by (meson homeomorphism_imp_open_map homeomorphism_of_subsets homhj hull_subset opeS opeU open_UNIV openin_open_eq) |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69918
diff
changeset
|
5113 |
have opn: "openin (top_of_set (affine hull S)) U \<Longrightarrow> open (h ` U)" for U |
69620 | 5114 |
using homeomorphism_imp_open_map [OF homhj] by simp |
5115 |
have "open (h ` U)" "open (h ` S)" |
|
5116 |
by (auto intro: opeS opeU openin_trans opn) |
|
5117 |
then obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" |
|
5118 |
and f: "\<And>x. x \<in> h ` K \<Longrightarrow> f x \<in> h ` U" |
|
5119 |
and sub: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> h ` S" |
|
5120 |
and bou: "bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
5121 |
apply (rule homeomorphism_grouping_points_exists [of "h ` U" "h ` S" "h ` K" "h ` T"]) |
|
5122 |
using assms by (auto simp: connh hUS) |
|
5123 |
have jf: "\<And>x. x \<in> affine hull S \<Longrightarrow> j (f (h x)) = x \<longleftrightarrow> f (h x) = h x" |
|
5124 |
by (metis h j) |
|
5125 |
have jg: "\<And>x. x \<in> affine hull S \<Longrightarrow> j (g (h x)) = x \<longleftrightarrow> g (h x) = h x" |
|
5126 |
by (metis h j) |
|
5127 |
have cont_hj: "continuous_on T h" "continuous_on Y j" for Y |
|
72372 | 5128 |
proof (rule continuous_on_subset [OF _ \<open>T \<subseteq> affine hull S\<close>]) |
5129 |
show "continuous_on (affine hull S) h" |
|
5130 |
using homeomorphism_def homhj by blast |
|
5131 |
qed (meson continuous_on_subset homeomorphism_def homhj top_greatest) |
|
69620 | 5132 |
define f' where "f' \<equiv> \<lambda>x. if x \<in> affine hull S then (j \<circ> f \<circ> h) x else x" |
5133 |
define g' where "g' \<equiv> \<lambda>x. if x \<in> affine hull S then (j \<circ> g \<circ> h) x else x" |
|
5134 |
show ?thesis |
|
5135 |
proof |
|
5136 |
show "homeomorphism T T f' g'" |
|
5137 |
proof |
|
5138 |
have "continuous_on T (j \<circ> f \<circ> h)" |
|
72372 | 5139 |
using hom homeomorphism_def by (intro continuous_on_compose cont_hj) blast |
69620 | 5140 |
then show "continuous_on T f'" |
5141 |
apply (rule continuous_on_eq) |
|
5142 |
using \<open>T \<subseteq> affine hull S\<close> f'_def by auto |
|
5143 |
have "continuous_on T (j \<circ> g \<circ> h)" |
|
72372 | 5144 |
using hom homeomorphism_def by (intro continuous_on_compose cont_hj) blast |
69620 | 5145 |
then show "continuous_on T g'" |
5146 |
apply (rule continuous_on_eq) |
|
5147 |
using \<open>T \<subseteq> affine hull S\<close> g'_def by auto |
|
5148 |
show "f' ` T \<subseteq> T" |
|
5149 |
proof (clarsimp simp: f'_def) |
|
5150 |
fix x assume "x \<in> T" |
|
5151 |
then have "f (h x) \<in> h ` T" |
|
5152 |
by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) |
|
5153 |
then show "j (f (h x)) \<in> T" |
|
5154 |
using \<open>T \<subseteq> affine hull S\<close> h by auto |
|
5155 |
qed |
|
5156 |
show "g' ` T \<subseteq> T" |
|
5157 |
proof (clarsimp simp: g'_def) |
|
5158 |
fix x assume "x \<in> T" |
|
5159 |
then have "g (h x) \<in> h ` T" |
|
5160 |
by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) |
|
5161 |
then show "j (g (h x)) \<in> T" |
|
5162 |
using \<open>T \<subseteq> affine hull S\<close> h by auto |
|
5163 |
qed |
|
5164 |
show "\<And>x. x \<in> T \<Longrightarrow> g' (f' x) = x" |
|
78474 | 5165 |
using h j hom homeomorphism_apply1 by (fastforce simp: f'_def g'_def) |
69620 | 5166 |
show "\<And>y. y \<in> T \<Longrightarrow> f' (g' y) = y" |
78474 | 5167 |
using h j hom homeomorphism_apply2 by (fastforce simp: f'_def g'_def) |
69620 | 5168 |
qed |
5169 |
next |
|
72372 | 5170 |
have \<section>: "\<And>x y. \<lbrakk>x \<in> affine hull S; h x = h y; y \<in> S\<rbrakk> \<Longrightarrow> x \<in> S" |
5171 |
by (metis h hull_inc) |
|
69620 | 5172 |
show "{x. \<not> (f' x = x \<and> g' x = x)} \<subseteq> S" |
72372 | 5173 |
using sub by (simp add: f'_def g'_def jf jg) (force elim: \<section>) |
69620 | 5174 |
next |
5175 |
have "compact (j ` closure {x. \<not> (f x = x \<and> g x = x)})" |
|
5176 |
using bou by (auto simp: compact_continuous_image cont_hj) |
|
5177 |
then have "bounded (j ` {x. \<not> (f x = x \<and> g x = x)})" |
|
5178 |
by (rule bounded_closure_image [OF compact_imp_bounded]) |
|
5179 |
moreover |
|
5180 |
have *: "{x \<in> affine hull S. j (f (h x)) \<noteq> x \<or> j (g (h x)) \<noteq> x} = j ` {x. (\<not> (f x = x \<and> g x = x))}" |
|
5181 |
using h j by (auto simp: image_iff; metis) |
|
5182 |
ultimately have "bounded {x \<in> affine hull S. j (f (h x)) \<noteq> x \<or> j (g (h x)) \<noteq> x}" |
|
5183 |
by metis |
|
5184 |
then show "bounded {x. \<not> (f' x = x \<and> g' x = x)}" |
|
5185 |
by (simp add: f'_def g'_def Collect_mono bounded_subset) |
|
5186 |
next |
|
5187 |
show "f' x \<in> U" if "x \<in> K" for x |
|
5188 |
proof - |
|
5189 |
have "U \<subseteq> S" |
|
5190 |
using opeU openin_imp_subset by blast |
|
5191 |
then have "j (f (h x)) \<in> U" |
|
5192 |
using f h hull_subset that by fastforce |
|
5193 |
then show "f' x \<in> U" |
|
5194 |
using \<open>K \<subseteq> S\<close> S f'_def that by auto |
|
5195 |
qed |
|
5196 |
qed |
|
5197 |
qed |
|
5198 |
qed |
|
5199 |
||
5200 |
||
5201 |
subsection\<open>Nullhomotopic mappings\<close> |
|
5202 |
||
5203 |
text\<open> A mapping out of a sphere is nullhomotopic iff it extends to the ball. |
|
5204 |
This even works out in the degenerate cases when the radius is \<open>\<le>\<close> 0, and |
|
5205 |
we also don't need to explicitly assume continuity since it's already implicit |
|
5206 |
in both sides of the equivalence.\<close> |
|
5207 |
||
5208 |
lemma nullhomotopic_from_lemma: |
|
5209 |
assumes contg: "continuous_on (cball a r - {a}) g" |
|
5210 |
and fa: "\<And>e. 0 < e |
|
5211 |
\<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>x. x \<noteq> a \<and> norm(x - a) < d \<longrightarrow> norm(g x - f a) < e)" |
|
5212 |
and r: "\<And>x. x \<in> cball a r \<and> x \<noteq> a \<Longrightarrow> f x = g x" |
|
5213 |
shows "continuous_on (cball a r) f" |
|
5214 |
proof (clarsimp simp: continuous_on_eq_continuous_within Ball_def) |
|
5215 |
fix x |
|
5216 |
assume x: "dist a x \<le> r" |
|
5217 |
show "continuous (at x within cball a r) f" |
|
5218 |
proof (cases "x=a") |
|
5219 |
case True |
|
5220 |
then show ?thesis |
|
5221 |
by (metis continuous_within_eps_delta fa dist_norm dist_self r) |
|
5222 |
next |
|
5223 |
case False |
|
5224 |
show ?thesis |
|
5225 |
proof (rule continuous_transform_within [where f=g and d = "norm(x-a)"]) |
|
5226 |
have "\<exists>d>0. \<forall>x'\<in>cball a r. |
|
5227 |
dist x' x < d \<longrightarrow> dist (g x') (g x) < e" if "e>0" for e |
|
5228 |
proof - |
|
5229 |
obtain d where "d > 0" |
|
5230 |
and d: "\<And>x'. \<lbrakk>dist x' a \<le> r; x' \<noteq> a; dist x' x < d\<rbrakk> \<Longrightarrow> |
|
5231 |
dist (g x') (g x) < e" |
|
5232 |
using contg False x \<open>e>0\<close> |
|
78474 | 5233 |
unfolding continuous_on_iff by (fastforce simp: dist_commute intro: that) |
69620 | 5234 |
show ?thesis |
5235 |
using \<open>d > 0\<close> \<open>x \<noteq> a\<close> |
|
5236 |
by (rule_tac x="min d (norm(x - a))" in exI) |
|
5237 |
(auto simp: dist_commute dist_norm [symmetric] intro!: d) |
|
5238 |
qed |
|
5239 |
then show "continuous (at x within cball a r) g" |
|
5240 |
using contg False by (auto simp: continuous_within_eps_delta) |
|
5241 |
show "0 < norm (x - a)" |
|
5242 |
using False by force |
|
5243 |
show "x \<in> cball a r" |
|
5244 |
by (simp add: x) |
|
5245 |
show "\<And>x'. \<lbrakk>x' \<in> cball a r; dist x' x < norm (x - a)\<rbrakk> |
|
5246 |
\<Longrightarrow> g x' = f x'" |
|
5247 |
by (metis dist_commute dist_norm less_le r) |
|
5248 |
qed |
|
5249 |
qed |
|
5250 |
qed |
|
5251 |
||
5252 |
proposition nullhomotopic_from_sphere_extension: |
|
5253 |
fixes f :: "'M::euclidean_space \<Rightarrow> 'a::real_normed_vector" |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
5254 |
shows "(\<exists>c. homotopic_with_canon (\<lambda>x. True) (sphere a r) S f (\<lambda>x. c)) \<longleftrightarrow> |
69620 | 5255 |
(\<exists>g. continuous_on (cball a r) g \<and> g ` (cball a r) \<subseteq> S \<and> |
5256 |
(\<forall>x \<in> sphere a r. g x = f x))" |
|
5257 |
(is "?lhs = ?rhs") |
|
5258 |
proof (cases r "0::real" rule: linorder_cases) |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
5259 |
case less |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
5260 |
then show ?thesis |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
5261 |
by (simp add: homotopic_on_emptyI) |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
5262 |
next |
69620 | 5263 |
case equal |
71768 | 5264 |
show ?thesis |
5265 |
proof |
|
5266 |
assume L: ?lhs |
|
5267 |
with equal have [simp]: "f a \<in> S" |
|
5268 |
using homotopic_with_imp_subset1 by fastforce |
|
5269 |
obtain h:: "real \<times> 'M \<Rightarrow> 'a" |
|
5270 |
where h: "continuous_on ({0..1} \<times> {a}) h" "h ` ({0..1} \<times> {a}) \<subseteq> S" "h (0, a) = f a" |
|
5271 |
using L equal by (auto simp: homotopic_with) |
|
5272 |
then have "continuous_on (cball a r) (\<lambda>x. h (0, a))" "(\<lambda>x. h (0, a)) ` cball a r \<subseteq> S" |
|
5273 |
by (auto simp: equal) |
|
5274 |
then show ?rhs |
|
5275 |
using h(3) local.equal by force |
|
5276 |
next |
|
5277 |
assume ?rhs |
|
5278 |
then show ?lhs |
|
78474 | 5279 |
using equal continuous_on_const by (force simp: homotopic_with) |
71768 | 5280 |
qed |
69620 | 5281 |
next |
5282 |
case greater |
|
5283 |
let ?P = "continuous_on {x. norm(x - a) = r} f \<and> f ` {x. norm(x - a) = r} \<subseteq> S" |
|
5284 |
have ?P if ?lhs using that |
|
5285 |
proof |
|
5286 |
fix c |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
5287 |
assume c: "homotopic_with_canon (\<lambda>x. True) (sphere a r) S f (\<lambda>x. c)" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
5288 |
then have contf: "continuous_on (sphere a r) f" |
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
5289 |
by (metis homotopic_with_imp_continuous) |
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
5290 |
moreover have fim: "f \<in> sphere a r \<rightarrow> S" |
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
5291 |
using homotopic_with_imp_subset1 that by blast |
69620 | 5292 |
show ?P |
5293 |
using contf fim by (auto simp: sphere_def dist_norm norm_minus_commute) |
|
5294 |
qed |
|
5295 |
moreover have ?P if ?rhs using that |
|
5296 |
proof |
|
5297 |
fix g |
|
5298 |
assume g: "continuous_on (cball a r) g \<and> g ` cball a r \<subseteq> S \<and> (\<forall>xa\<in>sphere a r. g xa = f xa)" |
|
71768 | 5299 |
then have "f ` {x. norm (x - a) = r} \<subseteq> S" |
5300 |
using sphere_cball [of a r] unfolding image_subset_iff sphere_def |
|
5301 |
by (metis dist_commute dist_norm mem_Collect_eq subset_eq) |
|
5302 |
with g show ?P |
|
5303 |
by (auto simp: dist_norm norm_minus_commute elim!: continuous_on_eq [OF continuous_on_subset]) |
|
69620 | 5304 |
qed |
5305 |
moreover have ?thesis if ?P |
|
5306 |
proof |
|
5307 |
assume ?lhs |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
5308 |
then obtain c where "homotopic_with_canon (\<lambda>x. True) (sphere a r) S (\<lambda>x. c) f" |
69620 | 5309 |
using homotopic_with_sym by blast |
5310 |
then obtain h where conth: "continuous_on ({0..1::real} \<times> sphere a r) h" |
|
5311 |
and him: "h ` ({0..1} \<times> sphere a r) \<subseteq> S" |
|
5312 |
and h: "\<And>x. h(0, x) = c" "\<And>x. h(1, x) = f x" |
|
5313 |
by (auto simp: homotopic_with_def) |
|
5314 |
obtain b1::'M where "b1 \<in> Basis" |
|
5315 |
using SOME_Basis by auto |
|
71768 | 5316 |
have "c \<in> h ` ({0..1} \<times> sphere a r)" |
5317 |
proof |
|
5318 |
show "c = h (0, a + r *\<^sub>R b1)" |
|
5319 |
by (simp add: h) |
|
5320 |
show "(0, a + r *\<^sub>R b1) \<in> {0..1::real} \<times> sphere a r" |
|
5321 |
using greater \<open>b1 \<in> Basis\<close> by (auto simp: dist_norm) |
|
5322 |
qed |
|
5323 |
then have "c \<in> S" |
|
5324 |
using him by blast |
|
69620 | 5325 |
have uconth: "uniformly_continuous_on ({0..1::real} \<times> (sphere a r)) h" |
5326 |
by (force intro: compact_Times conth compact_uniformly_continuous) |
|
5327 |
let ?g = "\<lambda>x. h (norm (x - a)/r, |
|
5328 |
a + (if x = a then r *\<^sub>R b1 else (r / norm(x - a)) *\<^sub>R (x - a)))" |
|
5329 |
let ?g' = "\<lambda>x. h (norm (x - a)/r, a + (r / norm(x - a)) *\<^sub>R (x - a))" |
|
5330 |
show ?rhs |
|
5331 |
proof (intro exI conjI) |
|
5332 |
have "continuous_on (cball a r - {a}) ?g'" |
|
72372 | 5333 |
using greater |
5334 |
by (force simp: dist_norm norm_minus_commute intro: continuous_on_compose2 [OF conth] continuous_intros) |
|
69620 | 5335 |
then show "continuous_on (cball a r) ?g" |
5336 |
proof (rule nullhomotopic_from_lemma) |
|
5337 |
show "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> norm (?g' x - ?g a) < e" if "0 < e" for e |
|
5338 |
proof - |
|
5339 |
obtain d where "0 < d" |
|
72372 | 5340 |
and d: "\<And>x x'. \<lbrakk>x \<in> {0..1} \<times> sphere a r; x' \<in> {0..1} \<times> sphere a r; norm ( x' - x) < d\<rbrakk> |
5341 |
\<Longrightarrow> norm (h x' - h x) < e" |
|
5342 |
using uniformly_continuous_onE [OF uconth \<open>0 < e\<close>] by (auto simp: dist_norm) |
|
69620 | 5343 |
have *: "norm (h (norm (x - a) / r, |
72372 | 5344 |
a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + r *\<^sub>R b1)) < e" (is "norm (?ha - ?hb) < e") |
69620 | 5345 |
if "x \<noteq> a" "norm (x - a) < r" "norm (x - a) < d * r" for x |
5346 |
proof - |
|
72372 | 5347 |
have "norm (?ha - ?hb) = norm (?ha - h (0, a + (r / norm (x - a)) *\<^sub>R (x - a)))" |
69620 | 5348 |
by (simp add: h) |
5349 |
also have "\<dots> < e" |
|
5350 |
using greater \<open>0 < d\<close> \<open>b1 \<in> Basis\<close> that |
|
72372 | 5351 |
by (intro d) (simp_all add: dist_norm, simp add: field_simps) |
69620 | 5352 |
finally show ?thesis . |
5353 |
qed |
|
5354 |
show ?thesis |
|
72372 | 5355 |
using greater \<open>0 < d\<close> |
5356 |
by (rule_tac x = "min r (d * r)" in exI) (auto simp: *) |
|
69620 | 5357 |
qed |
5358 |
show "\<And>x. x \<in> cball a r \<and> x \<noteq> a \<Longrightarrow> ?g x = ?g' x" |
|
5359 |
by auto |
|
5360 |
qed |
|
5361 |
next |
|
5362 |
show "?g ` cball a r \<subseteq> S" |
|
5363 |
using greater him \<open>c \<in> S\<close> |
|
5364 |
by (force simp: h dist_norm norm_minus_commute) |
|
5365 |
next |
|
5366 |
show "\<forall>x\<in>sphere a r. ?g x = f x" |
|
5367 |
using greater by (auto simp: h dist_norm norm_minus_commute) |
|
5368 |
qed |
|
5369 |
next |
|
5370 |
assume ?rhs |
|
5371 |
then obtain g where contg: "continuous_on (cball a r) g" |
|
5372 |
and gim: "g ` cball a r \<subseteq> S" |
|
5373 |
and gf: "\<forall>x \<in> sphere a r. g x = f x" |
|
5374 |
by auto |
|
5375 |
let ?h = "\<lambda>y. g (a + (fst y) *\<^sub>R (snd y - a))" |
|
5376 |
have "continuous_on ({0..1} \<times> sphere a r) ?h" |
|
70196
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents:
70136
diff
changeset
|
5377 |
proof (rule continuous_on_compose2 [OF contg]) |
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents:
70136
diff
changeset
|
5378 |
show "continuous_on ({0..1} \<times> sphere a r) (\<lambda>x. a + fst x *\<^sub>R (snd x - a))" |
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents:
70136
diff
changeset
|
5379 |
by (intro continuous_intros) |
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents:
70136
diff
changeset
|
5380 |
qed (auto simp: dist_norm norm_minus_commute mult_left_le_one_le) |
69620 | 5381 |
moreover |
82323
b022c013b04b
Function space instead of image closure
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
5382 |
have "?h \<in> ({0..1} \<times> sphere a r) \<rightarrow> S" |
69620 | 5383 |
by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gim [THEN subsetD]) |
5384 |
moreover |
|
5385 |
have "\<forall>x\<in>sphere a r. ?h (0, x) = g a" "\<forall>x\<in>sphere a r. ?h (1, x) = f x" |
|
5386 |
by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gf) |
|
70196
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents:
70136
diff
changeset
|
5387 |
ultimately have "homotopic_with_canon (\<lambda>x. True) (sphere a r) S (\<lambda>x. g a) f" |
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents:
70136
diff
changeset
|
5388 |
by (auto simp: homotopic_with) |
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents:
70136
diff
changeset
|
5389 |
then show ?lhs |
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents:
70136
diff
changeset
|
5390 |
using homotopic_with_symD by blast |
69620 | 5391 |
qed |
5392 |
ultimately |
|
5393 |
show ?thesis by meson |
|
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
5394 |
qed |
69620 | 5395 |
|
5396 |
end |