| author | nipkow | 
| Thu, 11 Aug 2022 13:23:00 +0200 | |
| changeset 75805 | 3581dcee70db | 
| parent 74123 | 7c5842b06114 | 
| child 76836 | 30182f9e1818 | 
| permissions | -rw-r--r-- | 
| 63627 | 1 | (* Title: HOL/Analysis/Path_Connected.thy | 
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61762diff
changeset | 2 | Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light | 
| 36583 | 3 | *) | 
| 4 | ||
| 69620 | 5 | section \<open>Path-Connectedness\<close> | 
| 36583 | 6 | |
| 7 | theory Path_Connected | |
| 71236 | 8 | imports | 
| 9 | Starlike | |
| 10 | T1_Spaces | |
| 36583 | 11 | begin | 
| 12 | ||
| 60420 | 13 | subsection \<open>Paths and Arcs\<close> | 
| 36583 | 14 | |
| 70136 | 15 | definition\<^marker>\<open>tag important\<close> path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool" | 
| 53640 | 16 |   where "path g \<longleftrightarrow> continuous_on {0..1} g"
 | 
| 36583 | 17 | |
| 70136 | 18 | definition\<^marker>\<open>tag important\<close> pathstart :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a" | 
| 36583 | 19 | where "pathstart g = g 0" | 
| 20 | ||
| 70136 | 21 | definition\<^marker>\<open>tag important\<close> pathfinish :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a" | 
| 36583 | 22 | where "pathfinish g = g 1" | 
| 23 | ||
| 70136 | 24 | definition\<^marker>\<open>tag important\<close> path_image :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a set" | 
| 36583 | 25 |   where "path_image g = g ` {0 .. 1}"
 | 
| 26 | ||
| 70136 | 27 | definition\<^marker>\<open>tag important\<close> reversepath :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a" | 
| 36583 | 28 | where "reversepath g = (\<lambda>x. g(1 - x))" | 
| 29 | ||
| 70136 | 30 | definition\<^marker>\<open>tag important\<close> joinpaths :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> real \<Rightarrow> 'a" | 
| 36583 | 31 | (infixr "+++" 75) | 
| 32 | where "g1 +++ g2 = (\<lambda>x. if x \<le> 1/2 then g1 (2 * x) else g2 (2 * x - 1))" | |
| 33 | ||
| 70136 | 34 | definition\<^marker>\<open>tag important\<close> simple_path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool" | 
| 36583 | 35 | where "simple_path g \<longleftrightarrow> | 
| 60303 | 36 |      path g \<and> (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0)"
 | 
| 36583 | 37 | |
| 70136 | 38 | definition\<^marker>\<open>tag important\<close> arc :: "(real \<Rightarrow> 'a :: topological_space) \<Rightarrow> bool" | 
| 60303 | 39 |   where "arc g \<longleftrightarrow> path g \<and> inj_on g {0..1}"
 | 
| 36583 | 40 | |
| 49653 | 41 | |
| 70136 | 42 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Invariance theorems\<close> | 
| 60303 | 43 | |
| 44 | lemma path_eq: "path p \<Longrightarrow> (\<And>t. t \<in> {0..1} \<Longrightarrow> p t = q t) \<Longrightarrow> path q"
 | |
| 45 | using continuous_on_eq path_def by blast | |
| 46 | ||
| 68096 | 47 | lemma path_continuous_image: "path g \<Longrightarrow> continuous_on (path_image g) f \<Longrightarrow> path(f \<circ> g)" | 
| 60303 | 48 | unfolding path_def path_image_def | 
| 49 | using continuous_on_compose by blast | |
| 50 | ||
| 72256 | 51 | lemma continuous_on_translation_eq: | 
| 52 | fixes g :: "'a :: real_normed_vector \<Rightarrow> 'b :: real_normed_vector" | |
| 53 | shows "continuous_on A ((+) a \<circ> g) = continuous_on A g" | |
| 60303 | 54 | proof - | 
| 68096 | 55 | have g: "g = (\<lambda>x. -a + x) \<circ> ((\<lambda>x. a + x) \<circ> g)" | 
| 60303 | 56 | by (rule ext) simp | 
| 57 | show ?thesis | |
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73795diff
changeset | 58 | by (metis (no_types, opaque_lifting) g continuous_on_compose homeomorphism_def homeomorphism_translation) | 
| 60303 | 59 | qed | 
| 60 | ||
| 72256 | 61 | lemma path_translation_eq: | 
| 62 | fixes g :: "real \<Rightarrow> 'a :: real_normed_vector" | |
| 63 | shows "path((\<lambda>x. a + x) \<circ> g) = path g" | |
| 64 | using continuous_on_translation_eq path_def by blast | |
| 65 | ||
| 60303 | 66 | lemma path_linear_image_eq: | 
| 67 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 68 | assumes "linear f" "inj f" | |
| 68096 | 69 | shows "path(f \<circ> g) = path g" | 
| 60303 | 70 | proof - | 
| 71 | from linear_injective_left_inverse [OF assms] | |
| 72 | obtain h where h: "linear h" "h \<circ> f = id" | |
| 73 | by blast | |
| 68096 | 74 | then have g: "g = h \<circ> (f \<circ> g)" | 
| 60303 | 75 | by (metis comp_assoc id_comp) | 
| 76 | show ?thesis | |
| 77 | unfolding path_def | |
| 78 | using h assms | |
| 79 | by (metis g continuous_on_compose linear_continuous_on linear_conv_bounded_linear) | |
| 80 | qed | |
| 81 | ||
| 68096 | 82 | lemma pathstart_translation: "pathstart((\<lambda>x. a + x) \<circ> g) = a + pathstart g" | 
| 60303 | 83 | by (simp add: pathstart_def) | 
| 84 | ||
| 68096 | 85 | lemma pathstart_linear_image_eq: "linear f \<Longrightarrow> pathstart(f \<circ> g) = f(pathstart g)" | 
| 60303 | 86 | by (simp add: pathstart_def) | 
| 87 | ||
| 68096 | 88 | lemma pathfinish_translation: "pathfinish((\<lambda>x. a + x) \<circ> g) = a + pathfinish g" | 
| 60303 | 89 | by (simp add: pathfinish_def) | 
| 90 | ||
| 68096 | 91 | lemma pathfinish_linear_image: "linear f \<Longrightarrow> pathfinish(f \<circ> g) = f(pathfinish g)" | 
| 60303 | 92 | by (simp add: pathfinish_def) | 
| 93 | ||
| 68096 | 94 | lemma path_image_translation: "path_image((\<lambda>x. a + x) \<circ> g) = (\<lambda>x. a + x) ` (path_image g)" | 
| 60303 | 95 | by (simp add: image_comp path_image_def) | 
| 96 | ||
| 68096 | 97 | lemma path_image_linear_image: "linear f \<Longrightarrow> path_image(f \<circ> g) = f ` (path_image g)" | 
| 60303 | 98 | by (simp add: image_comp path_image_def) | 
| 99 | ||
| 68096 | 100 | lemma reversepath_translation: "reversepath((\<lambda>x. a + x) \<circ> g) = (\<lambda>x. a + x) \<circ> reversepath g" | 
| 60303 | 101 | by (rule ext) (simp add: reversepath_def) | 
| 36583 | 102 | |
| 68096 | 103 | lemma reversepath_linear_image: "linear f \<Longrightarrow> reversepath(f \<circ> g) = f \<circ> reversepath g" | 
| 60303 | 104 | by (rule ext) (simp add: reversepath_def) | 
| 105 | ||
| 106 | lemma joinpaths_translation: | |
| 68096 | 107 | "((\<lambda>x. a + x) \<circ> g1) +++ ((\<lambda>x. a + x) \<circ> g2) = (\<lambda>x. a + x) \<circ> (g1 +++ g2)" | 
| 60303 | 108 | by (rule ext) (simp add: joinpaths_def) | 
| 109 | ||
| 68096 | 110 | lemma joinpaths_linear_image: "linear f \<Longrightarrow> (f \<circ> g1) +++ (f \<circ> g2) = f \<circ> (g1 +++ g2)" | 
| 60303 | 111 | by (rule ext) (simp add: joinpaths_def) | 
| 112 | ||
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 113 | lemma simple_path_translation_eq: | 
| 60303 | 114 | fixes g :: "real \<Rightarrow> 'a::euclidean_space" | 
| 68096 | 115 | shows "simple_path((\<lambda>x. a + x) \<circ> g) = simple_path g" | 
| 60303 | 116 | by (simp add: simple_path_def path_translation_eq) | 
| 117 | ||
| 118 | lemma simple_path_linear_image_eq: | |
| 119 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 120 | assumes "linear f" "inj f" | |
| 68096 | 121 | shows "simple_path(f \<circ> g) = simple_path g" | 
| 60303 | 122 | using assms inj_on_eq_iff [of f] | 
| 123 | by (auto simp: path_linear_image_eq simple_path_def path_translation_eq) | |
| 124 | ||
| 125 | lemma arc_translation_eq: | |
| 126 | fixes g :: "real \<Rightarrow> 'a::euclidean_space" | |
| 68096 | 127 | shows "arc((\<lambda>x. a + x) \<circ> g) = arc g" | 
| 60303 | 128 | by (auto simp: arc_def inj_on_def path_translation_eq) | 
| 129 | ||
| 130 | lemma arc_linear_image_eq: | |
| 131 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 132 | assumes "linear f" "inj f" | |
| 68096 | 133 | shows "arc(f \<circ> g) = arc g" | 
| 60303 | 134 | using assms inj_on_eq_iff [of f] | 
| 135 | by (auto simp: arc_def inj_on_def path_linear_image_eq) | |
| 136 | ||
| 69514 | 137 | |
| 70136 | 138 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Basic lemmas about paths\<close> | 
| 60303 | 139 | |
| 74007 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 140 | lemma path_of_real: "path complex_of_real" | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 141 | unfolding path_def by (intro continuous_intros) | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 142 | |
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 143 | lemma path_const: "path (\<lambda>t. a)" for a::"'a::real_normed_vector" | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 144 | unfolding path_def by (intro continuous_intros) | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 145 | |
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 146 | lemma path_minus: "path g \<Longrightarrow> path (\<lambda>t. - g t)" for g::"real\<Rightarrow>'a::real_normed_vector" | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 147 | unfolding path_def by (intro continuous_intros) | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 148 | |
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 149 | lemma path_add: "\<lbrakk>path f; path g\<rbrakk> \<Longrightarrow> path (\<lambda>t. f t + g t)" for f::"real\<Rightarrow>'a::real_normed_vector" | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 150 | unfolding path_def by (intro continuous_intros) | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 151 | |
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 152 | lemma path_diff: "\<lbrakk>path f; path g\<rbrakk> \<Longrightarrow> path (\<lambda>t. f t - g t)" for f::"real\<Rightarrow>'a::real_normed_vector" | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 153 | unfolding path_def by (intro continuous_intros) | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 154 | |
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 155 | lemma path_mult: "\<lbrakk>path f; path g\<rbrakk> \<Longrightarrow> path (\<lambda>t. f t * g t)" for f::"real\<Rightarrow>'a::real_normed_field" | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 156 | unfolding path_def by (intro continuous_intros) | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 157 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 158 | lemma pathin_iff_path_real [simp]: "pathin euclideanreal g \<longleftrightarrow> path g" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 159 | by (simp add: pathin_def path_def) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 160 | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 161 | lemma continuous_on_path: "path f \<Longrightarrow> t \<subseteq> {0..1} \<Longrightarrow> continuous_on t f"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 162 | using continuous_on_subset path_def by blast | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 163 | |
| 60303 | 164 | lemma arc_imp_simple_path: "arc g \<Longrightarrow> simple_path g" | 
| 165 | by (simp add: arc_def inj_on_def simple_path_def) | |
| 166 | ||
| 167 | lemma arc_imp_path: "arc g \<Longrightarrow> path g" | |
| 168 | using arc_def by blast | |
| 169 | ||
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 170 | lemma arc_imp_inj_on: "arc g \<Longrightarrow> inj_on g {0..1}"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 171 | by (auto simp: arc_def) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 172 | |
| 60303 | 173 | lemma simple_path_imp_path: "simple_path g \<Longrightarrow> path g" | 
| 174 | using simple_path_def by blast | |
| 175 | ||
| 176 | lemma simple_path_cases: "simple_path g \<Longrightarrow> arc g \<or> pathfinish g = pathstart g" | |
| 177 | unfolding simple_path_def arc_def inj_on_def pathfinish_def pathstart_def | |
| 68096 | 178 | by force | 
| 60303 | 179 | |
| 180 | lemma simple_path_imp_arc: "simple_path g \<Longrightarrow> pathfinish g \<noteq> pathstart g \<Longrightarrow> arc g" | |
| 181 | using simple_path_cases by auto | |
| 182 | ||
| 183 | lemma arc_distinct_ends: "arc g \<Longrightarrow> pathfinish g \<noteq> pathstart g" | |
| 184 | unfolding arc_def inj_on_def pathfinish_def pathstart_def | |
| 185 | by fastforce | |
| 186 | ||
| 187 | lemma arc_simple_path: "arc g \<longleftrightarrow> simple_path g \<and> pathfinish g \<noteq> pathstart g" | |
| 188 | using arc_distinct_ends arc_imp_simple_path simple_path_cases by blast | |
| 189 | ||
| 190 | lemma simple_path_eq_arc: "pathfinish g \<noteq> pathstart g \<Longrightarrow> (simple_path g = arc g)" | |
| 191 | by (simp add: arc_simple_path) | |
| 36583 | 192 | |
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 193 | lemma path_image_const [simp]: "path_image (\<lambda>t. a) = {a}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 194 | by (force simp: path_image_def) | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 195 | |
| 60974 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 paulson <lp15@cam.ac.uk> parents: 
60809diff
changeset | 196 | lemma path_image_nonempty [simp]: "path_image g \<noteq> {}"
 | 
| 56188 | 197 | unfolding path_image_def image_is_empty box_eq_empty | 
| 53640 | 198 | by auto | 
| 36583 | 199 | |
| 53640 | 200 | lemma pathstart_in_path_image[intro]: "pathstart g \<in> path_image g" | 
| 201 | unfolding pathstart_def path_image_def | |
| 202 | by auto | |
| 36583 | 203 | |
| 53640 | 204 | lemma pathfinish_in_path_image[intro]: "pathfinish g \<in> path_image g" | 
| 205 | unfolding pathfinish_def path_image_def | |
| 206 | by auto | |
| 207 | ||
| 208 | lemma connected_path_image[intro]: "path g \<Longrightarrow> connected (path_image g)" | |
| 36583 | 209 | unfolding path_def path_image_def | 
| 60303 | 210 | using connected_continuous_image connected_Icc by blast | 
| 36583 | 211 | |
| 53640 | 212 | lemma compact_path_image[intro]: "path g \<Longrightarrow> compact (path_image g)" | 
| 36583 | 213 | unfolding path_def path_image_def | 
| 60303 | 214 | using compact_continuous_image connected_Icc by blast | 
| 36583 | 215 | |
| 53640 | 216 | lemma reversepath_reversepath[simp]: "reversepath (reversepath g) = g" | 
| 217 | unfolding reversepath_def | |
| 218 | by auto | |
| 36583 | 219 | |
| 53640 | 220 | lemma pathstart_reversepath[simp]: "pathstart (reversepath g) = pathfinish g" | 
| 221 | unfolding pathstart_def reversepath_def pathfinish_def | |
| 222 | by auto | |
| 36583 | 223 | |
| 53640 | 224 | lemma pathfinish_reversepath[simp]: "pathfinish (reversepath g) = pathstart g" | 
| 225 | unfolding pathstart_def reversepath_def pathfinish_def | |
| 226 | by auto | |
| 36583 | 227 | |
| 73795 | 228 | lemma reversepath_o: "reversepath g = g \<circ> (-)1" | 
| 229 | by (auto simp: reversepath_def) | |
| 230 | ||
| 49653 | 231 | lemma pathstart_join[simp]: "pathstart (g1 +++ g2) = pathstart g1" | 
| 53640 | 232 | unfolding pathstart_def joinpaths_def pathfinish_def | 
| 233 | by auto | |
| 36583 | 234 | |
| 49653 | 235 | lemma pathfinish_join[simp]: "pathfinish (g1 +++ g2) = pathfinish g2" | 
| 53640 | 236 | unfolding pathstart_def joinpaths_def pathfinish_def | 
| 237 | by auto | |
| 36583 | 238 | |
| 53640 | 239 | lemma path_image_reversepath[simp]: "path_image (reversepath g) = path_image g" | 
| 49653 | 240 | proof - | 
| 53640 | 241 | have *: "\<And>g. path_image (reversepath g) \<subseteq> path_image g" | 
| 49653 | 242 | unfolding path_image_def subset_eq reversepath_def Ball_def image_iff | 
| 60303 | 243 | by force | 
| 49653 | 244 | show ?thesis | 
| 245 | using *[of g] *[of "reversepath g"] | |
| 53640 | 246 | unfolding reversepath_reversepath | 
| 247 | by auto | |
| 49653 | 248 | qed | 
| 36583 | 249 | |
| 53640 | 250 | lemma path_reversepath [simp]: "path (reversepath g) \<longleftrightarrow> path g" | 
| 49653 | 251 | proof - | 
| 252 | have *: "\<And>g. path g \<Longrightarrow> path (reversepath g)" | |
| 253 | unfolding path_def reversepath_def | |
| 254 | apply (rule continuous_on_compose[unfolded o_def, of _ "\<lambda>x. 1 - x"]) | |
| 68096 | 255 |     apply (auto intro: continuous_intros continuous_on_subset[of "{0..1}"])
 | 
| 49653 | 256 | done | 
| 257 | show ?thesis | |
| 72256 | 258 | using "*" by force | 
| 49653 | 259 | qed | 
| 260 | ||
| 60303 | 261 | lemma arc_reversepath: | 
| 262 | assumes "arc g" shows "arc(reversepath g)" | |
| 263 | proof - | |
| 264 |   have injg: "inj_on g {0..1}"
 | |
| 265 | using assms | |
| 266 | by (simp add: arc_def) | |
| 267 | have **: "\<And>x y::real. 1-x = 1-y \<Longrightarrow> x = y" | |
| 268 | by simp | |
| 269 | show ?thesis | |
| 68096 | 270 | using assms by (clarsimp simp: arc_def intro!: inj_onI) (simp add: inj_onD reversepath_def **) | 
| 60303 | 271 | qed | 
| 272 | ||
| 273 | lemma simple_path_reversepath: "simple_path g \<Longrightarrow> simple_path (reversepath g)" | |
| 274 | apply (simp add: simple_path_def) | |
| 275 | apply (force simp: reversepath_def) | |
| 276 | done | |
| 277 | ||
| 49653 | 278 | lemmas reversepath_simps = | 
| 279 | path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath | |
| 36583 | 280 | |
| 49653 | 281 | lemma path_join[simp]: | 
| 282 | assumes "pathfinish g1 = pathstart g2" | |
| 283 | shows "path (g1 +++ g2) \<longleftrightarrow> path g1 \<and> path g2" | |
| 284 | unfolding path_def pathfinish_def pathstart_def | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
50935diff
changeset | 285 | proof safe | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
50935diff
changeset | 286 |   assume cont: "continuous_on {0..1} (g1 +++ g2)"
 | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
50935diff
changeset | 287 |   have g1: "continuous_on {0..1} g1 \<longleftrightarrow> continuous_on {0..1} ((g1 +++ g2) \<circ> (\<lambda>x. x / 2))"
 | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
50935diff
changeset | 288 | by (intro continuous_on_cong refl) (auto simp: joinpaths_def) | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
50935diff
changeset | 289 |   have g2: "continuous_on {0..1} g2 \<longleftrightarrow> continuous_on {0..1} ((g1 +++ g2) \<circ> (\<lambda>x. x / 2 + 1/2))"
 | 
| 53640 | 290 | using assms | 
| 291 | by (intro continuous_on_cong refl) (auto simp: joinpaths_def pathfinish_def pathstart_def) | |
| 292 |   show "continuous_on {0..1} g1" and "continuous_on {0..1} g2"
 | |
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51478diff
changeset | 293 | unfolding g1 g2 | 
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56188diff
changeset | 294 | by (auto intro!: continuous_intros continuous_on_subset[OF cont] simp del: o_apply) | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
50935diff
changeset | 295 | next | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
50935diff
changeset | 296 |   assume g1g2: "continuous_on {0..1} g1" "continuous_on {0..1} g2"
 | 
| 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
50935diff
changeset | 297 |   have 01: "{0 .. 1} = {0..1/2} \<union> {1/2 .. 1::real}"
 | 
| 36583 | 298 | by auto | 
| 53640 | 299 |   {
 | 
| 300 | fix x :: real | |
| 301 | assume "0 \<le> x" and "x \<le> 1" | |
| 302 |     then have "x \<in> (\<lambda>x. x * 2) ` {0..1 / 2}"
 | |
| 303 | by (intro image_eqI[where x="x/2"]) auto | |
| 304 | } | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
50935diff
changeset | 305 | note 1 = this | 
| 53640 | 306 |   {
 | 
| 307 | fix x :: real | |
| 308 | assume "0 \<le> x" and "x \<le> 1" | |
| 309 |     then have "x \<in> (\<lambda>x. x * 2 - 1) ` {1 / 2..1}"
 | |
| 310 | by (intro image_eqI[where x="x/2 + 1/2"]) auto | |
| 311 | } | |
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
50935diff
changeset | 312 | note 2 = this | 
| 49653 | 313 |   show "continuous_on {0..1} (g1 +++ g2)"
 | 
| 53640 | 314 | using assms | 
| 315 | unfolding joinpaths_def 01 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56188diff
changeset | 316 | apply (intro continuous_on_cases closed_atLeastAtMost g1g2[THEN continuous_on_compose2] continuous_intros) | 
| 53640 | 317 | apply (auto simp: field_simps pathfinish_def pathstart_def intro!: 1 2) | 
| 318 | done | |
| 49653 | 319 | qed | 
| 36583 | 320 | |
| 69514 | 321 | |
| 70136 | 322 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Path Images\<close> | 
| 60303 | 323 | |
| 324 | lemma bounded_path_image: "path g \<Longrightarrow> bounded(path_image g)" | |
| 325 | by (simp add: compact_imp_bounded compact_path_image) | |
| 326 | ||
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 327 | lemma closed_path_image: | 
| 60303 | 328 | fixes g :: "real \<Rightarrow> 'a::t2_space" | 
| 329 | shows "path g \<Longrightarrow> closed(path_image g)" | |
| 330 | by (metis compact_path_image compact_imp_closed) | |
| 331 | ||
| 332 | lemma connected_simple_path_image: "simple_path g \<Longrightarrow> connected(path_image g)" | |
| 333 | by (metis connected_path_image simple_path_imp_path) | |
| 334 | ||
| 335 | lemma compact_simple_path_image: "simple_path g \<Longrightarrow> compact(path_image g)" | |
| 336 | by (metis compact_path_image simple_path_imp_path) | |
| 337 | ||
| 338 | lemma bounded_simple_path_image: "simple_path g \<Longrightarrow> bounded(path_image g)" | |
| 339 | by (metis bounded_path_image simple_path_imp_path) | |
| 340 | ||
| 341 | lemma closed_simple_path_image: | |
| 342 | fixes g :: "real \<Rightarrow> 'a::t2_space" | |
| 343 | shows "simple_path g \<Longrightarrow> closed(path_image g)" | |
| 344 | by (metis closed_path_image simple_path_imp_path) | |
| 345 | ||
| 346 | lemma connected_arc_image: "arc g \<Longrightarrow> connected(path_image g)" | |
| 347 | by (metis connected_path_image arc_imp_path) | |
| 348 | ||
| 349 | lemma compact_arc_image: "arc g \<Longrightarrow> compact(path_image g)" | |
| 350 | by (metis compact_path_image arc_imp_path) | |
| 351 | ||
| 352 | lemma bounded_arc_image: "arc g \<Longrightarrow> bounded(path_image g)" | |
| 353 | by (metis bounded_path_image arc_imp_path) | |
| 354 | ||
| 355 | lemma closed_arc_image: | |
| 356 | fixes g :: "real \<Rightarrow> 'a::t2_space" | |
| 357 | shows "arc g \<Longrightarrow> closed(path_image g)" | |
| 358 | by (metis closed_path_image arc_imp_path) | |
| 359 | ||
| 53640 | 360 | lemma path_image_join_subset: "path_image (g1 +++ g2) \<subseteq> path_image g1 \<union> path_image g2" | 
| 361 | unfolding path_image_def joinpaths_def | |
| 362 | by auto | |
| 36583 | 363 | |
| 364 | lemma subset_path_image_join: | |
| 53640 | 365 | assumes "path_image g1 \<subseteq> s" | 
| 366 | and "path_image g2 \<subseteq> s" | |
| 367 | shows "path_image (g1 +++ g2) \<subseteq> s" | |
| 368 | using path_image_join_subset[of g1 g2] and assms | |
| 369 | by auto | |
| 36583 | 370 | |
| 371 | lemma path_image_join: | |
| 72256 | 372 | assumes "pathfinish g1 = pathstart g2" | 
| 373 | shows "path_image(g1 +++ g2) = path_image g1 \<union> path_image g2" | |
| 374 | proof - | |
| 375 | have "path_image g1 \<subseteq> path_image (g1 +++ g2)" | |
| 376 | proof (clarsimp simp: path_image_def joinpaths_def) | |
| 377 | fix u::real | |
| 378 | assume "0 \<le> u" "u \<le> 1" | |
| 379 |     then show "g1 u \<in> (\<lambda>x. g1 (2 * x)) ` ({0..1} \<inter> {x. x * 2 \<le> 1})"
 | |
| 380 | by (rule_tac x="u/2" in image_eqI) auto | |
| 381 | qed | |
| 382 | moreover | |
| 383 |   have \<section>: "g2 u \<in> (\<lambda>x. g2 (2 * x - 1)) ` ({0..1} \<inter> {x. \<not> x * 2 \<le> 1})" 
 | |
| 384 | if "0 < u" "u \<le> 1" for u | |
| 385 | using that assms | |
| 386 | by (rule_tac x="(u+1)/2" in image_eqI) (auto simp: field_simps pathfinish_def pathstart_def) | |
| 387 |   have "g2 0 \<in> (\<lambda>x. g1 (2 * x)) ` ({0..1} \<inter> {x. x * 2 \<le> 1})"
 | |
| 388 | using assms | |
| 389 | by (rule_tac x="1/2" in image_eqI) (auto simp: pathfinish_def pathstart_def) | |
| 390 | then have "path_image g2 \<subseteq> path_image (g1 +++ g2)" | |
| 391 | by (auto simp: path_image_def joinpaths_def intro!: \<section>) | |
| 392 | ultimately show ?thesis | |
| 393 | using path_image_join_subset by blast | |
| 394 | qed | |
| 36583 | 395 | |
| 396 | lemma not_in_path_image_join: | |
| 53640 | 397 | assumes "x \<notin> path_image g1" | 
| 398 | and "x \<notin> path_image g2" | |
| 399 | shows "x \<notin> path_image (g1 +++ g2)" | |
| 400 | using assms and path_image_join_subset[of g1 g2] | |
| 401 | by auto | |
| 36583 | 402 | |
| 68096 | 403 | lemma pathstart_compose: "pathstart(f \<circ> p) = f(pathstart p)" | 
| 60303 | 404 | by (simp add: pathstart_def) | 
| 405 | ||
| 68096 | 406 | lemma pathfinish_compose: "pathfinish(f \<circ> p) = f(pathfinish p)" | 
| 60303 | 407 | by (simp add: pathfinish_def) | 
| 408 | ||
| 68096 | 409 | lemma path_image_compose: "path_image (f \<circ> p) = f ` (path_image p)" | 
| 60303 | 410 | by (simp add: image_comp path_image_def) | 
| 411 | ||
| 68096 | 412 | lemma path_compose_join: "f \<circ> (p +++ q) = (f \<circ> p) +++ (f \<circ> q)" | 
| 60303 | 413 | by (rule ext) (simp add: joinpaths_def) | 
| 414 | ||
| 68096 | 415 | lemma path_compose_reversepath: "f \<circ> reversepath p = reversepath(f \<circ> p)" | 
| 60303 | 416 | by (rule ext) (simp add: reversepath_def) | 
| 417 | ||
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61738diff
changeset | 418 | lemma joinpaths_eq: | 
| 60303 | 419 |   "(\<And>t. t \<in> {0..1} \<Longrightarrow> p t = p' t) \<Longrightarrow>
 | 
| 420 |    (\<And>t. t \<in> {0..1} \<Longrightarrow> q t = q' t)
 | |
| 421 |    \<Longrightarrow>  t \<in> {0..1} \<Longrightarrow> (p +++ q) t = (p' +++ q') t"
 | |
| 422 | by (auto simp: joinpaths_def) | |
| 423 | ||
| 424 | lemma simple_path_inj_on: "simple_path g \<Longrightarrow> inj_on g {0<..<1}"
 | |
| 425 | by (auto simp: simple_path_def path_image_def inj_on_def less_eq_real_def Ball_def) | |
| 426 | ||
| 427 | ||
| 70136 | 428 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Simple paths with the endpoints removed\<close> | 
| 60303 | 429 | |
| 430 | lemma simple_path_endless: | |
| 72256 | 431 | assumes "simple_path c" | 
| 432 |   shows "path_image c - {pathstart c,pathfinish c} = c ` {0<..<1}" (is "?lhs = ?rhs")
 | |
| 433 | proof | |
| 434 | show "?lhs \<subseteq> ?rhs" | |
| 435 | using less_eq_real_def by (auto simp: path_image_def pathstart_def pathfinish_def) | |
| 436 | show "?rhs \<subseteq> ?lhs" | |
| 437 | using assms | |
| 438 | apply (auto simp: simple_path_def path_image_def pathstart_def pathfinish_def Ball_def) | |
| 439 | using less_eq_real_def zero_le_one by blast+ | |
| 440 | qed | |
| 36583 | 441 | |
| 60303 | 442 | lemma connected_simple_path_endless: | 
| 72256 | 443 | assumes "simple_path c" | 
| 444 |   shows "connected(path_image c - {pathstart c,pathfinish c})"
 | |
| 445 | proof - | |
| 446 |   have "continuous_on {0<..<1} c"
 | |
| 447 | using assms by (simp add: simple_path_def continuous_on_path path_def subset_iff) | |
| 448 |   then have "connected (c ` {0<..<1})"
 | |
| 449 | using connected_Ioo connected_continuous_image by blast | |
| 450 | then show ?thesis | |
| 451 | using assms by (simp add: simple_path_endless) | |
| 452 | qed | |
| 60303 | 453 | |
| 454 | lemma nonempty_simple_path_endless: | |
| 455 |     "simple_path c \<Longrightarrow> path_image c - {pathstart c,pathfinish c} \<noteq> {}"
 | |
| 456 | by (simp add: simple_path_endless) | |
| 457 | ||
| 458 | ||
| 70136 | 459 | subsection\<^marker>\<open>tag unimportant\<close>\<open>The operations on paths\<close> | 
| 60303 | 460 | |
| 461 | lemma path_image_subset_reversepath: "path_image(reversepath g) \<le> path_image g" | |
| 72256 | 462 | by simp | 
| 60303 | 463 | |
| 464 | lemma path_imp_reversepath: "path g \<Longrightarrow> path(reversepath g)" | |
| 72256 | 465 | by simp | 
| 60303 | 466 | |
| 61204 | 467 | lemma half_bounded_equal: "1 \<le> x * 2 \<Longrightarrow> x * 2 \<le> 1 \<longleftrightarrow> x = (1/2::real)" | 
| 468 | by simp | |
| 60303 | 469 | |
| 470 | lemma continuous_on_joinpaths: | |
| 471 |   assumes "continuous_on {0..1} g1" "continuous_on {0..1} g2" "pathfinish g1 = pathstart g2"
 | |
| 472 |     shows "continuous_on {0..1} (g1 +++ g2)"
 | |
| 473 | proof - | |
| 72256 | 474 |   have "{0..1::real} = {0..1/2} \<union> {1/2..1}"
 | 
| 60303 | 475 | by auto | 
| 72256 | 476 | then show ?thesis | 
| 477 | using assms by (metis path_def path_join) | |
| 60303 | 478 | qed | 
| 479 | ||
| 480 | lemma path_join_imp: "\<lbrakk>path g1; path g2; pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> path(g1 +++ g2)" | |
| 72256 | 481 | by simp | 
| 60303 | 482 | |
| 36583 | 483 | lemma simple_path_join_loop: | 
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 484 | assumes "arc g1" "arc g2" | 
| 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 485 | "pathfinish g1 = pathstart g2" "pathfinish g2 = pathstart g1" | 
| 60303 | 486 |           "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
 | 
| 487 | shows "simple_path(g1 +++ g2)" | |
| 488 | proof - | |
| 489 |   have injg1: "inj_on g1 {0..1}"
 | |
| 490 | using assms | |
| 491 | by (simp add: arc_def) | |
| 492 |   have injg2: "inj_on g2 {0..1}"
 | |
| 493 | using assms | |
| 494 | by (simp add: arc_def) | |
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 495 | have g12: "g1 1 = g2 0" | 
| 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 496 | and g21: "g2 1 = g1 0" | 
| 60303 | 497 |    and sb:  "g1 ` {0..1} \<inter> g2 ` {0..1} \<subseteq> {g1 0, g2 0}"
 | 
| 498 | using assms | |
| 499 | by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) | |
| 500 |   { fix x and y::real
 | |
| 72256 | 501 | assume g2_eq: "g2 (2 * x - 1) = g1 (2 * y)" | 
| 502 | and xyI: "x \<noteq> 1 \<or> y \<noteq> 0" | |
| 503 | and xy: "x \<le> 1" "0 \<le> y" " y * 2 \<le> 1" "\<not> x * 2 \<le> 1" | |
| 504 | then consider "g1 (2 * y) = g1 0" | "g1 (2 * y) = g2 0" | |
| 505 | using sb by force | |
| 506 | then have False | |
| 507 | proof cases | |
| 508 | case 1 | |
| 509 | then have "y = 0" | |
| 510 | using xy g2_eq by (auto dest!: inj_onD [OF injg1]) | |
| 511 | then show ?thesis | |
| 512 | using xy g2_eq xyI by (auto dest: inj_onD [OF injg2] simp flip: g21) | |
| 513 | next | |
| 514 | case 2 | |
| 515 | then have "2*x = 1" | |
| 516 | using g2_eq g12 inj_onD [OF injg2] atLeastAtMost_iff xy(1) xy(4) by fastforce | |
| 517 | with xy show False by auto | |
| 518 | qed | |
| 519 | } note * = this | |
| 60303 | 520 |   { fix x and y::real
 | 
| 72256 | 521 | assume xy: "g1 (2 * x) = g2 (2 * y - 1)" "y \<le> 1" "0 \<le> x" "\<not> y * 2 \<le> 1" "x * 2 \<le> 1" | 
| 522 | then have "x = 0 \<and> y = 1" | |
| 523 | using * xy by force | |
| 60303 | 524 | } note ** = this | 
| 525 | show ?thesis | |
| 526 | using assms | |
| 72256 | 527 | apply (simp add: arc_def simple_path_def) | 
| 528 | apply (auto simp: joinpaths_def split: if_split_asm | |
| 529 | dest!: * ** dest: inj_onD [OF injg1] inj_onD [OF injg2]) | |
| 60303 | 530 | done | 
| 531 | qed | |
| 532 | ||
| 533 | lemma arc_join: | |
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 534 | assumes "arc g1" "arc g2" | 
| 60303 | 535 | "pathfinish g1 = pathstart g2" | 
| 536 |           "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g2}"
 | |
| 537 | shows "arc(g1 +++ g2)" | |
| 538 | proof - | |
| 539 |   have injg1: "inj_on g1 {0..1}"
 | |
| 540 | using assms | |
| 541 | by (simp add: arc_def) | |
| 542 |   have injg2: "inj_on g2 {0..1}"
 | |
| 543 | using assms | |
| 544 | by (simp add: arc_def) | |
| 545 | have g11: "g1 1 = g2 0" | |
| 546 |    and sb:  "g1 ` {0..1} \<inter> g2 ` {0..1} \<subseteq> {g2 0}"
 | |
| 547 | using assms | |
| 548 | by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) | |
| 549 |   { fix x and y::real
 | |
| 72256 | 550 | assume xy: "g2 (2 * x - 1) = g1 (2 * y)" "x \<le> 1" "0 \<le> y" " y * 2 \<le> 1" "\<not> x * 2 \<le> 1" | 
| 551 | then have "g1 (2 * y) = g2 0" | |
| 552 | using sb by force | |
| 553 | then have False | |
| 554 | using xy inj_onD injg2 by fastforce | |
| 60303 | 555 | } note * = this | 
| 556 | show ?thesis | |
| 72256 | 557 | using assms | 
| 60303 | 558 | apply (simp add: arc_def inj_on_def) | 
| 72256 | 559 | apply (auto simp: joinpaths_def arc_imp_path split: if_split_asm | 
| 560 | dest: * *[OF sym] inj_onD [OF injg1] inj_onD [OF injg2]) | |
| 60303 | 561 | done | 
| 562 | qed | |
| 563 | ||
| 564 | lemma reversepath_joinpaths: | |
| 565 | "pathfinish g1 = pathstart g2 \<Longrightarrow> reversepath(g1 +++ g2) = reversepath g2 +++ reversepath g1" | |
| 566 | unfolding reversepath_def pathfinish_def pathstart_def joinpaths_def | |
| 567 | by (rule ext) (auto simp: mult.commute) | |
| 568 | ||
| 569 | ||
| 70136 | 570 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Some reversed and "if and only if" versions of joining theorems\<close> | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 571 | |
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 572 | lemma path_join_path_ends: | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 573 | fixes g1 :: "real \<Rightarrow> 'a::metric_space" | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 574 | assumes "path(g1 +++ g2)" "path g2" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 575 | shows "pathfinish g1 = pathstart g2" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 576 | proof (rule ccontr) | 
| 63040 | 577 | define e where "e = dist (g1 1) (g2 0)" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 578 | assume Neg: "pathfinish g1 \<noteq> pathstart g2" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 579 | then have "0 < dist (pathfinish g1) (pathstart g2)" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 580 | by auto | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 581 | then have "e > 0" | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 582 | by (metis e_def pathfinish_def pathstart_def) | 
| 72256 | 583 |   then have "\<forall>e>0. \<exists>d>0. \<forall>x'\<in>{0..1}. dist x' 0 < d \<longrightarrow> dist (g2 x') (g2 0) < e"
 | 
| 584 | using \<open>path g2\<close> atLeastAtMost_iff zero_le_one unfolding path_def continuous_on_iff | |
| 585 | by blast | |
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 586 | then obtain d1 where "d1 > 0" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 587 |        and d1: "\<And>x'. \<lbrakk>x'\<in>{0..1}; norm x' < d1\<rbrakk> \<Longrightarrow> dist (g2 x') (g2 0) < e/2"
 | 
| 72256 | 588 | by (metis \<open>0 < e\<close> half_gt_zero_iff norm_conv_dist) | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 589 | obtain d2 where "d2 > 0" | 
| 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 590 |        and d2: "\<And>x'. \<lbrakk>x'\<in>{0..1}; dist x' (1/2) < d2\<rbrakk>
 | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 591 | \<Longrightarrow> dist ((g1 +++ g2) x') (g1 1) < e/2" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 592 | using assms(1) \<open>e > 0\<close> unfolding path_def continuous_on_iff | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 593 | apply (drule_tac x="1/2" in bspec, simp) | 
| 72256 | 594 | apply (drule_tac x="e/2" in spec, force simp: joinpaths_def) | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 595 | done | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 596 |   have int01_1: "min (1/2) (min d1 d2) / 2 \<in> {0..1}"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 597 | using \<open>d1 > 0\<close> \<open>d2 > 0\<close> by (simp add: min_def) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 598 | have dist1: "norm (min (1 / 2) (min d1 d2) / 2) < d1" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 599 | using \<open>d1 > 0\<close> \<open>d2 > 0\<close> by (simp add: min_def dist_norm) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 600 |   have int01_2: "1/2 + min (1/2) (min d1 d2) / 4 \<in> {0..1}"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 601 | using \<open>d1 > 0\<close> \<open>d2 > 0\<close> by (simp add: min_def) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 602 | have dist2: "dist (1 / 2 + min (1 / 2) (min d1 d2) / 4) (1 / 2) < d2" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 603 | using \<open>d1 > 0\<close> \<open>d2 > 0\<close> by (simp add: min_def dist_norm) | 
| 69508 | 604 | have [simp]: "\<not> min (1 / 2) (min d1 d2) \<le> 0" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 605 | using \<open>d1 > 0\<close> \<open>d2 > 0\<close> by (simp add: min_def) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 606 | have "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g1 1) < e/2" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 607 | "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g2 0) < e/2" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 608 | using d1 [OF int01_1 dist1] d2 [OF int01_2 dist2] by (simp_all add: joinpaths_def) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 609 | then have "dist (g1 1) (g2 0) < e/2 + e/2" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 610 | using dist_triangle_half_r e_def by blast | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 611 | then show False | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 612 | by (simp add: e_def [symmetric]) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 613 | qed | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 614 | |
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 615 | lemma path_join_eq [simp]: | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 616 | fixes g1 :: "real \<Rightarrow> 'a::metric_space" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 617 | assumes "path g1" "path g2" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 618 | shows "path(g1 +++ g2) \<longleftrightarrow> pathfinish g1 = pathstart g2" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 619 | using assms by (metis path_join_path_ends path_join_imp) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 620 | |
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 621 | lemma simple_path_joinE: | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 622 | assumes "simple_path(g1 +++ g2)" and "pathfinish g1 = pathstart g2" | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 623 | obtains "arc g1" "arc g2" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 624 |           "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 625 | proof - | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 626 | have *: "\<And>x y. \<lbrakk>0 \<le> x; x \<le> 1; 0 \<le> y; y \<le> 1; (g1 +++ g2) x = (g1 +++ g2) y\<rbrakk> | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 627 | \<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 628 | using assms by (simp add: simple_path_def) | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 629 | have "path g1" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 630 | using assms path_join simple_path_imp_path by blast | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 631 |   moreover have "inj_on g1 {0..1}"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 632 | proof (clarsimp simp: inj_on_def) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 633 | fix x y | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 634 | assume "g1 x = g1 y" "0 \<le> x" "x \<le> 1" "0 \<le> y" "y \<le> 1" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 635 | then show "x = y" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 636 | using * [of "x/2" "y/2"] by (simp add: joinpaths_def split_ifs) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 637 | qed | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 638 | ultimately have "arc g1" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 639 | using assms by (simp add: arc_def) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 640 | have [simp]: "g2 0 = g1 1" | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 641 | using assms by (metis pathfinish_def pathstart_def) | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 642 | have "path g2" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 643 | using assms path_join simple_path_imp_path by blast | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 644 |   moreover have "inj_on g2 {0..1}"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 645 | proof (clarsimp simp: inj_on_def) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 646 | fix x y | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 647 | assume "g2 x = g2 y" "0 \<le> x" "x \<le> 1" "0 \<le> y" "y \<le> 1" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 648 | then show "x = y" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 649 | using * [of "(x + 1) / 2" "(y + 1) / 2"] | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 650 | by (force simp: joinpaths_def split_ifs field_split_simps) | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 651 | qed | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 652 | ultimately have "arc g2" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 653 | using assms by (simp add: arc_def) | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 654 | have "g2 y = g1 0 \<or> g2 y = g1 1" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 655 | if "g1 x = g2 y" "0 \<le> x" "x \<le> 1" "0 \<le> y" "y \<le> 1" for x y | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 656 | using * [of "x / 2" "(y + 1) / 2"] that | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 657 | by (auto simp: joinpaths_def split_ifs field_split_simps) | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 658 |   then have "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 659 | by (fastforce simp: pathstart_def pathfinish_def path_image_def) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 660 | with \<open>arc g1\<close> \<open>arc g2\<close> show ?thesis using that by blast | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 661 | qed | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 662 | |
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 663 | lemma simple_path_join_loop_eq: | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 664 | assumes "pathfinish g2 = pathstart g1" "pathfinish g1 = pathstart g2" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 665 | shows "simple_path(g1 +++ g2) \<longleftrightarrow> | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 666 |              arc g1 \<and> arc g2 \<and> path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 667 | by (metis assms simple_path_joinE simple_path_join_loop) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 668 | |
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 669 | lemma arc_join_eq: | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 670 | assumes "pathfinish g1 = pathstart g2" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 671 | shows "arc(g1 +++ g2) \<longleftrightarrow> | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 672 |            arc g1 \<and> arc g2 \<and> path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g2}"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 673 | (is "?lhs = ?rhs") | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 674 | proof | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 675 | assume ?lhs | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 676 | then have "simple_path(g1 +++ g2)" by (rule arc_imp_simple_path) | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 677 | then have *: "\<And>x y. \<lbrakk>0 \<le> x; x \<le> 1; 0 \<le> y; y \<le> 1; (g1 +++ g2) x = (g1 +++ g2) y\<rbrakk> | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 678 | \<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 679 | using assms by (simp add: simple_path_def) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 680 | have False if "g1 0 = g2 u" "0 \<le> u" "u \<le> 1" for u | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 681 | using * [of 0 "(u + 1) / 2"] that assms arc_distinct_ends [OF \<open>?lhs\<close>] | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 682 | by (auto simp: joinpaths_def pathstart_def pathfinish_def split_ifs field_split_simps) | 
| 69508 | 683 | then have n1: "pathstart g1 \<notin> path_image g2" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 684 | unfolding pathstart_def path_image_def | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 685 | using atLeastAtMost_iff by blast | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 686 | show ?rhs using \<open>?lhs\<close> | 
| 72256 | 687 | using \<open>simple_path (g1 +++ g2)\<close> assms n1 simple_path_joinE by auto | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 688 | next | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 689 | assume ?rhs then show ?lhs | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 690 | using assms | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 691 | by (fastforce simp: pathfinish_def pathstart_def intro!: arc_join) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 692 | qed | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 693 | |
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 694 | lemma arc_join_eq_alt: | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 695 | "pathfinish g1 = pathstart g2 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 696 | \<Longrightarrow> (arc(g1 +++ g2) \<longleftrightarrow> | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 697 | arc g1 \<and> arc g2 \<and> | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 698 |              path_image g1 \<inter> path_image g2 = {pathstart g2})"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 699 | using pathfinish_in_path_image by (fastforce simp: arc_join_eq) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 700 | |
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 701 | |
| 70136 | 702 | subsection\<^marker>\<open>tag unimportant\<close>\<open>The joining of paths is associative\<close> | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 703 | |
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 704 | lemma path_assoc: | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 705 | "\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart r\<rbrakk> | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 706 | \<Longrightarrow> path(p +++ (q +++ r)) \<longleftrightarrow> path((p +++ q) +++ r)" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 707 | by simp | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 708 | |
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 709 | lemma simple_path_assoc: | 
| 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 710 | assumes "pathfinish p = pathstart q" "pathfinish q = pathstart r" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 711 | shows "simple_path (p +++ (q +++ r)) \<longleftrightarrow> simple_path ((p +++ q) +++ r)" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 712 | proof (cases "pathstart p = pathfinish r") | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 713 | case True show ?thesis | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 714 | proof | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 715 | assume "simple_path (p +++ q +++ r)" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 716 | with assms True show "simple_path ((p +++ q) +++ r)" | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 717 | by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 718 | dest: arc_distinct_ends [of r]) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 719 | next | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 720 | assume 0: "simple_path ((p +++ q) +++ r)" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 721 | with assms True have q: "pathfinish r \<notin> path_image q" | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 722 | using arc_distinct_ends | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 723 | by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 724 | have "pathstart r \<notin> path_image p" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 725 | using assms | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 726 | by (metis 0 IntI arc_distinct_ends arc_join_eq_alt empty_iff insert_iff | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 727 | pathfinish_in_path_image pathfinish_join simple_path_joinE) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 728 | with assms 0 q True show "simple_path (p +++ q +++ r)" | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 729 | by (auto simp: simple_path_join_loop_eq arc_join_eq path_image_join | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 730 | dest!: subsetD [OF _ IntI]) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 731 | qed | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 732 | next | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 733 | case False | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 734 |   { fix x :: 'a
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 735 |     assume a: "path_image p \<inter> path_image q \<subseteq> {pathstart q}"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 736 |               "(path_image p \<union> path_image q) \<inter> path_image r \<subseteq> {pathstart r}"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 737 | "x \<in> path_image p" "x \<in> path_image r" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 738 | have "pathstart r \<in> path_image q" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 739 | by (metis assms(2) pathfinish_in_path_image) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 740 | with a have "x = pathstart q" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 741 | by blast | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 742 | } | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 743 | with False assms show ?thesis | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 744 | by (auto simp: simple_path_eq_arc simple_path_join_loop_eq arc_join_eq path_image_join) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 745 | qed | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 746 | |
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 747 | lemma arc_assoc: | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 748 | "\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart r\<rbrakk> | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 749 | \<Longrightarrow> arc(p +++ (q +++ r)) \<longleftrightarrow> arc((p +++ q) +++ r)" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 750 | by (simp add: arc_simple_path simple_path_assoc) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 751 | |
| 70136 | 752 | subsubsection\<^marker>\<open>tag unimportant\<close>\<open>Symmetry and loops\<close> | 
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 753 | |
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 754 | lemma path_sym: | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 755 | "\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart p\<rbrakk> \<Longrightarrow> path(p +++ q) \<longleftrightarrow> path(q +++ p)" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 756 | by auto | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 757 | |
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 758 | lemma simple_path_sym: | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 759 | "\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart p\<rbrakk> | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 760 | \<Longrightarrow> simple_path(p +++ q) \<longleftrightarrow> simple_path(q +++ p)" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 761 | by (metis (full_types) inf_commute insert_commute simple_path_joinE simple_path_join_loop) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 762 | |
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 763 | lemma path_image_sym: | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 764 | "\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart p\<rbrakk> | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 765 | \<Longrightarrow> path_image(p +++ q) = path_image(q +++ p)" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 766 | by (simp add: path_image_join sup_commute) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 767 | |
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 768 | |
| 69518 | 769 | subsection\<open>Subpath\<close> | 
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 770 | |
| 70136 | 771 | definition\<^marker>\<open>tag important\<close> subpath :: "real \<Rightarrow> real \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> real \<Rightarrow> 'a::real_normed_vector" | 
| 60303 | 772 | where "subpath a b g \<equiv> \<lambda>x. g((b - a) * x + a)" | 
| 773 | ||
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61738diff
changeset | 774 | lemma path_image_subpath_gen: | 
| 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61738diff
changeset | 775 | fixes g :: "_ \<Rightarrow> 'a::real_normed_vector" | 
| 60303 | 776 | shows "path_image(subpath u v g) = g ` (closed_segment u v)" | 
| 69661 | 777 | by (auto simp add: closed_segment_real_eq path_image_def subpath_def) | 
| 60303 | 778 | |
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61738diff
changeset | 779 | lemma path_image_subpath: | 
| 60303 | 780 | fixes g :: "real \<Rightarrow> 'a::real_normed_vector" | 
| 781 |   shows "path_image(subpath u v g) = (if u \<le> v then g ` {u..v} else g ` {v..u})"
 | |
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61738diff
changeset | 782 | by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) | 
| 60303 | 783 | |
| 65038 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 784 | lemma path_image_subpath_commute: | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 785 | fixes g :: "real \<Rightarrow> 'a::real_normed_vector" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 786 | shows "path_image(subpath u v g) = path_image(subpath v u g)" | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 787 | by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) | 
| 
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 788 | |
| 60303 | 789 | lemma path_subpath [simp]: | 
| 790 | fixes g :: "real \<Rightarrow> 'a::real_normed_vector" | |
| 791 |   assumes "path g" "u \<in> {0..1}" "v \<in> {0..1}"
 | |
| 792 | shows "path(subpath u v g)" | |
| 793 | proof - | |
| 68096 | 794 |   have "continuous_on {0..1} (g \<circ> (\<lambda>x. ((v-u) * x+ u)))"
 | 
| 60303 | 795 | using assms | 
| 72256 | 796 | apply (intro continuous_intros; simp add: image_affinity_atLeastAtMost [where c=u]) | 
| 60303 | 797 | apply (auto simp: path_def continuous_on_subset) | 
| 798 | done | |
| 799 | then show ?thesis | |
| 800 | by (simp add: path_def subpath_def) | |
| 49653 | 801 | qed | 
| 36583 | 802 | |
| 60303 | 803 | lemma pathstart_subpath [simp]: "pathstart(subpath u v g) = g(u)" | 
| 804 | by (simp add: pathstart_def subpath_def) | |
| 805 | ||
| 806 | lemma pathfinish_subpath [simp]: "pathfinish(subpath u v g) = g(v)" | |
| 807 | by (simp add: pathfinish_def subpath_def) | |
| 808 | ||
| 809 | lemma subpath_trivial [simp]: "subpath 0 1 g = g" | |
| 810 | by (simp add: subpath_def) | |
| 811 | ||
| 812 | lemma subpath_reversepath: "subpath 1 0 g = reversepath g" | |
| 813 | by (simp add: reversepath_def subpath_def) | |
| 814 | ||
| 815 | lemma reversepath_subpath: "reversepath(subpath u v g) = subpath v u g" | |
| 816 | by (simp add: reversepath_def subpath_def algebra_simps) | |
| 817 | ||
| 68096 | 818 | lemma subpath_translation: "subpath u v ((\<lambda>x. a + x) \<circ> g) = (\<lambda>x. a + x) \<circ> subpath u v g" | 
| 60303 | 819 | by (rule ext) (simp add: subpath_def) | 
| 820 | ||
| 70971 | 821 | lemma subpath_image: "subpath u v (f \<circ> g) = f \<circ> subpath u v g" | 
| 60303 | 822 | by (rule ext) (simp add: subpath_def) | 
| 823 | ||
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 824 | lemma affine_ineq: | 
| 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 825 | fixes x :: "'a::linordered_idom" | 
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61738diff
changeset | 826 | assumes "x \<le> 1" "v \<le> u" | 
| 60303 | 827 | shows "v + x * u \<le> u + x * v" | 
| 828 | proof - | |
| 829 | have "(1-x)*(u-v) \<ge> 0" | |
| 830 | using assms by auto | |
| 831 | then show ?thesis | |
| 832 | by (simp add: algebra_simps) | |
| 49653 | 833 | qed | 
| 36583 | 834 | |
| 61711 
21d7910d6816
Theory of homotopic paths (from HOL Light), plus comments and minor refinements
 paulson <lp15@cam.ac.uk> parents: 
61699diff
changeset | 835 | lemma sum_le_prod1: | 
| 
21d7910d6816
Theory of homotopic paths (from HOL Light), plus comments and minor refinements
 paulson <lp15@cam.ac.uk> parents: 
61699diff
changeset | 836 | fixes a::real shows "\<lbrakk>a \<le> 1; b \<le> 1\<rbrakk> \<Longrightarrow> a + b \<le> 1 + a * b" | 
| 71172 | 837 | by (metis add.commute affine_ineq mult.right_neutral) | 
| 61711 
21d7910d6816
Theory of homotopic paths (from HOL Light), plus comments and minor refinements
 paulson <lp15@cam.ac.uk> parents: 
61699diff
changeset | 838 | |
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 839 | lemma simple_path_subpath_eq: | 
| 60303 | 840 | "simple_path(subpath u v g) \<longleftrightarrow> | 
| 841 | path(subpath u v g) \<and> u\<noteq>v \<and> | |
| 842 | (\<forall>x y. x \<in> closed_segment u v \<and> y \<in> closed_segment u v \<and> g x = g y | |
| 843 | \<longrightarrow> x = y \<or> x = u \<and> y = v \<or> x = v \<and> y = u)" | |
| 844 | (is "?lhs = ?rhs") | |
| 72256 | 845 | proof | 
| 60303 | 846 | assume ?lhs | 
| 847 | then have p: "path (\<lambda>x. g ((v - u) * x + u))" | |
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 848 |         and sim: "(\<And>x y. \<lbrakk>x\<in>{0..1}; y\<in>{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\<rbrakk>
 | 
| 60303 | 849 | \<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0)" | 
| 850 | by (auto simp: simple_path_def subpath_def) | |
| 851 |   { fix x y
 | |
| 852 | assume "x \<in> closed_segment u v" "y \<in> closed_segment u v" "g x = g y" | |
| 853 | then have "x = y \<or> x = u \<and> y = v \<or> x = v \<and> y = u" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 854 | using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 855 | by (auto split: if_split_asm simp add: closed_segment_real_eq image_affinity_atLeastAtMost) | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 856 | (simp_all add: field_split_simps) | 
| 60303 | 857 | } moreover | 
| 858 | have "path(subpath u v g) \<and> u\<noteq>v" | |
| 859 | using sim [of "1/3" "2/3"] p | |
| 860 | by (auto simp: subpath_def) | |
| 861 | ultimately show ?rhs | |
| 862 | by metis | |
| 863 | next | |
| 864 | assume ?rhs | |
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 865 | then | 
| 60303 | 866 | have d1: "\<And>x y. \<lbrakk>g x = g y; u \<le> x; x \<le> v; u \<le> y; y \<le> v\<rbrakk> \<Longrightarrow> x = y \<or> x = u \<and> y = v \<or> x = v \<and> y = u" | 
| 867 | and d2: "\<And>x y. \<lbrakk>g x = g y; v \<le> x; x \<le> u; v \<le> y; y \<le> u\<rbrakk> \<Longrightarrow> x = y \<or> x = u \<and> y = v \<or> x = v \<and> y = u" | |
| 868 | and ne: "u < v \<or> v < u" | |
| 869 | and psp: "path (subpath u v g)" | |
| 870 | by (auto simp: closed_segment_real_eq image_affinity_atLeastAtMost) | |
| 871 | have [simp]: "\<And>x. u + x * v = v + x * u \<longleftrightarrow> u=v \<or> x=1" | |
| 872 | by algebra | |
| 873 | show ?lhs using psp ne | |
| 874 | unfolding simple_path_def subpath_def | |
| 875 | by (fastforce simp add: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) | |
| 876 | qed | |
| 877 | ||
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 878 | lemma arc_subpath_eq: | 
| 60303 | 879 | "arc(subpath u v g) \<longleftrightarrow> path(subpath u v g) \<and> u\<noteq>v \<and> inj_on g (closed_segment u v)" | 
| 880 | (is "?lhs = ?rhs") | |
| 72256 | 881 | proof | 
| 60303 | 882 | assume ?lhs | 
| 883 | then have p: "path (\<lambda>x. g ((v - u) * x + u))" | |
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 884 |         and sim: "(\<And>x y. \<lbrakk>x\<in>{0..1}; y\<in>{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\<rbrakk>
 | 
| 60303 | 885 | \<Longrightarrow> x = y)" | 
| 886 | by (auto simp: arc_def inj_on_def subpath_def) | |
| 887 |   { fix x y
 | |
| 888 | assume "x \<in> closed_segment u v" "y \<in> closed_segment u v" "g x = g y" | |
| 889 | then have "x = y" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 890 | using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 891 | by (cases "v = u") | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 892 | (simp_all split: if_split_asm add: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost, | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 893 | simp add: field_simps) | 
| 60303 | 894 | } moreover | 
| 895 | have "path(subpath u v g) \<and> u\<noteq>v" | |
| 896 | using sim [of "1/3" "2/3"] p | |
| 897 | by (auto simp: subpath_def) | |
| 898 | ultimately show ?rhs | |
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 899 | unfolding inj_on_def | 
| 60303 | 900 | by metis | 
| 901 | next | |
| 902 | assume ?rhs | |
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 903 | then | 
| 60303 | 904 | have d1: "\<And>x y. \<lbrakk>g x = g y; u \<le> x; x \<le> v; u \<le> y; y \<le> v\<rbrakk> \<Longrightarrow> x = y" | 
| 905 | and d2: "\<And>x y. \<lbrakk>g x = g y; v \<le> x; x \<le> u; v \<le> y; y \<le> u\<rbrakk> \<Longrightarrow> x = y" | |
| 906 | and ne: "u < v \<or> v < u" | |
| 907 | and psp: "path (subpath u v g)" | |
| 908 | by (auto simp: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost) | |
| 909 | show ?lhs using psp ne | |
| 910 | unfolding arc_def subpath_def inj_on_def | |
| 911 | by (auto simp: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) | |
| 912 | qed | |
| 913 | ||
| 914 | ||
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 915 | lemma simple_path_subpath: | 
| 60303 | 916 |   assumes "simple_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<noteq> v"
 | 
| 917 | shows "simple_path(subpath u v g)" | |
| 918 | using assms | |
| 919 | apply (simp add: simple_path_subpath_eq simple_path_imp_path) | |
| 920 | apply (simp add: simple_path_def closed_segment_real_eq image_affinity_atLeastAtMost, fastforce) | |
| 921 | done | |
| 922 | ||
| 923 | lemma arc_simple_path_subpath: | |
| 924 |     "\<lbrakk>simple_path g; u \<in> {0..1}; v \<in> {0..1}; g u \<noteq> g v\<rbrakk> \<Longrightarrow> arc(subpath u v g)"
 | |
| 925 | by (force intro: simple_path_subpath simple_path_imp_arc) | |
| 926 | ||
| 927 | lemma arc_subpath_arc: | |
| 928 |     "\<lbrakk>arc g; u \<in> {0..1}; v \<in> {0..1}; u \<noteq> v\<rbrakk> \<Longrightarrow> arc(subpath u v g)"
 | |
| 929 | by (meson arc_def arc_imp_simple_path arc_simple_path_subpath inj_onD) | |
| 930 | ||
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 931 | lemma arc_simple_path_subpath_interior: | 
| 60303 | 932 |     "\<lbrakk>simple_path g; u \<in> {0..1}; v \<in> {0..1}; u \<noteq> v; \<bar>u-v\<bar> < 1\<rbrakk> \<Longrightarrow> arc(subpath u v g)"
 | 
| 72256 | 933 | by (force simp: simple_path_def intro: arc_simple_path_subpath) | 
| 60303 | 934 | |
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 935 | lemma path_image_subpath_subset: | 
| 68532 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68310diff
changeset | 936 |     "\<lbrakk>u \<in> {0..1}; v \<in> {0..1}\<rbrakk> \<Longrightarrow> path_image(subpath u v g) \<subseteq> path_image g"
 | 
| 72256 | 937 | by (metis atLeastAtMost_iff atLeastatMost_subset_iff path_image_def path_image_subpath subset_image_iff) | 
| 60303 | 938 | |
| 939 | lemma join_subpaths_middle: "subpath (0) ((1 / 2)) p +++ subpath ((1 / 2)) 1 p = p" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 940 | by (rule ext) (simp add: joinpaths_def subpath_def field_split_simps) | 
| 53640 | 941 | |
| 69514 | 942 | |
| 70136 | 943 | subsection\<^marker>\<open>tag unimportant\<close>\<open>There is a subpath to the frontier\<close> | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 944 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 945 | lemma subpath_to_frontier_explicit: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 946 | fixes S :: "'a::metric_space set" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 947 | assumes g: "path g" and "pathfinish g \<notin> S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 948 | obtains u where "0 \<le> u" "u \<le> 1" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 949 | "\<And>x. 0 \<le> x \<and> x < u \<Longrightarrow> g x \<in> interior S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 950 | "(g u \<notin> interior S)" "(u = 0 \<or> g u \<in> closure S)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 951 | proof - | 
| 72256 | 952 |   have gcon: "continuous_on {0..1} g"     
 | 
| 953 | using g by (simp add: path_def) | |
| 954 |   moreover have "bounded ({u. g u \<in> closure (- S)} \<inter> {0..1})"
 | |
| 955 | using compact_eq_bounded_closed by fastforce | |
| 956 |   ultimately have com: "compact ({0..1} \<inter> {u. g u \<in> closure (- S)})"
 | |
| 957 | using closed_vimage_Int | |
| 958 | by (metis (full_types) Int_commute closed_atLeastAtMost closed_closure compact_eq_bounded_closed vimage_def) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 959 |   have "1 \<in> {u. g u \<in> closure (- S)}"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 960 | using assms by (simp add: pathfinish_def closure_def) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 961 |   then have dis: "{0..1} \<inter> {u. g u \<in> closure (- S)} \<noteq> {}"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 962 | using atLeastAtMost_iff zero_le_one by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 963 | then obtain u where "0 \<le> u" "u \<le> 1" and gu: "g u \<in> closure (- S)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 964 | and umin: "\<And>t. \<lbrakk>0 \<le> t; t \<le> 1; g t \<in> closure (- S)\<rbrakk> \<Longrightarrow> u \<le> t" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 965 | using compact_attains_inf [OF com dis] by fastforce | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 966 | then have umin': "\<And>t. \<lbrakk>0 \<le> t; t \<le> 1; t < u\<rbrakk> \<Longrightarrow> g t \<in> S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 967 | using closure_def by fastforce | 
| 72256 | 968 | have \<section>: "g u \<in> closure S" if "u \<noteq> 0" | 
| 969 | proof - | |
| 970 | have "u > 0" using that \<open>0 \<le> u\<close> by auto | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 971 |     { fix e::real assume "e > 0"
 | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62381diff
changeset | 972 |       obtain d where "d>0" and d: "\<And>x'. \<lbrakk>x' \<in> {0..1}; dist x' u \<le> d\<rbrakk> \<Longrightarrow> dist (g x') (g u) < e"
 | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62381diff
changeset | 973 | using continuous_onE [OF gcon _ \<open>e > 0\<close>] \<open>0 \<le> _\<close> \<open>_ \<le> 1\<close> atLeastAtMost_iff by auto | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62381diff
changeset | 974 | have *: "dist (max 0 (u - d / 2)) u \<le> d" | 
| 61808 | 975 | using \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> \<open>d > 0\<close> by (simp add: dist_real_def) | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 976 | have "\<exists>y\<in>S. dist y (g u) < e" | 
| 61808 | 977 | using \<open>0 < u\<close> \<open>u \<le> 1\<close> \<open>d > 0\<close> | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 978 | by (force intro: d [OF _ *] umin') | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 979 | } | 
| 72256 | 980 | then show ?thesis | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 981 | by (simp add: frontier_def closure_approachable) | 
| 72256 | 982 | qed | 
| 983 | show ?thesis | |
| 984 | proof | |
| 985 | show "\<And>x. 0 \<le> x \<and> x < u \<Longrightarrow> g x \<in> interior S" | |
| 986 | using \<open>u \<le> 1\<close> interior_closure umin by fastforce | |
| 987 | show "g u \<notin> interior S" | |
| 988 | by (simp add: gu interior_closure) | |
| 989 | qed (use \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> \<section> in auto) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 990 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 991 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 992 | lemma subpath_to_frontier_strong: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 993 | assumes g: "path g" and "pathfinish g \<notin> S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 994 | obtains u where "0 \<le> u" "u \<le> 1" "g u \<notin> interior S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 995 | "u = 0 \<or> (\<forall>x. 0 \<le> x \<and> x < 1 \<longrightarrow> subpath 0 u g x \<in> interior S) \<and> g u \<in> closure S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 996 | proof - | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 997 | obtain u where "0 \<le> u" "u \<le> 1" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 998 | and gxin: "\<And>x. 0 \<le> x \<and> x < u \<Longrightarrow> g x \<in> interior S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 999 | and gunot: "(g u \<notin> interior S)" and u0: "(u = 0 \<or> g u \<in> closure S)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1000 | using subpath_to_frontier_explicit [OF assms] by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1001 | show ?thesis | 
| 72256 | 1002 | proof | 
| 1003 | show "g u \<notin> interior S" | |
| 1004 | using gunot by blast | |
| 1005 | qed (use \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> u0 in \<open>(force simp: subpath_def gxin)+\<close>) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1006 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1007 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1008 | lemma subpath_to_frontier: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1009 | assumes g: "path g" and g0: "pathstart g \<in> closure S" and g1: "pathfinish g \<notin> S" | 
| 72256 | 1010 |     obtains u where "0 \<le> u" "u \<le> 1" "g u \<in> frontier S" "path_image(subpath 0 u g) - {g u} \<subseteq> interior S"
 | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1011 | proof - | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1012 | obtain u where "0 \<le> u" "u \<le> 1" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1013 | and notin: "g u \<notin> interior S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1014 | and disj: "u = 0 \<or> | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1015 | (\<forall>x. 0 \<le> x \<and> x < 1 \<longrightarrow> subpath 0 u g x \<in> interior S) \<and> g u \<in> closure S" | 
| 72256 | 1016 | (is "_ \<or> ?P") | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1017 | using subpath_to_frontier_strong [OF g g1] by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1018 | show ?thesis | 
| 72256 | 1019 | proof | 
| 1020 | show "g u \<in> frontier S" | |
| 1021 | by (metis DiffI disj frontier_def g0 notin pathstart_def) | |
| 1022 |     show "path_image (subpath 0 u g) - {g u} \<subseteq> interior S"
 | |
| 1023 | using disj | |
| 1024 | proof | |
| 1025 | assume "u = 0" | |
| 1026 | then show ?thesis | |
| 1027 | by (simp add: path_image_subpath) | |
| 1028 | next | |
| 1029 | assume P: ?P | |
| 1030 | show ?thesis | |
| 1031 | proof (clarsimp simp add: path_image_subpath_gen) | |
| 1032 | fix y | |
| 1033 | assume y: "y \<in> closed_segment 0 u" "g y \<notin> interior S" | |
| 1034 | with \<open>0 \<le> u\<close> have "0 \<le> y" "y \<le> u" | |
| 1035 | by (auto simp: closed_segment_eq_real_ivl split: if_split_asm) | |
| 1036 | then have "y=u \<or> subpath 0 u g (y/u) \<in> interior S" | |
| 1037 | using P less_eq_real_def by force | |
| 1038 | then show "g y = g u" | |
| 1039 | using y by (auto simp: subpath_def split: if_split_asm) | |
| 1040 | qed | |
| 1041 | qed | |
| 1042 | qed (use \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> in auto) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1043 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1044 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1045 | lemma exists_path_subpath_to_frontier: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1046 | fixes S :: "'a::real_normed_vector set" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1047 | assumes "path g" "pathstart g \<in> closure S" "pathfinish g \<notin> S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1048 | obtains h where "path h" "pathstart h = pathstart g" "path_image h \<subseteq> path_image g" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1049 |                     "path_image h - {pathfinish h} \<subseteq> interior S"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1050 | "pathfinish h \<in> frontier S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1051 | proof - | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1052 |   obtain u where u: "0 \<le> u" "u \<le> 1" "g u \<in> frontier S" "(path_image(subpath 0 u g) - {g u}) \<subseteq> interior S"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1053 | using subpath_to_frontier [OF assms] by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1054 | show ?thesis | 
| 72256 | 1055 | proof | 
| 1056 | show "path_image (subpath 0 u g) \<subseteq> path_image g" | |
| 1057 | by (simp add: path_image_subpath_subset u) | |
| 1058 | show "pathstart (subpath 0 u g) = pathstart g" | |
| 1059 | by (metis pathstart_def pathstart_subpath) | |
| 1060 | qed (use assms u in \<open>auto simp: path_image_subpath\<close>) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1061 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1062 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1063 | lemma exists_path_subpath_to_frontier_closed: | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1064 | fixes S :: "'a::real_normed_vector set" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1065 | assumes S: "closed S" and g: "path g" and g0: "pathstart g \<in> S" and g1: "pathfinish g \<notin> S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1066 | obtains h where "path h" "pathstart h = pathstart g" "path_image h \<subseteq> path_image g \<inter> S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1067 | "pathfinish h \<in> frontier S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1068 | proof - | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1069 | obtain h where h: "path h" "pathstart h = pathstart g" "path_image h \<subseteq> path_image g" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1070 |                     "path_image h - {pathfinish h} \<subseteq> interior S"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1071 | "pathfinish h \<in> frontier S" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1072 | using exists_path_subpath_to_frontier [OF g _ g1] closure_closed [OF S] g0 by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1073 | show ?thesis | 
| 72256 | 1074 | proof | 
| 1075 | show "path_image h \<subseteq> path_image g \<inter> S" | |
| 1076 | using assms h interior_subset [of S] by (auto simp: frontier_def) | |
| 1077 | qed (use h in auto) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1078 | qed | 
| 49653 | 1079 | |
| 69514 | 1080 | |
| 1081 | subsection \<open>Shift Path to Start at Some Given Point\<close> | |
| 36583 | 1082 | |
| 70136 | 1083 | definition\<^marker>\<open>tag important\<close> shiftpath :: "real \<Rightarrow> (real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a" | 
| 53640 | 1084 | where "shiftpath a f = (\<lambda>x. if (a + x) \<le> 1 then f (a + x) else f (a + x - 1))" | 
| 36583 | 1085 | |
| 71189 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1086 | lemma shiftpath_alt_def: "shiftpath a f = (\<lambda>x. if x \<le> 1-a then f (a + x) else f (a + x - 1))" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1087 | by (auto simp: shiftpath_def) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1088 | |
| 53640 | 1089 | lemma pathstart_shiftpath: "a \<le> 1 \<Longrightarrow> pathstart (shiftpath a g) = g a" | 
| 36583 | 1090 | unfolding pathstart_def shiftpath_def by auto | 
| 1091 | ||
| 49653 | 1092 | lemma pathfinish_shiftpath: | 
| 53640 | 1093 | assumes "0 \<le> a" | 
| 1094 | and "pathfinish g = pathstart g" | |
| 1095 | shows "pathfinish (shiftpath a g) = g a" | |
| 1096 | using assms | |
| 1097 | unfolding pathstart_def pathfinish_def shiftpath_def | |
| 36583 | 1098 | by auto | 
| 1099 | ||
| 1100 | lemma endpoints_shiftpath: | |
| 53640 | 1101 | assumes "pathfinish g = pathstart g" | 
| 1102 |     and "a \<in> {0 .. 1}"
 | |
| 1103 | shows "pathfinish (shiftpath a g) = g a" | |
| 1104 | and "pathstart (shiftpath a g) = g a" | |
| 1105 | using assms | |
| 1106 | by (auto intro!: pathfinish_shiftpath pathstart_shiftpath) | |
| 36583 | 1107 | |
| 1108 | lemma closed_shiftpath: | |
| 53640 | 1109 | assumes "pathfinish g = pathstart g" | 
| 1110 |     and "a \<in> {0..1}"
 | |
| 1111 | shows "pathfinish (shiftpath a g) = pathstart (shiftpath a g)" | |
| 1112 | using endpoints_shiftpath[OF assms] | |
| 1113 | by auto | |
| 36583 | 1114 | |
| 1115 | lemma path_shiftpath: | |
| 53640 | 1116 | assumes "path g" | 
| 1117 | and "pathfinish g = pathstart g" | |
| 1118 |     and "a \<in> {0..1}"
 | |
| 1119 | shows "path (shiftpath a g)" | |
| 49653 | 1120 | proof - | 
| 53640 | 1121 |   have *: "{0 .. 1} = {0 .. 1-a} \<union> {1-a .. 1}"
 | 
| 1122 | using assms(3) by auto | |
| 49653 | 1123 | have **: "\<And>x. x + a = 1 \<Longrightarrow> g (x + a - 1) = g (x + a)" | 
| 53640 | 1124 | using assms(2)[unfolded pathfinish_def pathstart_def] | 
| 1125 | by auto | |
| 49653 | 1126 | show ?thesis | 
| 1127 | unfolding path_def shiftpath_def * | |
| 68096 | 1128 | proof (rule continuous_on_closed_Un) | 
| 1129 |     have contg: "continuous_on {0..1} g"
 | |
| 1130 | using \<open>path g\<close> path_def by blast | |
| 1131 |     show "continuous_on {0..1-a} (\<lambda>x. if a + x \<le> 1 then g (a + x) else g (a + x - 1))"
 | |
| 1132 | proof (rule continuous_on_eq) | |
| 1133 |       show "continuous_on {0..1-a} (g \<circ> (+) a)"
 | |
| 1134 |         by (intro continuous_intros continuous_on_subset [OF contg]) (use \<open>a \<in> {0..1}\<close> in auto)
 | |
| 1135 | qed auto | |
| 1136 |     show "continuous_on {1-a..1} (\<lambda>x. if a + x \<le> 1 then g (a + x) else g (a + x - 1))"
 | |
| 1137 | proof (rule continuous_on_eq) | |
| 1138 |       show "continuous_on {1-a..1} (g \<circ> (+) (a - 1))"
 | |
| 1139 |         by (intro continuous_intros continuous_on_subset [OF contg]) (use \<open>a \<in> {0..1}\<close> in auto)
 | |
| 1140 | qed (auto simp: "**" add.commute add_diff_eq) | |
| 1141 | qed auto | |
| 49653 | 1142 | qed | 
| 36583 | 1143 | |
| 49653 | 1144 | lemma shiftpath_shiftpath: | 
| 53640 | 1145 | assumes "pathfinish g = pathstart g" | 
| 1146 |     and "a \<in> {0..1}"
 | |
| 1147 |     and "x \<in> {0..1}"
 | |
| 36583 | 1148 | shows "shiftpath (1 - a) (shiftpath a g) x = g x" | 
| 53640 | 1149 | using assms | 
| 1150 | unfolding pathfinish_def pathstart_def shiftpath_def | |
| 1151 | by auto | |
| 36583 | 1152 | |
| 1153 | lemma path_image_shiftpath: | |
| 68096 | 1154 |   assumes a: "a \<in> {0..1}"
 | 
| 53640 | 1155 | and "pathfinish g = pathstart g" | 
| 1156 | shows "path_image (shiftpath a g) = path_image g" | |
| 49653 | 1157 | proof - | 
| 1158 |   { fix x
 | |
| 68096 | 1159 |     assume g: "g 1 = g 0" "x \<in> {0..1::real}" and gne: "\<And>y. y\<in>{0..1} \<inter> {x. \<not> a + x \<le> 1} \<Longrightarrow> g x \<noteq> g (a + y - 1)"
 | 
| 49654 | 1160 |     then have "\<exists>y\<in>{0..1} \<inter> {x. a + x \<le> 1}. g x = g (a + y)"
 | 
| 49653 | 1161 | proof (cases "a \<le> x") | 
| 1162 | case False | |
| 49654 | 1163 | then show ?thesis | 
| 49653 | 1164 | apply (rule_tac x="1 + x - a" in bexI) | 
| 72256 | 1165 | using g gne[of "1 + x - a"] a by (force simp: field_simps)+ | 
| 49653 | 1166 | next | 
| 1167 | case True | |
| 53640 | 1168 | then show ?thesis | 
| 68096 | 1169 | using g a by (rule_tac x="x - a" in bexI) (auto simp: field_simps) | 
| 49653 | 1170 | qed | 
| 1171 | } | |
| 49654 | 1172 | then show ?thesis | 
| 53640 | 1173 | using assms | 
| 1174 | unfolding shiftpath_def path_image_def pathfinish_def pathstart_def | |
| 68096 | 1175 | by (auto simp: image_iff) | 
| 49653 | 1176 | qed | 
| 1177 | ||
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1178 | lemma simple_path_shiftpath: | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1179 | assumes "simple_path g" "pathfinish g = pathstart g" and a: "0 \<le> a" "a \<le> 1" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1180 | shows "simple_path (shiftpath a g)" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1181 | unfolding simple_path_def | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1182 | proof (intro conjI impI ballI) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1183 | show "path (shiftpath a g)" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1184 | by (simp add: assms path_shiftpath simple_path_imp_path) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1185 |   have *: "\<And>x y. \<lbrakk>g x = g y; x \<in> {0..1}; y \<in> {0..1}\<rbrakk> \<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1186 | using assms by (simp add: simple_path_def) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1187 | show "x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1188 |     if "x \<in> {0..1}" "y \<in> {0..1}" "shiftpath a g x = shiftpath a g y" for x y
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1189 | using that a unfolding shiftpath_def | 
| 68096 | 1190 | by (force split: if_split_asm dest!: *) | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1191 | qed | 
| 36583 | 1192 | |
| 69514 | 1193 | |
| 1194 | subsection \<open>Straight-Line Paths\<close> | |
| 36583 | 1195 | |
| 70136 | 1196 | definition\<^marker>\<open>tag important\<close> linepath :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> real \<Rightarrow> 'a" | 
| 49653 | 1197 | where "linepath a b = (\<lambda>x. (1 - x) *\<^sub>R a + x *\<^sub>R b)" | 
| 36583 | 1198 | |
| 53640 | 1199 | lemma pathstart_linepath[simp]: "pathstart (linepath a b) = a" | 
| 1200 | unfolding pathstart_def linepath_def | |
| 1201 | by auto | |
| 36583 | 1202 | |
| 53640 | 1203 | lemma pathfinish_linepath[simp]: "pathfinish (linepath a b) = b" | 
| 1204 | unfolding pathfinish_def linepath_def | |
| 1205 | by auto | |
| 36583 | 1206 | |
| 68721 | 1207 | lemma linepath_inner: "linepath a b x \<bullet> v = linepath (a \<bullet> v) (b \<bullet> v) x" | 
| 1208 | by (simp add: linepath_def algebra_simps) | |
| 1209 | ||
| 1210 | lemma Re_linepath': "Re (linepath a b x) = linepath (Re a) (Re b) x" | |
| 1211 | by (simp add: linepath_def) | |
| 1212 | ||
| 1213 | lemma Im_linepath': "Im (linepath a b x) = linepath (Im a) (Im b) x" | |
| 1214 | by (simp add: linepath_def) | |
| 1215 | ||
| 1216 | lemma linepath_0': "linepath a b 0 = a" | |
| 1217 | by (simp add: linepath_def) | |
| 1218 | ||
| 1219 | lemma linepath_1': "linepath a b 1 = b" | |
| 1220 | by (simp add: linepath_def) | |
| 1221 | ||
| 36583 | 1222 | lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)" | 
| 53640 | 1223 | unfolding linepath_def | 
| 1224 | by (intro continuous_intros) | |
| 36583 | 1225 | |
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61738diff
changeset | 1226 | lemma continuous_on_linepath [intro,continuous_intros]: "continuous_on s (linepath a b)" | 
| 53640 | 1227 | using continuous_linepath_at | 
| 1228 | by (auto intro!: continuous_at_imp_continuous_on) | |
| 36583 | 1229 | |
| 62618 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1230 | lemma path_linepath[iff]: "path (linepath a b)" | 
| 53640 | 1231 | unfolding path_def | 
| 1232 | by (rule continuous_on_linepath) | |
| 36583 | 1233 | |
| 53640 | 1234 | lemma path_image_linepath[simp]: "path_image (linepath a b) = closed_segment a b" | 
| 49653 | 1235 | unfolding path_image_def segment linepath_def | 
| 60303 | 1236 | by auto | 
| 49653 | 1237 | |
| 53640 | 1238 | lemma reversepath_linepath[simp]: "reversepath (linepath a b) = linepath b a" | 
| 49653 | 1239 | unfolding reversepath_def linepath_def | 
| 36583 | 1240 | by auto | 
| 1241 | ||
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61738diff
changeset | 1242 | lemma linepath_0 [simp]: "linepath 0 b x = x *\<^sub>R b" | 
| 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61738diff
changeset | 1243 | by (simp add: linepath_def) | 
| 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61738diff
changeset | 1244 | |
| 68721 | 1245 | lemma linepath_cnj: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" | 
| 1246 | by (simp add: linepath_def) | |
| 1247 | ||
| 60303 | 1248 | lemma arc_linepath: | 
| 62618 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1249 | assumes "a \<noteq> b" shows [simp]: "arc (linepath a b)" | 
| 36583 | 1250 | proof - | 
| 53640 | 1251 |   {
 | 
| 1252 | fix x y :: "real" | |
| 36583 | 1253 | assume "x *\<^sub>R b + y *\<^sub>R a = x *\<^sub>R a + y *\<^sub>R b" | 
| 53640 | 1254 | then have "(x - y) *\<^sub>R a = (x - y) *\<^sub>R b" | 
| 1255 | by (simp add: algebra_simps) | |
| 1256 | with assms have "x = y" | |
| 1257 | by simp | |
| 1258 | } | |
| 49654 | 1259 | then show ?thesis | 
| 60809 
457abb82fb9e
the Cauchy integral theorem and related material
 paulson <lp15@cam.ac.uk> parents: 
60420diff
changeset | 1260 | unfolding arc_def inj_on_def | 
| 68096 | 1261 | by (fastforce simp: algebra_simps linepath_def) | 
| 49653 | 1262 | qed | 
| 36583 | 1263 | |
| 53640 | 1264 | lemma simple_path_linepath[intro]: "a \<noteq> b \<Longrightarrow> simple_path (linepath a b)" | 
| 68096 | 1265 | by (simp add: arc_imp_simple_path) | 
| 49653 | 1266 | |
| 61711 
21d7910d6816
Theory of homotopic paths (from HOL Light), plus comments and minor refinements
 paulson <lp15@cam.ac.uk> parents: 
61699diff
changeset | 1267 | lemma linepath_trivial [simp]: "linepath a a x = a" | 
| 
21d7910d6816
Theory of homotopic paths (from HOL Light), plus comments and minor refinements
 paulson <lp15@cam.ac.uk> parents: 
61699diff
changeset | 1268 | by (simp add: linepath_def real_vector.scale_left_diff_distrib) | 
| 61738 
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
 paulson <lp15@cam.ac.uk> parents: 
61711diff
changeset | 1269 | |
| 64394 | 1270 | lemma linepath_refl: "linepath a a = (\<lambda>x. a)" | 
| 1271 | by auto | |
| 1272 | ||
| 61711 
21d7910d6816
Theory of homotopic paths (from HOL Light), plus comments and minor refinements
 paulson <lp15@cam.ac.uk> parents: 
61699diff
changeset | 1273 | lemma subpath_refl: "subpath a a g = linepath (g a) (g a)" | 
| 
21d7910d6816
Theory of homotopic paths (from HOL Light), plus comments and minor refinements
 paulson <lp15@cam.ac.uk> parents: 
61699diff
changeset | 1274 | by (simp add: subpath_def linepath_def algebra_simps) | 
| 
21d7910d6816
Theory of homotopic paths (from HOL Light), plus comments and minor refinements
 paulson <lp15@cam.ac.uk> parents: 
61699diff
changeset | 1275 | |
| 62618 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1276 | lemma linepath_of_real: "(linepath (of_real a) (of_real b) x) = of_real ((1 - x)*a + x*b)" | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1277 | by (simp add: scaleR_conv_of_real linepath_def) | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1278 | |
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1279 | lemma of_real_linepath: "of_real (linepath a b x) = linepath (of_real a) (of_real b) x" | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1280 | by (metis linepath_of_real mult.right_neutral of_real_def real_scaleR_def) | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1281 | |
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1282 | lemma inj_on_linepath: | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1283 |   assumes "a \<noteq> b" shows "inj_on (linepath a b) {0..1}"
 | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1284 | proof (clarsimp simp: inj_on_def linepath_def) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1285 | fix x y | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1286 | assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b" "0 \<le> x" "x \<le> 1" "0 \<le> y" "y \<le> 1" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1287 | then have "x *\<^sub>R (a - b) = y *\<^sub>R (a - b)" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1288 | by (auto simp: algebra_simps) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1289 | then show "x=y" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1290 | using assms by auto | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1291 | qed | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1292 | |
| 69144 
f13b82281715
new theory Abstract_Topology with lots of stuff from HOL Light's metric.sml
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 1293 | lemma linepath_le_1: | 
| 
f13b82281715
new theory Abstract_Topology with lots of stuff from HOL Light's metric.sml
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 1294 | fixes a::"'a::linordered_idom" shows "\<lbrakk>a \<le> 1; b \<le> 1; 0 \<le> u; u \<le> 1\<rbrakk> \<Longrightarrow> (1 - u) * a + u * b \<le> 1" | 
| 
f13b82281715
new theory Abstract_Topology with lots of stuff from HOL Light's metric.sml
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 1295 | using mult_left_le [of a "1-u"] mult_left_le [of b u] by auto | 
| 
f13b82281715
new theory Abstract_Topology with lots of stuff from HOL Light's metric.sml
 paulson <lp15@cam.ac.uk> parents: 
69064diff
changeset | 1296 | |
| 71189 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1297 | lemma linepath_in_path: | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1298 |   shows "x \<in> {0..1} \<Longrightarrow> linepath a b x \<in> closed_segment a b"
 | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1299 | by (auto simp: segment linepath_def) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1300 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1301 | lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b"
 | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1302 | by (auto simp: segment linepath_def) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1303 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1304 | lemma linepath_in_convex_hull: | 
| 72256 | 1305 | fixes x::real | 
| 1306 | assumes a: "a \<in> convex hull S" | |
| 1307 | and b: "b \<in> convex hull S" | |
| 1308 | and x: "0\<le>x" "x\<le>1" | |
| 1309 | shows "linepath a b x \<in> convex hull S" | |
| 1310 | proof - | |
| 1311 | have "linepath a b x \<in> closed_segment a b" | |
| 1312 | using x by (auto simp flip: linepath_image_01) | |
| 1313 | then show ?thesis | |
| 1314 | using a b convex_contains_segment by blast | |
| 1315 | qed | |
| 71189 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1316 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1317 | lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1318 | by (simp add: linepath_def) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1319 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1320 | lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1321 | by (simp add: linepath_def) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1322 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1323 | lemma bounded_linear_linepath: | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1324 | assumes "bounded_linear f" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1325 | shows "f (linepath a b x) = linepath (f a) (f b) x" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1326 | proof - | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1327 | interpret f: bounded_linear f by fact | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1328 | show ?thesis by (simp add: linepath_def f.add f.scale) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1329 | qed | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1330 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1331 | lemma bounded_linear_linepath': | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1332 | assumes "bounded_linear f" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1333 | shows "f \<circ> linepath a b = linepath (f a) (f b)" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1334 | using bounded_linear_linepath[OF assms] by (simp add: fun_eq_iff) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1335 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1336 | lemma linepath_cnj': "cnj \<circ> linepath a b = linepath (cnj a) (cnj b)" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1337 | by (simp add: linepath_def fun_eq_iff) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1338 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1339 | lemma differentiable_linepath [intro]: "linepath a b differentiable at x within A" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1340 | by (auto simp: linepath_def) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1341 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1342 | lemma has_vector_derivative_linepath_within: | 
| 72256 | 1343 | "(linepath a b has_vector_derivative (b - a)) (at x within S)" | 
| 1344 | by (force intro: derivative_eq_intros simp add: linepath_def has_vector_derivative_def algebra_simps) | |
| 71189 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 1345 | |
| 62618 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1346 | |
| 70136 | 1347 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Segments via convex hulls\<close> | 
| 62618 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1348 | |
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1349 | lemma segments_subset_convex_hull: | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1350 |     "closed_segment a b \<subseteq> (convex hull {a,b,c})"
 | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1351 |     "closed_segment a c \<subseteq> (convex hull {a,b,c})"
 | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1352 |     "closed_segment b c \<subseteq> (convex hull {a,b,c})"
 | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1353 |     "closed_segment b a \<subseteq> (convex hull {a,b,c})"
 | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1354 |     "closed_segment c a \<subseteq> (convex hull {a,b,c})"
 | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1355 |     "closed_segment c b \<subseteq> (convex hull {a,b,c})"
 | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1356 | by (auto simp: segment_convex_hull linepath_of_real elim!: rev_subsetD [OF _ hull_mono]) | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1357 | |
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1358 | lemma midpoints_in_convex_hull: | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1359 | assumes "x \<in> convex hull s" "y \<in> convex hull s" | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1360 | shows "midpoint x y \<in> convex hull s" | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1361 | proof - | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1362 | have "(1 - inverse(2)) *\<^sub>R x + inverse(2) *\<^sub>R y \<in> convex hull s" | 
| 68096 | 1363 | by (rule convexD_alt) (use assms in auto) | 
| 62618 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1364 | then show ?thesis | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1365 | by (simp add: midpoint_def algebra_simps) | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1366 | qed | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1367 | |
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1368 | lemma not_in_interior_convex_hull_3: | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1369 | fixes a :: "complex" | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1370 |   shows "a \<notin> interior(convex hull {a,b,c})"
 | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1371 |         "b \<notin> interior(convex hull {a,b,c})"
 | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1372 |         "c \<notin> interior(convex hull {a,b,c})"
 | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1373 | by (auto simp: card_insert_le_m1 not_in_interior_convex_hull) | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1374 | |
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1375 | lemma midpoint_in_closed_segment [simp]: "midpoint a b \<in> closed_segment a b" | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1376 | using midpoints_in_convex_hull segment_convex_hull by blast | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1377 | |
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1378 | lemma midpoint_in_open_segment [simp]: "midpoint a b \<in> open_segment a b \<longleftrightarrow> a \<noteq> b" | 
| 64122 | 1379 | by (simp add: open_segment_def) | 
| 1380 | ||
| 1381 | lemma continuous_IVT_local_extremum: | |
| 1382 | fixes f :: "'a::euclidean_space \<Rightarrow> real" | |
| 1383 | assumes contf: "continuous_on (closed_segment a b) f" | |
| 1384 | and "a \<noteq> b" "f a = f b" | |
| 1385 | obtains z where "z \<in> open_segment a b" | |
| 1386 | "(\<forall>w \<in> closed_segment a b. (f w) \<le> (f z)) \<or> | |
| 1387 | (\<forall>w \<in> closed_segment a b. (f z) \<le> (f w))" | |
| 1388 | proof - | |
| 1389 | obtain c where "c \<in> closed_segment a b" and c: "\<And>y. y \<in> closed_segment a b \<Longrightarrow> f y \<le> f c" | |
| 1390 | using continuous_attains_sup [of "closed_segment a b" f] contf by auto | |
| 1391 | obtain d where "d \<in> closed_segment a b" and d: "\<And>y. y \<in> closed_segment a b \<Longrightarrow> f d \<le> f y" | |
| 1392 | using continuous_attains_inf [of "closed_segment a b" f] contf by auto | |
| 1393 | show ?thesis | |
| 1394 | proof (cases "c \<in> open_segment a b \<or> d \<in> open_segment a b") | |
| 1395 | case True | |
| 1396 | then show ?thesis | |
| 1397 | using c d that by blast | |
| 1398 | next | |
| 1399 | case False | |
| 1400 | then have "(c = a \<or> c = b) \<and> (d = a \<or> d = b)" | |
| 1401 | by (simp add: \<open>c \<in> closed_segment a b\<close> \<open>d \<in> closed_segment a b\<close> open_segment_def) | |
| 1402 | with \<open>a \<noteq> b\<close> \<open>f a = f b\<close> c d show ?thesis | |
| 1403 | by (rule_tac z = "midpoint a b" in that) (fastforce+) | |
| 1404 | qed | |
| 1405 | qed | |
| 1406 | ||
| 1407 | text\<open>An injective map into R is also an open map w.r.T. the universe, and conversely. \<close> | |
| 1408 | proposition injective_eq_1d_open_map_UNIV: | |
| 1409 | fixes f :: "real \<Rightarrow> real" | |
| 1410 | assumes contf: "continuous_on S f" and S: "is_interval S" | |
| 1411 | shows "inj_on f S \<longleftrightarrow> (\<forall>T. open T \<and> T \<subseteq> S \<longrightarrow> open(f ` T))" | |
| 1412 | (is "?lhs = ?rhs") | |
| 1413 | proof safe | |
| 1414 | fix T | |
| 1415 | assume injf: ?lhs and "open T" and "T \<subseteq> S" | |
| 1416 | have "\<exists>U. open U \<and> f x \<in> U \<and> U \<subseteq> f ` T" if "x \<in> T" for x | |
| 1417 | proof - | |
| 1418 | obtain \<delta> where "\<delta> > 0" and \<delta>: "cball x \<delta> \<subseteq> T" | |
| 1419 | using \<open>open T\<close> \<open>x \<in> T\<close> open_contains_cball_eq by blast | |
| 1420 | show ?thesis | |
| 1421 | proof (intro exI conjI) | |
| 1422 |       have "closed_segment (x-\<delta>) (x+\<delta>) = {x-\<delta>..x+\<delta>}"
 | |
| 1423 | using \<open>0 < \<delta>\<close> by (auto simp: closed_segment_eq_real_ivl) | |
| 68096 | 1424 | also have "\<dots> \<subseteq> S" | 
| 64122 | 1425 | using \<delta> \<open>T \<subseteq> S\<close> by (auto simp: dist_norm subset_eq) | 
| 1426 | finally have "f ` (open_segment (x-\<delta>) (x+\<delta>)) = open_segment (f (x-\<delta>)) (f (x+\<delta>))" | |
| 1427 | using continuous_injective_image_open_segment_1 | |
| 1428 | by (metis continuous_on_subset [OF contf] inj_on_subset [OF injf]) | |
| 1429 |       then show "open (f ` {x-\<delta><..<x+\<delta>})"
 | |
| 1430 | using \<open>0 < \<delta>\<close> by (simp add: open_segment_eq_real_ivl) | |
| 1431 |       show "f x \<in> f ` {x - \<delta><..<x + \<delta>}"
 | |
| 1432 | by (auto simp: \<open>\<delta> > 0\<close>) | |
| 1433 |       show "f ` {x - \<delta><..<x + \<delta>} \<subseteq> f ` T"
 | |
| 1434 | using \<delta> by (auto simp: dist_norm subset_iff) | |
| 1435 | qed | |
| 1436 | qed | |
| 1437 | with open_subopen show "open (f ` T)" | |
| 1438 | by blast | |
| 1439 | next | |
| 1440 | assume R: ?rhs | |
| 1441 | have False if xy: "x \<in> S" "y \<in> S" and "f x = f y" "x \<noteq> y" for x y | |
| 1442 | proof - | |
| 1443 | have "open (f ` open_segment x y)" | |
| 1444 | using R | |
| 1445 | by (metis S convex_contains_open_segment is_interval_convex open_greaterThanLessThan open_segment_eq_real_ivl xy) | |
| 1446 | moreover | |
| 1447 | have "continuous_on (closed_segment x y) f" | |
| 1448 | by (meson S closed_segment_subset contf continuous_on_subset is_interval_convex that) | |
| 1449 | then obtain \<xi> where "\<xi> \<in> open_segment x y" | |
| 1450 | and \<xi>: "(\<forall>w \<in> closed_segment x y. (f w) \<le> (f \<xi>)) \<or> | |
| 1451 | (\<forall>w \<in> closed_segment x y. (f \<xi>) \<le> (f w))" | |
| 1452 | using continuous_IVT_local_extremum [of x y f] \<open>f x = f y\<close> \<open>x \<noteq> y\<close> by blast | |
| 1453 | ultimately obtain e where "e>0" and e: "\<And>u. dist u (f \<xi>) < e \<Longrightarrow> u \<in> f ` open_segment x y" | |
| 1454 | using open_dist by (metis image_eqI) | |
| 1455 | have fin: "f \<xi> + (e/2) \<in> f ` open_segment x y" "f \<xi> - (e/2) \<in> f ` open_segment x y" | |
| 1456 | using e [of "f \<xi> + (e/2)"] e [of "f \<xi> - (e/2)"] \<open>e > 0\<close> by (auto simp: dist_norm) | |
| 1457 | show ?thesis | |
| 1458 | using \<xi> \<open>0 < e\<close> fin open_closed_segment by fastforce | |
| 1459 | qed | |
| 1460 | then show ?lhs | |
| 1461 | by (force simp: inj_on_def) | |
| 1462 | qed | |
| 36583 | 1463 | |
| 69514 | 1464 | |
| 70136 | 1465 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Bounding a point away from a path\<close> | 
| 36583 | 1466 | |
| 1467 | lemma not_on_path_ball: | |
| 1468 | fixes g :: "real \<Rightarrow> 'a::heine_borel" | |
| 53640 | 1469 | assumes "path g" | 
| 68096 | 1470 | and z: "z \<notin> path_image g" | 
| 53640 | 1471 |   shows "\<exists>e > 0. ball z e \<inter> path_image g = {}"
 | 
| 49653 | 1472 | proof - | 
| 68096 | 1473 | have "closed (path_image g)" | 
| 1474 | by (simp add: \<open>path g\<close> closed_path_image) | |
| 1475 | then obtain a where "a \<in> path_image g" "\<forall>y \<in> path_image g. dist z a \<le> dist z y" | |
| 1476 | by (auto intro: distance_attains_inf[OF _ path_image_nonempty, of g z]) | |
| 49654 | 1477 | then show ?thesis | 
| 68096 | 1478 | by (rule_tac x="dist z a" in exI) (use dist_commute z in auto) | 
| 49653 | 1479 | qed | 
| 36583 | 1480 | |
| 1481 | lemma not_on_path_cball: | |
| 1482 | fixes g :: "real \<Rightarrow> 'a::heine_borel" | |
| 53640 | 1483 | assumes "path g" | 
| 1484 | and "z \<notin> path_image g" | |
| 49653 | 1485 |   shows "\<exists>e>0. cball z e \<inter> (path_image g) = {}"
 | 
| 1486 | proof - | |
| 53640 | 1487 |   obtain e where "ball z e \<inter> path_image g = {}" "e > 0"
 | 
| 49653 | 1488 | using not_on_path_ball[OF assms] by auto | 
| 53640 | 1489 | moreover have "cball z (e/2) \<subseteq> ball z e" | 
| 60420 | 1490 | using \<open>e > 0\<close> by auto | 
| 53640 | 1491 | ultimately show ?thesis | 
| 68096 | 1492 | by (rule_tac x="e/2" in exI) auto | 
| 49653 | 1493 | qed | 
| 1494 | ||
| 69518 | 1495 | subsection \<open>Path component\<close> | 
| 1496 | ||
| 1497 | text \<open>Original formalization by Tom Hales\<close> | |
| 36583 | 1498 | |
| 72256 | 1499 | definition\<^marker>\<open>tag important\<close> "path_component S x y \<equiv> | 
| 1500 | (\<exists>g. path g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y)" | |
| 36583 | 1501 | |
| 70136 | 1502 | abbreviation\<^marker>\<open>tag important\<close> | 
| 72256 | 1503 | "path_component_set S x \<equiv> Collect (path_component S x)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1504 | |
| 53640 | 1505 | lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def | 
| 36583 | 1506 | |
| 49653 | 1507 | lemma path_component_mem: | 
| 72256 | 1508 | assumes "path_component S x y" | 
| 1509 | shows "x \<in> S" and "y \<in> S" | |
| 53640 | 1510 | using assms | 
| 1511 | unfolding path_defs | |
| 1512 | by auto | |
| 36583 | 1513 | |
| 49653 | 1514 | lemma path_component_refl: | 
| 72256 | 1515 | assumes "x \<in> S" | 
| 1516 | shows "path_component S x x" | |
| 53640 | 1517 | using assms | 
| 72256 | 1518 | unfolding path_defs | 
| 1519 | by (metis (full_types) assms continuous_on_const image_subset_iff path_image_def) | |
| 1520 | ||
| 1521 | lemma path_component_refl_eq: "path_component S x x \<longleftrightarrow> x \<in> S" | |
| 49653 | 1522 | by (auto intro!: path_component_mem path_component_refl) | 
| 36583 | 1523 | |
| 72256 | 1524 | lemma path_component_sym: "path_component S x y \<Longrightarrow> path_component S y x" | 
| 49653 | 1525 | unfolding path_component_def | 
| 72256 | 1526 | by (metis (no_types) path_image_reversepath path_reversepath pathfinish_reversepath pathstart_reversepath) | 
| 36583 | 1527 | |
| 49653 | 1528 | lemma path_component_trans: | 
| 72256 | 1529 | assumes "path_component S x y" and "path_component S y z" | 
| 1530 | shows "path_component S x z" | |
| 49653 | 1531 | using assms | 
| 1532 | unfolding path_component_def | |
| 72256 | 1533 | by (metis path_join pathfinish_join pathstart_join subset_path_image_join) | 
| 1534 | ||
| 1535 | lemma path_component_of_subset: "S \<subseteq> T \<Longrightarrow> path_component S x y \<Longrightarrow> path_component T x y" | |
| 36583 | 1536 | unfolding path_component_def by auto | 
| 1537 | ||
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1538 | lemma path_component_linepath: | 
| 72256 | 1539 | fixes S :: "'a::real_normed_vector set" | 
| 1540 | shows "closed_segment a b \<subseteq> S \<Longrightarrow> path_component S a b" | |
| 68096 | 1541 | unfolding path_component_def | 
| 1542 | by (rule_tac x="linepath a b" in exI, auto) | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1543 | |
| 70136 | 1544 | subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Path components as sets\<close> | 
| 36583 | 1545 | |
| 49653 | 1546 | lemma path_component_set: | 
| 72256 | 1547 | "path_component_set S x = | 
| 1548 |     {y. (\<exists>g. path g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y)}"
 | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1549 | by (auto simp: path_component_def) | 
| 36583 | 1550 | |
| 72256 | 1551 | lemma path_component_subset: "path_component_set S x \<subseteq> S" | 
| 68096 | 1552 | by (auto simp: path_component_mem(2)) | 
| 36583 | 1553 | |
| 72256 | 1554 | lemma path_component_eq_empty: "path_component_set S x = {} \<longleftrightarrow> x \<notin> S"
 | 
| 60303 | 1555 | using path_component_mem path_component_refl_eq | 
| 1556 | by fastforce | |
| 36583 | 1557 | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1558 | lemma path_component_mono: | 
| 72256 | 1559 | "S \<subseteq> T \<Longrightarrow> (path_component_set S x) \<subseteq> (path_component_set T x)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1560 | by (simp add: Collect_mono path_component_of_subset) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1561 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1562 | lemma path_component_eq: | 
| 72256 | 1563 | "y \<in> path_component_set S x \<Longrightarrow> path_component_set S y = path_component_set S x" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1564 | by (metis (no_types, lifting) Collect_cong mem_Collect_eq path_component_sym path_component_trans) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1565 | |
| 69514 | 1566 | |
| 60420 | 1567 | subsection \<open>Path connectedness of a space\<close> | 
| 36583 | 1568 | |
| 72256 | 1569 | definition\<^marker>\<open>tag important\<close> "path_connected S \<longleftrightarrow> | 
| 1570 | (\<forall>x\<in>S. \<forall>y\<in>S. \<exists>g. path g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y)" | |
| 36583 | 1571 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1572 | lemma path_connectedin_iff_path_connected_real [simp]: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1573 | "path_connectedin euclideanreal S \<longleftrightarrow> path_connected S" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1574 | by (simp add: path_connectedin path_connected_def path_defs) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1575 | |
| 72256 | 1576 | lemma path_connected_component: "path_connected S \<longleftrightarrow> (\<forall>x\<in>S. \<forall>y\<in>S. path_component S x y)" | 
| 36583 | 1577 | unfolding path_connected_def path_component_def by auto | 
| 1578 | ||
| 72256 | 1579 | lemma path_connected_component_set: "path_connected S \<longleftrightarrow> (\<forall>x\<in>S. path_component_set S x = S)" | 
| 61694 
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 1580 | unfolding path_connected_component path_component_subset | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1581 | using path_component_mem by blast | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1582 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1583 | lemma path_component_maximal: | 
| 72256 | 1584 | "\<lbrakk>x \<in> T; path_connected T; T \<subseteq> S\<rbrakk> \<Longrightarrow> T \<subseteq> (path_component_set S x)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1585 | by (metis path_component_mono path_connected_component_set) | 
| 36583 | 1586 | |
| 1587 | lemma convex_imp_path_connected: | |
| 72256 | 1588 | fixes S :: "'a::real_normed_vector set" | 
| 1589 | assumes "convex S" | |
| 1590 | shows "path_connected S" | |
| 49653 | 1591 | unfolding path_connected_def | 
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 1592 | using assms convex_contains_segment by fastforce | 
| 36583 | 1593 | |
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1594 | lemma path_connected_UNIV [iff]: "path_connected (UNIV :: 'a::real_normed_vector set)" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1595 | by (simp add: convex_imp_path_connected) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1596 | |
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1597 | lemma path_component_UNIV: "path_component_set UNIV x = (UNIV :: 'a::real_normed_vector set)" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1598 | using path_connected_component_set by auto | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1599 | |
| 49653 | 1600 | lemma path_connected_imp_connected: | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1601 | assumes "path_connected S" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1602 | shows "connected S" | 
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 1603 | proof (rule connectedI) | 
| 49653 | 1604 | fix e1 e2 | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1605 |   assume as: "open e1" "open e2" "S \<subseteq> e1 \<union> e2" "e1 \<inter> e2 \<inter> S = {}" "e1 \<inter> S \<noteq> {}" "e2 \<inter> S \<noteq> {}"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1606 | then obtain x1 x2 where obt:"x1 \<in> e1 \<inter> S" "x2 \<in> e2 \<inter> S" | 
| 53640 | 1607 | by auto | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1608 | then obtain g where g: "path g" "path_image g \<subseteq> S" "pathstart g = x1" "pathfinish g = x2" | 
| 36583 | 1609 | using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto | 
| 49653 | 1610 |   have *: "connected {0..1::real}"
 | 
| 71172 | 1611 | by (auto intro!: convex_connected) | 
| 49653 | 1612 |   have "{0..1} \<subseteq> {x \<in> {0..1}. g x \<in> e1} \<union> {x \<in> {0..1}. g x \<in> e2}"
 | 
| 1613 | using as(3) g(2)[unfolded path_defs] by blast | |
| 1614 |   moreover have "{x \<in> {0..1}. g x \<in> e1} \<inter> {x \<in> {0..1}. g x \<in> e2} = {}"
 | |
| 53640 | 1615 | using as(4) g(2)[unfolded path_defs] | 
| 1616 | unfolding subset_eq | |
| 1617 | by auto | |
| 49653 | 1618 |   moreover have "{x \<in> {0..1}. g x \<in> e1} \<noteq> {} \<and> {x \<in> {0..1}. g x \<in> e2} \<noteq> {}"
 | 
| 53640 | 1619 | using g(3,4)[unfolded path_defs] | 
| 1620 | using obt | |
| 36583 | 1621 | by (simp add: ex_in_conv [symmetric], metis zero_le_one order_refl) | 
| 49653 | 1622 | ultimately show False | 
| 53640 | 1623 | using *[unfolded connected_local not_ex, rule_format, | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1624 |       of "{0..1} \<inter> g -` e1" "{0..1} \<inter> g -` e2"]
 | 
| 63301 | 1625 | using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(1)] | 
| 1626 | using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(2)] | |
| 49653 | 1627 | by auto | 
| 1628 | qed | |
| 36583 | 1629 | |
| 1630 | lemma open_path_component: | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1631 | fixes S :: "'a::real_normed_vector set" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1632 | assumes "open S" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1633 | shows "open (path_component_set S x)" | 
| 49653 | 1634 | unfolding open_contains_ball | 
| 1635 | proof | |
| 1636 | fix y | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1637 | assume as: "y \<in> path_component_set S x" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1638 | then have "y \<in> S" | 
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 1639 | by (simp add: path_component_mem(2)) | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1640 | then obtain e where e: "e > 0" "ball y e \<subseteq> S" | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 1641 | using assms openE by blast | 
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 1642 | have "\<And>u. dist y u < e \<Longrightarrow> path_component S x u" | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 1643 | by (metis (full_types) as centre_in_ball convex_ball convex_imp_path_connected e mem_Collect_eq mem_ball path_component_eq path_component_of_subset path_connected_component) | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 1644 | then show "\<exists>e > 0. ball y e \<subseteq> path_component_set S x" | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 1645 | using \<open>e>0\<close> by auto | 
| 49653 | 1646 | qed | 
| 36583 | 1647 | |
| 1648 | lemma open_non_path_component: | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1649 | fixes S :: "'a::real_normed_vector set" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1650 | assumes "open S" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1651 | shows "open (S - path_component_set S x)" | 
| 49653 | 1652 | unfolding open_contains_ball | 
| 1653 | proof | |
| 1654 | fix y | |
| 68096 | 1655 | assume y: "y \<in> S - path_component_set S x" | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1656 | then obtain e where e: "e > 0" "ball y e \<subseteq> S" | 
| 68096 | 1657 | using assms openE by auto | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1658 | show "\<exists>e>0. ball y e \<subseteq> S - path_component_set S x" | 
| 68096 | 1659 | proof (intro exI conjI subsetI DiffI notI) | 
| 1660 | show "\<And>x. x \<in> ball y e \<Longrightarrow> x \<in> S" | |
| 1661 | using e by blast | |
| 1662 | show False if "z \<in> ball y e" "z \<in> path_component_set S x" for z | |
| 1663 | proof - | |
| 1664 | have "y \<in> path_component_set S z" | |
| 1665 | by (meson assms convex_ball convex_imp_path_connected e open_contains_ball_eq open_path_component path_component_maximal that(1)) | |
| 1666 | then have "y \<in> path_component_set S x" | |
| 1667 | using path_component_eq that(2) by blast | |
| 1668 | then show False | |
| 1669 | using y by blast | |
| 1670 | qed | |
| 1671 | qed (use e in auto) | |
| 49653 | 1672 | qed | 
| 36583 | 1673 | |
| 1674 | lemma connected_open_path_connected: | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1675 | fixes S :: "'a::real_normed_vector set" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1676 | assumes "open S" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1677 | and "connected S" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1678 | shows "path_connected S" | 
| 49653 | 1679 | unfolding path_connected_component_set | 
| 1680 | proof (rule, rule, rule path_component_subset, rule) | |
| 1681 | fix x y | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1682 | assume "x \<in> S" and "y \<in> S" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1683 | show "y \<in> path_component_set S x" | 
| 49653 | 1684 | proof (rule ccontr) | 
| 53640 | 1685 | assume "\<not> ?thesis" | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1686 |     moreover have "path_component_set S x \<inter> S \<noteq> {}"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1687 | using \<open>x \<in> S\<close> path_component_eq_empty path_component_subset[of S x] | 
| 53640 | 1688 | by auto | 
| 49653 | 1689 | ultimately | 
| 1690 | show False | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1691 | using \<open>y \<in> S\<close> open_non_path_component[OF assms(1)] open_path_component[OF assms(1)] | 
| 53640 | 1692 | using assms(2)[unfolded connected_def not_ex, rule_format, | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1693 | of "path_component_set S x" "S - path_component_set S x"] | 
| 49653 | 1694 | by auto | 
| 1695 | qed | |
| 1696 | qed | |
| 36583 | 1697 | |
| 1698 | lemma path_connected_continuous_image: | |
| 72256 | 1699 | assumes contf: "continuous_on S f" | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1700 | and "path_connected S" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1701 | shows "path_connected (f ` S)" | 
| 49653 | 1702 | unfolding path_connected_def | 
| 1703 | proof (rule, rule) | |
| 1704 | fix x' y' | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1705 | assume "x' \<in> f ` S" "y' \<in> f ` S" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1706 | then obtain x y where x: "x \<in> S" and y: "y \<in> S" and x': "x' = f x" and y': "y' = f y" | 
| 53640 | 1707 | by auto | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1708 | from x y obtain g where "path g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y" | 
| 53640 | 1709 | using assms(2)[unfolded path_connected_def] by fast | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1710 | then show "\<exists>g. path g \<and> path_image g \<subseteq> f ` S \<and> pathstart g = x' \<and> pathfinish g = y'" | 
| 72256 | 1711 | unfolding x' y' path_defs | 
| 1712 | by (fastforce intro: continuous_on_compose continuous_on_subset[OF contf]) | |
| 49653 | 1713 | qed | 
| 36583 | 1714 | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1715 | lemma path_connected_translationI: | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1716 | fixes a :: "'a :: topological_group_add" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1717 | assumes "path_connected S" shows "path_connected ((\<lambda>x. a + x) ` S)" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1718 | by (intro path_connected_continuous_image assms continuous_intros) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1719 | |
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1720 | lemma path_connected_translation: | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1721 | fixes a :: "'a :: topological_group_add" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1722 | shows "path_connected ((\<lambda>x. a + x) ` S) = path_connected S" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1723 | proof - | 
| 67399 | 1724 | have "\<forall>x y. (+) (x::'a) ` (+) (0 - x) ` y = y" | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1725 | by (simp add: image_image) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1726 | then show ?thesis | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1727 | by (metis (no_types) path_connected_translationI) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1728 | qed | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1729 | |
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1730 | lemma path_connected_segment [simp]: | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1731 | fixes a :: "'a::real_normed_vector" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1732 | shows "path_connected (closed_segment a b)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1733 | by (simp add: convex_imp_path_connected) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1734 | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1735 | lemma path_connected_open_segment [simp]: | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1736 | fixes a :: "'a::real_normed_vector" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1737 | shows "path_connected (open_segment a b)" | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1738 | by (simp add: convex_imp_path_connected) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 1739 | |
| 36583 | 1740 | lemma homeomorphic_path_connectedness: | 
| 68096 | 1741 | "S homeomorphic T \<Longrightarrow> path_connected S \<longleftrightarrow> path_connected T" | 
| 61738 
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
 paulson <lp15@cam.ac.uk> parents: 
61711diff
changeset | 1742 | unfolding homeomorphic_def homeomorphism_def by (metis path_connected_continuous_image) | 
| 36583 | 1743 | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1744 | lemma path_connected_empty [simp]: "path_connected {}"
 | 
| 36583 | 1745 | unfolding path_connected_def by auto | 
| 1746 | ||
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1747 | lemma path_connected_singleton [simp]: "path_connected {a}"
 | 
| 36583 | 1748 | unfolding path_connected_def pathstart_def pathfinish_def path_image_def | 
| 72256 | 1749 | using path_def by fastforce | 
| 36583 | 1750 | |
| 49653 | 1751 | lemma path_connected_Un: | 
| 68096 | 1752 | assumes "path_connected S" | 
| 1753 | and "path_connected T" | |
| 1754 |     and "S \<inter> T \<noteq> {}"
 | |
| 1755 | shows "path_connected (S \<union> T)" | |
| 49653 | 1756 | unfolding path_connected_component | 
| 68096 | 1757 | proof (intro ballI) | 
| 49653 | 1758 | fix x y | 
| 68096 | 1759 | assume x: "x \<in> S \<union> T" and y: "y \<in> S \<union> T" | 
| 1760 | from assms obtain z where z: "z \<in> S" "z \<in> T" | |
| 53640 | 1761 | by auto | 
| 68096 | 1762 | show "path_component (S \<union> T) x y" | 
| 1763 | using x y | |
| 1764 | proof safe | |
| 1765 | assume "x \<in> S" "y \<in> S" | |
| 1766 | then show "path_component (S \<union> T) x y" | |
| 1767 | by (meson Un_upper1 \<open>path_connected S\<close> path_component_of_subset path_connected_component) | |
| 1768 | next | |
| 1769 | assume "x \<in> S" "y \<in> T" | |
| 1770 | then show "path_component (S \<union> T) x y" | |
| 1771 | by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component) | |
| 1772 | next | |
| 1773 | assume "x \<in> T" "y \<in> S" | |
| 1774 | then show "path_component (S \<union> T) x y" | |
| 1775 | by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component) | |
| 1776 | next | |
| 1777 | assume "x \<in> T" "y \<in> T" | |
| 1778 | then show "path_component (S \<union> T) x y" | |
| 1779 | by (metis Un_upper1 assms(2) path_component_of_subset path_connected_component sup_commute) | |
| 1780 | qed | |
| 49653 | 1781 | qed | 
| 36583 | 1782 | |
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 1783 | lemma path_connected_UNION: | 
| 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 1784 | assumes "\<And>i. i \<in> A \<Longrightarrow> path_connected (S i)" | 
| 49653 | 1785 | and "\<And>i. i \<in> A \<Longrightarrow> z \<in> S i" | 
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 1786 | shows "path_connected (\<Union>i\<in>A. S i)" | 
| 49653 | 1787 | unfolding path_connected_component | 
| 1788 | proof clarify | |
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 1789 | fix x i y j | 
| 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 1790 | assume *: "i \<in> A" "x \<in> S i" "j \<in> A" "y \<in> S j" | 
| 49654 | 1791 | then have "path_component (S i) x z" and "path_component (S j) z y" | 
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 1792 | using assms by (simp_all add: path_connected_component) | 
| 49654 | 1793 | then have "path_component (\<Union>i\<in>A. S i) x z" and "path_component (\<Union>i\<in>A. S i) z y" | 
| 48125 
602dc0215954
tuned proofs -- prefer direct "rotated" instead of old-style COMP;
 wenzelm parents: 
44647diff
changeset | 1794 | using *(1,3) by (auto elim!: path_component_of_subset [rotated]) | 
| 49654 | 1795 | then show "path_component (\<Union>i\<in>A. S i) x y" | 
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 1796 | by (rule path_component_trans) | 
| 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 1797 | qed | 
| 36583 | 1798 | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1799 | lemma path_component_path_image_pathstart: | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1800 | assumes p: "path p" and x: "x \<in> path_image p" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1801 | shows "path_component (path_image p) (pathstart p) x" | 
| 68096 | 1802 | proof - | 
| 1803 | obtain y where x: "x = p y" and y: "0 \<le> y" "y \<le> 1" | |
| 1804 | using x by (auto simp: path_image_def) | |
| 1805 | show ?thesis | |
| 1806 | unfolding path_component_def | |
| 1807 | proof (intro exI conjI) | |
| 72256 | 1808 |     have "continuous_on ((*) y ` {0..1}) p"
 | 
| 1809 | by (simp add: continuous_on_path image_mult_atLeastAtMost_if p y) | |
| 1810 |     then have "continuous_on {0..1} (p \<circ> ((*) y))"
 | |
| 1811 | using continuous_on_compose continuous_on_mult_const by blast | |
| 68096 | 1812 | then show "path (\<lambda>u. p (y * u))" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1813 | by (simp add: path_def) | 
| 68096 | 1814 | show "path_image (\<lambda>u. p (y * u)) \<subseteq> path_image p" | 
| 1815 | using y mult_le_one by (fastforce simp: path_image_def image_iff) | |
| 1816 | qed (auto simp: pathstart_def pathfinish_def x) | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1817 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1818 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1819 | lemma path_connected_path_image: "path p \<Longrightarrow> path_connected(path_image p)" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1820 | unfolding path_connected_component | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1821 | by (meson path_component_path_image_pathstart path_component_sym path_component_trans) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1822 | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1823 | lemma path_connected_path_component [simp]: | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1824 | "path_connected (path_component_set s x)" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1825 | proof - | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1826 |   { fix y z
 | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1827 | assume pa: "path_component s x y" "path_component s x z" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1828 | then have pae: "path_component_set s x = path_component_set s y" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1829 | using path_component_eq by auto | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1830 | have yz: "path_component s y z" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1831 | using pa path_component_sym path_component_trans by blast | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1832 | then have "\<exists>g. path g \<and> path_image g \<subseteq> path_component_set s x \<and> pathstart g = y \<and> pathfinish g = z" | 
| 72256 | 1833 | apply (simp add: path_component_def) | 
| 1834 | by (metis pae path_component_maximal path_connected_path_image pathstart_in_path_image) | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1835 | } | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1836 | then show ?thesis | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1837 | by (simp add: path_connected_def) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1838 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1839 | |
| 68532 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68310diff
changeset | 1840 | lemma path_component: "path_component S x y \<longleftrightarrow> (\<exists>t. path_connected t \<and> t \<subseteq> S \<and> x \<in> t \<and> y \<in> t)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1841 | apply (intro iffI) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1842 | apply (metis path_connected_path_image path_defs(5) pathfinish_in_path_image pathstart_in_path_image) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1843 | using path_component_of_subset path_connected_component by blast | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1844 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1845 | lemma path_component_path_component [simp]: | 
| 68532 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68310diff
changeset | 1846 | "path_component_set (path_component_set S x) x = path_component_set S x" | 
| 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68310diff
changeset | 1847 | proof (cases "x \<in> S") | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1848 | case True show ?thesis | 
| 72256 | 1849 | by (metis True mem_Collect_eq path_component_refl path_connected_component_set path_connected_path_component) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1850 | next | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1851 | case False then show ?thesis | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1852 | by (metis False empty_iff path_component_eq_empty) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1853 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1854 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1855 | lemma path_component_subset_connected_component: | 
| 68532 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68310diff
changeset | 1856 | "(path_component_set S x) \<subseteq> (connected_component_set S x)" | 
| 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68310diff
changeset | 1857 | proof (cases "x \<in> S") | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1858 | case True show ?thesis | 
| 72256 | 1859 | by (simp add: True connected_component_maximal path_component_refl path_component_subset path_connected_imp_connected) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1860 | next | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1861 | case False then show ?thesis | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1862 | using path_component_eq_empty by auto | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 1863 | qed | 
| 49653 | 1864 | |
| 69514 | 1865 | |
| 70136 | 1866 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Lemmas about path-connectedness\<close> | 
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1867 | |
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1868 | lemma path_connected_linear_image: | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1869 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 68532 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68310diff
changeset | 1870 | assumes "path_connected S" "bounded_linear f" | 
| 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68310diff
changeset | 1871 | shows "path_connected(f ` S)" | 
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1872 | by (auto simp: linear_continuous_on assms path_connected_continuous_image) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1873 | |
| 68532 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68310diff
changeset | 1874 | lemma is_interval_path_connected: "is_interval S \<Longrightarrow> path_connected S" | 
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1875 | by (simp add: convex_imp_path_connected is_interval_convex) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1876 | |
| 71025 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1877 | lemma path_connected_Ioi[simp]: "path_connected {a<..}" for a :: real
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1878 | by (simp add: convex_imp_path_connected) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1879 | |
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1880 | lemma path_connected_Ici[simp]: "path_connected {a..}" for a :: real
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1881 | by (simp add: convex_imp_path_connected) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1882 | |
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1883 | lemma path_connected_Iio[simp]: "path_connected {..<a}" for a :: real
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1884 | by (simp add: convex_imp_path_connected) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1885 | |
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1886 | lemma path_connected_Iic[simp]: "path_connected {..a}" for a :: real
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1887 | by (simp add: convex_imp_path_connected) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1888 | |
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1889 | lemma path_connected_Ioo[simp]: "path_connected {a<..<b}" for a b :: real
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1890 | by (simp add: convex_imp_path_connected) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1891 | |
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1892 | lemma path_connected_Ioc[simp]: "path_connected {a<..b}" for a b :: real
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1893 | by (simp add: convex_imp_path_connected) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1894 | |
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1895 | lemma path_connected_Ico[simp]: "path_connected {a..<b}" for a b :: real
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1896 | by (simp add: convex_imp_path_connected) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70971diff
changeset | 1897 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1898 | lemma path_connectedin_path_image: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1899 |   assumes "pathin X g" shows "path_connectedin X (g ` ({0..1}))"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1900 | unfolding pathin_def | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1901 | proof (rule path_connectedin_continuous_map_image) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1902 |   show "continuous_map (subtopology euclideanreal {0..1}) X g"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1903 | using assms pathin_def by blast | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1904 | qed (auto simp: is_interval_1 is_interval_path_connected) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1905 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1906 | lemma path_connected_space_subconnected: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1907 | "path_connected_space X \<longleftrightarrow> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1908 | (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. \<exists>S. path_connectedin X S \<and> x \<in> S \<and> y \<in> S)" | 
| 72256 | 1909 | by (metis path_connectedin path_connectedin_topspace path_connected_space_def) | 
| 1910 | ||
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1911 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1912 | lemma connectedin_path_image: "pathin X g \<Longrightarrow> connectedin X (g ` ({0..1}))"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1913 | by (simp add: path_connectedin_imp_connectedin path_connectedin_path_image) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1914 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1915 | lemma compactin_path_image: "pathin X g \<Longrightarrow> compactin X (g ` ({0..1}))"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1916 | unfolding pathin_def | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1917 |   by (rule image_compactin [of "top_of_set {0..1}"]) auto
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 1918 | |
| 62843 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62626diff
changeset | 1919 | lemma linear_homeomorphism_image: | 
| 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62626diff
changeset | 1920 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1921 | assumes "linear f" "inj f" | 
| 72256 | 1922 | obtains g where "homeomorphism (f ` S) S g f" | 
| 1923 | proof - | |
| 1924 | obtain g where "linear g" "g \<circ> f = id" | |
| 1925 | using assms linear_injective_left_inverse by blast | |
| 1926 | then have "homeomorphism (f ` S) S g f" | |
| 1927 | using assms unfolding homeomorphism_def | |
| 1928 | by (auto simp: eq_id_iff [symmetric] image_comp linear_conv_bounded_linear linear_continuous_on) | |
| 1929 | then show thesis .. | |
| 1930 | qed | |
| 62843 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62626diff
changeset | 1931 | |
| 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62626diff
changeset | 1932 | lemma linear_homeomorphic_image: | 
| 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62626diff
changeset | 1933 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62626diff
changeset | 1934 | assumes "linear f" "inj f" | 
| 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62626diff
changeset | 1935 | shows "S homeomorphic f ` S" | 
| 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62626diff
changeset | 1936 | by (meson homeomorphic_def homeomorphic_sym linear_homeomorphism_image [OF assms]) | 
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1937 | |
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1938 | lemma path_connected_Times: | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1939 | assumes "path_connected s" "path_connected t" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1940 | shows "path_connected (s \<times> t)" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1941 | proof (simp add: path_connected_def Sigma_def, clarify) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1942 | fix x1 y1 x2 y2 | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1943 | assume "x1 \<in> s" "y1 \<in> t" "x2 \<in> s" "y2 \<in> t" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1944 | obtain g where "path g" and g: "path_image g \<subseteq> s" and gs: "pathstart g = x1" and gf: "pathfinish g = x2" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1945 | using \<open>x1 \<in> s\<close> \<open>x2 \<in> s\<close> assms by (force simp: path_connected_def) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1946 | obtain h where "path h" and h: "path_image h \<subseteq> t" and hs: "pathstart h = y1" and hf: "pathfinish h = y2" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1947 | using \<open>y1 \<in> t\<close> \<open>y2 \<in> t\<close> assms by (force simp: path_connected_def) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1948 | have "path (\<lambda>z. (x1, h z))" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1949 | using \<open>path h\<close> | 
| 72256 | 1950 | unfolding path_def | 
| 1951 | by (intro continuous_intros continuous_on_compose2 [where g = "Pair _"]; force) | |
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1952 | moreover have "path (\<lambda>z. (g z, y2))" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1953 | using \<open>path g\<close> | 
| 72256 | 1954 | unfolding path_def | 
| 1955 | by (intro continuous_intros continuous_on_compose2 [where g = "Pair _"]; force) | |
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1956 | ultimately have 1: "path ((\<lambda>z. (x1, h z)) +++ (\<lambda>z. (g z, y2)))" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1957 | by (metis hf gs path_join_imp pathstart_def pathfinish_def) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1958 | have "path_image ((\<lambda>z. (x1, h z)) +++ (\<lambda>z. (g z, y2))) \<subseteq> path_image (\<lambda>z. (x1, h z)) \<union> path_image (\<lambda>z. (g z, y2))" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1959 | by (rule Path_Connected.path_image_join_subset) | 
| 68096 | 1960 |   also have "\<dots> \<subseteq> (\<Union>x\<in>s. \<Union>x1\<in>t. {(x, x1)})"
 | 
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1961 | using g h \<open>x1 \<in> s\<close> \<open>y2 \<in> t\<close> by (force simp: path_image_def) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1962 |   finally have 2: "path_image ((\<lambda>z. (x1, h z)) +++ (\<lambda>z. (g z, y2))) \<subseteq> (\<Union>x\<in>s. \<Union>x1\<in>t. {(x, x1)})" .
 | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1963 |   show "\<exists>g. path g \<and> path_image g \<subseteq> (\<Union>x\<in>s. \<Union>x1\<in>t. {(x, x1)}) \<and>
 | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1964 | pathstart g = (x1, y1) \<and> pathfinish g = (x2, y2)" | 
| 72256 | 1965 | using 1 2 gf hs | 
| 1966 | by (metis (no_types, lifting) pathfinish_def pathfinish_join pathstart_def pathstart_join) | |
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1967 | qed | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1968 | |
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1969 | lemma is_interval_path_connected_1: | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1970 | fixes s :: "real set" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1971 | shows "is_interval s \<longleftrightarrow> path_connected s" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1972 | using is_interval_connected_1 is_interval_path_connected path_connected_imp_connected by blast | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1973 | |
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 1974 | |
| 70136 | 1975 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Path components\<close> | 
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 1976 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1977 | lemma Union_path_component [simp]: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1978 |    "Union {path_component_set S x |x. x \<in> S} = S"
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1979 | apply (rule subset_antisym) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1980 | using path_component_subset apply force | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1981 | using path_component_refl by auto | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1982 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1983 | lemma path_component_disjoint: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1984 | "disjnt (path_component_set S a) (path_component_set S b) \<longleftrightarrow> | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1985 | (a \<notin> path_component_set S b)" | 
| 72256 | 1986 | unfolding disjnt_iff | 
| 1987 | using path_component_sym path_component_trans by blast | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1988 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1989 | lemma path_component_eq_eq: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1990 | "path_component S x = path_component S y \<longleftrightarrow> | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 1991 | (x \<notin> S) \<and> (y \<notin> S) \<or> x \<in> S \<and> y \<in> S \<and> path_component S x y" | 
| 72256 | 1992 | (is "?lhs = ?rhs") | 
| 1993 | proof | |
| 1994 | assume ?lhs then show ?rhs | |
| 1995 | by (metis (no_types) path_component_mem(1) path_component_refl) | |
| 1996 | next | |
| 1997 | assume ?rhs then show ?lhs | |
| 1998 | proof | |
| 1999 | assume "x \<notin> S \<and> y \<notin> S" then show ?lhs | |
| 2000 | by (metis Collect_empty_eq_bot path_component_eq_empty) | |
| 2001 | next | |
| 2002 | assume S: "x \<in> S \<and> y \<in> S \<and> path_component S x y" show ?lhs | |
| 2003 | by (rule ext) (metis S path_component_trans path_component_sym) | |
| 2004 | qed | |
| 2005 | qed | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2006 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2007 | lemma path_component_unique: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2008 | assumes "x \<in> c" "c \<subseteq> S" "path_connected c" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2009 | "\<And>c'. \<lbrakk>x \<in> c'; c' \<subseteq> S; path_connected c'\<rbrakk> \<Longrightarrow> c' \<subseteq> c" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2010 | shows "path_component_set S x = c" | 
| 72256 | 2011 | (is "?lhs = ?rhs") | 
| 2012 | proof | |
| 2013 | show "?lhs \<subseteq> ?rhs" | |
| 2014 | using assms | |
| 2015 | by (metis mem_Collect_eq path_component_refl path_component_subset path_connected_path_component subsetD) | |
| 2016 | qed (simp add: assms path_component_maximal) | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2017 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2018 | lemma path_component_intermediate_subset: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2019 | "path_component_set u a \<subseteq> t \<and> t \<subseteq> u | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2020 | \<Longrightarrow> path_component_set t a = path_component_set u a" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2021 | by (metis (no_types) path_component_mono path_component_path_component subset_antisym) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2022 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2023 | lemma complement_path_component_Union: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2024 | fixes x :: "'a :: topological_space" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2025 | shows "S - path_component_set S x = | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2026 |          \<Union>({path_component_set S y| y. y \<in> S} - {path_component_set S x})"
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2027 | proof - | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2028 |   have *: "(\<And>x. x \<in> S - {a} \<Longrightarrow> disjnt a x) \<Longrightarrow> \<Union>S - a = \<Union>(S - {a})"
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2029 | for a::"'a set" and S | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2030 | by (auto simp: disjnt_def) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2031 |   have "\<And>y. y \<in> {path_component_set S x |x. x \<in> S} - {path_component_set S x}
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2032 | \<Longrightarrow> disjnt (path_component_set S x) y" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2033 | using path_component_disjoint path_component_eq by fastforce | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2034 |   then have "\<Union>{path_component_set S x |x. x \<in> S} - path_component_set S x =
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2035 |              \<Union>({path_component_set S y |y. y \<in> S} - {path_component_set S x})"
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2036 | by (meson *) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2037 | then show ?thesis by simp | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2038 | qed | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2039 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2040 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2041 | subsection\<open>Path components\<close> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2042 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2043 | definition path_component_of | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2044 | where "path_component_of X x y \<equiv> \<exists>g. pathin X g \<and> g 0 = x \<and> g 1 = y" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2045 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2046 | abbreviation path_component_of_set | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2047 | where "path_component_of_set X x \<equiv> Collect (path_component_of X x)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2048 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2049 | definition path_components_of :: "'a topology \<Rightarrow> 'a set set" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2050 | where "path_components_of X \<equiv> path_component_of_set X ` topspace X" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2051 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 2052 | lemma pathin_canon_iff: "pathin (top_of_set T) g \<longleftrightarrow> path g \<and> g ` {0..1} \<subseteq> T"
 | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 2053 | by (simp add: path_def pathin_def) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 2054 | |
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 2055 | lemma path_component_of_canon_iff [simp]: | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 2056 | "path_component_of (top_of_set T) a b \<longleftrightarrow> path_component T a b" | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 2057 | by (simp add: path_component_of_def pathin_canon_iff path_defs) | 
| 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 2058 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2059 | lemma path_component_in_topspace: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2060 | "path_component_of X x y \<Longrightarrow> x \<in> topspace X \<and> y \<in> topspace X" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2061 | by (auto simp: path_component_of_def pathin_def continuous_map_def) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2062 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2063 | lemma path_component_of_refl: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2064 | "path_component_of X x x \<longleftrightarrow> x \<in> topspace X" | 
| 72256 | 2065 | by (metis path_component_in_topspace path_component_of_def pathin_const) | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2066 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2067 | lemma path_component_of_sym: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2068 | assumes "path_component_of X x y" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2069 | shows "path_component_of X y x" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2070 | using assms | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2071 | apply (clarsimp simp: path_component_of_def pathin_def) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2072 | apply (rule_tac x="g \<circ> (\<lambda>t. 1 - t)" in exI) | 
| 72256 | 2073 | apply (auto intro!: continuous_map_compose simp: continuous_map_in_subtopology continuous_on_op_minus) | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2074 | done | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2075 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2076 | lemma path_component_of_sym_iff: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2077 | "path_component_of X x y \<longleftrightarrow> path_component_of X y x" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2078 | by (metis path_component_of_sym) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2079 | |
| 71236 | 2080 | lemma continuous_map_cases_le: | 
| 2081 | assumes contp: "continuous_map X euclideanreal p" | |
| 2082 | and contq: "continuous_map X euclideanreal q" | |
| 2083 |     and contf: "continuous_map (subtopology X {x. x \<in> topspace X \<and> p x \<le> q x}) Y f"
 | |
| 2084 |     and contg: "continuous_map (subtopology X {x. x \<in> topspace X \<and> q x \<le> p x}) Y g"
 | |
| 2085 | and fg: "\<And>x. \<lbrakk>x \<in> topspace X; p x = q x\<rbrakk> \<Longrightarrow> f x = g x" | |
| 2086 | shows "continuous_map X Y (\<lambda>x. if p x \<le> q x then f x else g x)" | |
| 2087 | proof - | |
| 2088 |   have "continuous_map X Y (\<lambda>x. if q x - p x \<in> {0..} then f x else g x)"
 | |
| 2089 | proof (rule continuous_map_cases_function) | |
| 2090 | show "continuous_map X euclideanreal (\<lambda>x. q x - p x)" | |
| 2091 | by (intro contp contq continuous_intros) | |
| 2092 |     show "continuous_map (subtopology X {x \<in> topspace X. q x - p x \<in> euclideanreal closure_of {0..}}) Y f"
 | |
| 2093 | by (simp add: contf) | |
| 2094 |     show "continuous_map (subtopology X {x \<in> topspace X. q x - p x \<in> euclideanreal closure_of (topspace euclideanreal - {0..})}) Y g"
 | |
| 2095 | by (simp add: contg flip: Compl_eq_Diff_UNIV) | |
| 2096 | qed (auto simp: fg) | |
| 2097 | then show ?thesis | |
| 2098 | by simp | |
| 2099 | qed | |
| 2100 | ||
| 2101 | lemma continuous_map_cases_lt: | |
| 2102 | assumes contp: "continuous_map X euclideanreal p" | |
| 2103 | and contq: "continuous_map X euclideanreal q" | |
| 2104 |     and contf: "continuous_map (subtopology X {x. x \<in> topspace X \<and> p x \<le> q x}) Y f"
 | |
| 2105 |     and contg: "continuous_map (subtopology X {x. x \<in> topspace X \<and> q x \<le> p x}) Y g"
 | |
| 2106 | and fg: "\<And>x. \<lbrakk>x \<in> topspace X; p x = q x\<rbrakk> \<Longrightarrow> f x = g x" | |
| 2107 | shows "continuous_map X Y (\<lambda>x. if p x < q x then f x else g x)" | |
| 2108 | proof - | |
| 2109 |   have "continuous_map X Y (\<lambda>x. if q x - p x \<in> {0<..} then f x else g x)"
 | |
| 2110 | proof (rule continuous_map_cases_function) | |
| 2111 | show "continuous_map X euclideanreal (\<lambda>x. q x - p x)" | |
| 2112 | by (intro contp contq continuous_intros) | |
| 2113 |     show "continuous_map (subtopology X {x \<in> topspace X. q x - p x \<in> euclideanreal closure_of {0<..}}) Y f"
 | |
| 2114 | by (simp add: contf) | |
| 2115 |     show "continuous_map (subtopology X {x \<in> topspace X. q x - p x \<in> euclideanreal closure_of (topspace euclideanreal - {0<..})}) Y g"
 | |
| 2116 | by (simp add: contg flip: Compl_eq_Diff_UNIV) | |
| 2117 | qed (auto simp: fg) | |
| 2118 | then show ?thesis | |
| 2119 | by simp | |
| 2120 | qed | |
| 2121 | ||
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2122 | lemma path_component_of_trans: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2123 | assumes "path_component_of X x y" and "path_component_of X y z" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2124 | shows "path_component_of X x z" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2125 | unfolding path_component_of_def pathin_def | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2126 | proof - | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2127 |   let ?T01 = "top_of_set {0..1::real}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2128 | obtain g1 g2 where g1: "continuous_map ?T01 X g1" "x = g1 0" "y = g1 1" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2129 | and g2: "continuous_map ?T01 X g2" "g2 0 = g1 1" "z = g2 1" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2130 | using assms unfolding path_component_of_def pathin_def by blast | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2131 | let ?g = "\<lambda>x. if x \<le> 1/2 then (g1 \<circ> (\<lambda>t. 2 * t)) x else (g2 \<circ> (\<lambda>t. 2 * t -1)) x" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2132 | show "\<exists>g. continuous_map ?T01 X g \<and> g 0 = x \<and> g 1 = z" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2133 | proof (intro exI conjI) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2134 |     show "continuous_map (subtopology euclideanreal {0..1}) X ?g"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2135 | proof (intro continuous_map_cases_le continuous_map_compose, force, force) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2136 |       show "continuous_map (subtopology ?T01 {x \<in> topspace ?T01. x \<le> 1/2}) ?T01 ((*) 2)"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2137 | by (auto simp: continuous_map_in_subtopology continuous_map_from_subtopology) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2138 | have "continuous_map | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2139 |              (subtopology (top_of_set {0..1}) {x. 0 \<le> x \<and> x \<le> 1 \<and> 1 \<le> x * 2})
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2140 | euclideanreal (\<lambda>t. 2 * t - 1)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2141 | by (intro continuous_intros) (force intro: continuous_map_from_subtopology) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2142 |       then show "continuous_map (subtopology ?T01 {x \<in> topspace ?T01. 1/2 \<le> x}) ?T01 (\<lambda>t. 2 * t - 1)"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2143 | by (force simp: continuous_map_in_subtopology) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2144 | show "(g1 \<circ> (*) 2) x = (g2 \<circ> (\<lambda>t. 2 * t - 1)) x" if "x \<in> topspace ?T01" "x = 1/2" for x | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2145 | using that by (simp add: g2(2) mult.commute continuous_map_from_subtopology) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2146 | qed (auto simp: g1 g2) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2147 | qed (auto simp: g1 g2) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2148 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2149 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2150 | lemma path_component_of_mono: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2151 | "\<lbrakk>path_component_of (subtopology X S) x y; S \<subseteq> T\<rbrakk> \<Longrightarrow> path_component_of (subtopology X T) x y" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2152 | unfolding path_component_of_def | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2153 | by (metis subsetD pathin_subtopology) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2154 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2155 | lemma path_component_of: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2156 | "path_component_of X x y \<longleftrightarrow> (\<exists>T. path_connectedin X T \<and> x \<in> T \<and> y \<in> T)" | 
| 72256 | 2157 | (is "?lhs = ?rhs") | 
| 2158 | proof | |
| 2159 | assume ?lhs then show ?rhs | |
| 2160 | by (metis atLeastAtMost_iff image_eqI order_refl path_component_of_def path_connectedin_path_image zero_le_one) | |
| 2161 | next | |
| 2162 | assume ?rhs then show ?lhs | |
| 2163 | by (metis path_component_of_def path_connectedin) | |
| 2164 | qed | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2165 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2166 | lemma path_component_of_set: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2167 | "path_component_of X x y \<longleftrightarrow> (\<exists>g. pathin X g \<and> g 0 = x \<and> g 1 = y)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2168 | by (auto simp: path_component_of_def) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2169 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2170 | lemma path_component_of_subset_topspace: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2171 | "Collect(path_component_of X x) \<subseteq> topspace X" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2172 | using path_component_in_topspace by fastforce | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2173 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2174 | lemma path_component_of_eq_empty: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2175 |    "Collect(path_component_of X x) = {} \<longleftrightarrow> (x \<notin> topspace X)"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2176 | using path_component_in_topspace path_component_of_refl by fastforce | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2177 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2178 | lemma path_connected_space_iff_path_component: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2179 | "path_connected_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. path_component_of X x y)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2180 | by (simp add: path_component_of path_connected_space_subconnected) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2181 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2182 | lemma path_connected_space_imp_path_component_of: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2183 | "\<lbrakk>path_connected_space X; a \<in> topspace X; b \<in> topspace X\<rbrakk> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2184 | \<Longrightarrow> path_component_of X a b" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2185 | by (simp add: path_connected_space_iff_path_component) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2186 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2187 | lemma path_connected_space_path_component_set: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2188 | "path_connected_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. Collect(path_component_of X x) = topspace X)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2189 | using path_component_of_subset_topspace path_connected_space_iff_path_component by fastforce | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2190 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2191 | lemma path_component_of_maximal: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2192 | "\<lbrakk>path_connectedin X s; x \<in> s\<rbrakk> \<Longrightarrow> s \<subseteq> Collect(path_component_of X x)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2193 | using path_component_of by fastforce | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2194 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2195 | lemma path_component_of_equiv: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2196 | "path_component_of X x y \<longleftrightarrow> x \<in> topspace X \<and> y \<in> topspace X \<and> path_component_of X x = path_component_of X y" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2197 | (is "?lhs = ?rhs") | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2198 | proof | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2199 | assume ?lhs | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2200 | then show ?rhs | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2201 | apply (simp add: fun_eq_iff path_component_in_topspace) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2202 | apply (meson path_component_of_sym path_component_of_trans) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2203 | done | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2204 | qed (simp add: path_component_of_refl) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2205 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2206 | lemma path_component_of_disjoint: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2207 | "disjnt (Collect (path_component_of X x)) (Collect (path_component_of X y)) \<longleftrightarrow> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2208 | ~(path_component_of X x y)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2209 | by (force simp: disjnt_def path_component_of_eq_empty path_component_of_equiv) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2210 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2211 | lemma path_component_of_eq: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2212 | "path_component_of X x = path_component_of X y \<longleftrightarrow> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2213 | (x \<notin> topspace X) \<and> (y \<notin> topspace X) \<or> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2214 | x \<in> topspace X \<and> y \<in> topspace X \<and> path_component_of X x y" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2215 | by (metis Collect_empty_eq_bot path_component_of_eq_empty path_component_of_equiv) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2216 | |
| 72256 | 2217 | lemma path_component_of_aux: | 
| 2218 | "path_component_of X x y | |
| 2219 | \<Longrightarrow> path_component_of (subtopology X (Collect (path_component_of X x))) x y" | |
| 2220 | by (meson path_component_of path_component_of_maximal path_connectedin_subtopology) | |
| 2221 | ||
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2222 | lemma path_connectedin_path_component_of: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2223 | "path_connectedin X (Collect (path_component_of X x))" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2224 | proof - | 
| 72256 | 2225 | have "topspace (subtopology X (path_component_of_set X x)) = path_component_of_set X x" | 
| 2226 | by (meson path_component_of_subset_topspace topspace_subtopology_subset) | |
| 2227 | then have "path_connected_space (subtopology X (path_component_of_set X x))" | |
| 2228 | by (metis (full_types) path_component_of_aux mem_Collect_eq path_component_of_equiv path_connected_space_iff_path_component) | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2229 | then show ?thesis | 
| 72256 | 2230 | by (simp add: path_component_of_subset_topspace path_connectedin_def) | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2231 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2232 | |
| 70178 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 2233 | lemma path_connectedin_euclidean [simp]: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 2234 | "path_connectedin euclidean S \<longleftrightarrow> path_connected S" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 2235 | by (auto simp: path_connectedin_def path_connected_space_iff_path_component path_connected_component) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 2236 | |
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 2237 | lemma path_connected_space_euclidean_subtopology [simp]: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 2238 | "path_connected_space(subtopology euclidean S) \<longleftrightarrow> path_connected S" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 2239 | using path_connectedin_topspace by force | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 2240 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2241 | lemma Union_path_components_of: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2242 | "\<Union>(path_components_of X) = topspace X" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2243 | by (auto simp: path_components_of_def path_component_of_equiv) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2244 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2245 | lemma path_components_of_maximal: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2246 | "\<lbrakk>C \<in> path_components_of X; path_connectedin X S; ~disjnt C S\<rbrakk> \<Longrightarrow> S \<subseteq> C" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2247 | apply (auto simp: path_components_of_def path_component_of_equiv) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2248 | using path_component_of_maximal path_connectedin_def apply fastforce | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2249 | by (meson disjnt_subset2 path_component_of_disjoint path_component_of_equiv path_component_of_maximal) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2250 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2251 | lemma pairwise_disjoint_path_components_of: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2252 | "pairwise disjnt (path_components_of X)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2253 | by (auto simp: path_components_of_def pairwise_def path_component_of_disjoint path_component_of_equiv) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2254 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2255 | lemma complement_path_components_of_Union: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2256 | "C \<in> path_components_of X | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2257 |         \<Longrightarrow> topspace X - C = \<Union>(path_components_of X - {C})"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2258 | by (metis Diff_cancel Diff_subset Union_path_components_of cSup_singleton diff_Union_pairwise_disjoint insert_subset pairwise_disjoint_path_components_of) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2259 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2260 | lemma nonempty_path_components_of: | 
| 72256 | 2261 |   assumes "C \<in> path_components_of X" shows "C \<noteq> {}"
 | 
| 2262 | proof - | |
| 2263 | have "C \<in> path_component_of_set X ` topspace X" | |
| 2264 | using assms path_components_of_def by blast | |
| 2265 | then show ?thesis | |
| 2266 | using path_component_of_refl by fastforce | |
| 2267 | qed | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2268 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2269 | lemma path_components_of_subset: "C \<in> path_components_of X \<Longrightarrow> C \<subseteq> topspace X" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2270 | by (auto simp: path_components_of_def path_component_of_equiv) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2271 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2272 | lemma path_connectedin_path_components_of: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2273 | "C \<in> path_components_of X \<Longrightarrow> path_connectedin X C" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2274 | by (auto simp: path_components_of_def path_connectedin_path_component_of) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2275 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2276 | lemma path_component_in_path_components_of: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2277 | "Collect (path_component_of X a) \<in> path_components_of X \<longleftrightarrow> a \<in> topspace X" | 
| 72256 | 2278 | by (metis imageI nonempty_path_components_of path_component_of_eq_empty path_components_of_def) | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2279 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2280 | lemma path_connectedin_Union: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2281 |   assumes \<A>: "\<And>S. S \<in> \<A> \<Longrightarrow> path_connectedin X S" "\<Inter>\<A> \<noteq> {}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2282 | shows "path_connectedin X (\<Union>\<A>)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2283 | proof - | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2284 | obtain a where "\<And>S. S \<in> \<A> \<Longrightarrow> a \<in> S" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2285 | using assms by blast | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2286 | then have "\<And>x. x \<in> topspace (subtopology X (\<Union>\<A>)) \<Longrightarrow> path_component_of (subtopology X (\<Union>\<A>)) a x" | 
| 72256 | 2287 | by simp (meson Union_upper \<A> path_component_of path_connectedin_subtopology) | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2288 | then show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2289 | using \<A> unfolding path_connectedin_def | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2290 | by (metis Sup_le_iff path_component_of_equiv path_connected_space_iff_path_component) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2291 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2292 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2293 | lemma path_connectedin_Un: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2294 |    "\<lbrakk>path_connectedin X S; path_connectedin X T; S \<inter> T \<noteq> {}\<rbrakk>
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2295 | \<Longrightarrow> path_connectedin X (S \<union> T)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2296 |   by (blast intro: path_connectedin_Union [of "{S,T}", simplified])
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2297 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2298 | lemma path_connected_space_iff_components_eq: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2299 | "path_connected_space X \<longleftrightarrow> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2300 | (\<forall>C \<in> path_components_of X. \<forall>C' \<in> path_components_of X. C = C')" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2301 | unfolding path_components_of_def | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2302 | proof (intro iffI ballI) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2303 | assume "\<forall>C \<in> path_component_of_set X ` topspace X. | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2304 | \<forall>C' \<in> path_component_of_set X ` topspace X. C = C'" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2305 | then show "path_connected_space X" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2306 | using path_component_of_refl path_connected_space_iff_path_component by fastforce | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2307 | qed (auto simp: path_connected_space_path_component_set) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2308 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2309 | lemma path_components_of_eq_empty: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2310 |    "path_components_of X = {} \<longleftrightarrow> topspace X = {}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2311 | using Union_path_components_of nonempty_path_components_of by fastforce | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2312 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2313 | lemma path_components_of_empty_space: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2314 |    "topspace X = {} \<Longrightarrow> path_components_of X = {}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2315 | by (simp add: path_components_of_eq_empty) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2316 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2317 | lemma path_components_of_subset_singleton: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2318 |   "path_components_of X \<subseteq> {S} \<longleftrightarrow>
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2319 |         path_connected_space X \<and> (topspace X = {} \<or> topspace X = S)"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2320 | proof (cases "topspace X = {}")
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2321 | case True | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2322 | then show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2323 | by (auto simp: path_components_of_empty_space path_connected_space_topspace_empty) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2324 | next | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2325 | case False | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2326 |   have "(path_components_of X = {S}) \<longleftrightarrow> (path_connected_space X \<and> topspace X = S)"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2327 | proof (intro iffI conjI) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2328 |     assume L: "path_components_of X = {S}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2329 | then show "path_connected_space X" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2330 | by (simp add: path_connected_space_iff_components_eq) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2331 | show "topspace X = S" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2332 | by (metis L ccpo_Sup_singleton [of S] Union_path_components_of) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2333 | next | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2334 | assume R: "path_connected_space X \<and> topspace X = S" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2335 |     then show "path_components_of X = {S}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2336 | using ccpo_Sup_singleton [of S] | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2337 | by (metis False all_not_in_conv insert_iff mk_disjoint_insert path_component_in_path_components_of path_connected_space_iff_components_eq path_connected_space_path_component_set) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2338 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2339 | with False show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2340 | by (simp add: path_components_of_eq_empty subset_singleton_iff) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2341 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2342 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2343 | lemma path_connected_space_iff_components_subset_singleton: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2344 |    "path_connected_space X \<longleftrightarrow> (\<exists>a. path_components_of X \<subseteq> {a})"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2345 | by (simp add: path_components_of_subset_singleton) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2346 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2347 | lemma path_components_of_eq_singleton: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2348 |    "path_components_of X = {S} \<longleftrightarrow> path_connected_space X \<and> topspace X \<noteq> {} \<and> S = topspace X"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2349 | by (metis cSup_singleton insert_not_empty path_components_of_subset_singleton subset_singleton_iff) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2350 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2351 | lemma path_components_of_path_connected_space: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2352 |    "path_connected_space X \<Longrightarrow> path_components_of X = (if topspace X = {} then {} else {topspace X})"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2353 | by (simp add: path_components_of_eq_empty path_components_of_eq_singleton) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2354 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2355 | lemma path_component_subset_connected_component_of: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2356 | "path_component_of_set X x \<subseteq> connected_component_of_set X x" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2357 | proof (cases "x \<in> topspace X") | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2358 | case True | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2359 | then show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2360 | by (simp add: connected_component_of_maximal path_component_of_refl path_connectedin_imp_connectedin path_connectedin_path_component_of) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2361 | next | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2362 | case False | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2363 | then show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2364 | using path_component_of_eq_empty by fastforce | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2365 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2366 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2367 | lemma exists_path_component_of_superset: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2368 |   assumes S: "path_connectedin X S" and ne: "topspace X \<noteq> {}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2369 | obtains C where "C \<in> path_components_of X" "S \<subseteq> C" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2370 | proof (cases "S = {}")
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2371 | case True | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2372 | then show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2373 | using ne path_components_of_eq_empty that by fastforce | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2374 | next | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2375 | case False | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2376 | then obtain a where "a \<in> S" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2377 | by blast | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2378 | show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2379 | proof | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2380 | show "Collect (path_component_of X a) \<in> path_components_of X" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2381 | by (meson \<open>a \<in> S\<close> S subsetD path_component_in_path_components_of path_connectedin_subset_topspace) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2382 | show "S \<subseteq> Collect (path_component_of X a)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2383 | by (simp add: S \<open>a \<in> S\<close> path_component_of_maximal) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2384 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2385 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2386 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2387 | lemma path_component_of_eq_overlap: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2388 | "path_component_of X x = path_component_of X y \<longleftrightarrow> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2389 | (x \<notin> topspace X) \<and> (y \<notin> topspace X) \<or> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2390 |       Collect (path_component_of X x) \<inter> Collect (path_component_of X y) \<noteq> {}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2391 | by (metis disjnt_def empty_iff inf_bot_right mem_Collect_eq path_component_of_disjoint path_component_of_eq path_component_of_eq_empty) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2392 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2393 | lemma path_component_of_nonoverlap: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2394 |    "Collect (path_component_of X x) \<inter> Collect (path_component_of X y) = {} \<longleftrightarrow>
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2395 | (x \<notin> topspace X) \<or> (y \<notin> topspace X) \<or> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2396 | path_component_of X x \<noteq> path_component_of X y" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2397 | by (metis inf.idem path_component_of_eq_empty path_component_of_eq_overlap) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2398 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2399 | lemma path_component_of_overlap: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2400 |    "Collect (path_component_of X x) \<inter> Collect (path_component_of X y) \<noteq> {} \<longleftrightarrow>
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2401 | x \<in> topspace X \<and> y \<in> topspace X \<and> path_component_of X x = path_component_of X y" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2402 | by (meson path_component_of_nonoverlap) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2403 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2404 | lemma path_components_of_disjoint: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2405 | "\<lbrakk>C \<in> path_components_of X; C' \<in> path_components_of X\<rbrakk> \<Longrightarrow> disjnt C C' \<longleftrightarrow> C \<noteq> C'" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2406 | by (auto simp: path_components_of_def path_component_of_disjoint path_component_of_equiv) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2407 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2408 | lemma path_components_of_overlap: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2409 |     "\<lbrakk>C \<in> path_components_of X; C' \<in> path_components_of X\<rbrakk> \<Longrightarrow> C \<inter> C' \<noteq> {} \<longleftrightarrow> C = C'"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2410 | by (auto simp: path_components_of_def path_component_of_equiv) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2411 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2412 | lemma path_component_of_unique: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2413 | "\<lbrakk>x \<in> C; path_connectedin X C; \<And>C'. \<lbrakk>x \<in> C'; path_connectedin X C'\<rbrakk> \<Longrightarrow> C' \<subseteq> C\<rbrakk> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2414 | \<Longrightarrow> Collect (path_component_of X x) = C" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2415 | by (meson subsetD eq_iff path_component_of_maximal path_connectedin_path_component_of) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2416 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2417 | lemma path_component_of_discrete_topology [simp]: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2418 |   "Collect (path_component_of (discrete_topology U) x) = (if x \<in> U then {x} else {})"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2419 | proof - | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2420 |   have "\<And>C'. \<lbrakk>x \<in> C'; path_connectedin (discrete_topology U) C'\<rbrakk> \<Longrightarrow> C' \<subseteq> {x}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2421 | by (metis path_connectedin_discrete_topology subsetD singletonD) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2422 |   then have "x \<in> U \<Longrightarrow> Collect (path_component_of (discrete_topology U) x) = {x}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2423 | by (simp add: path_component_of_unique) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2424 | then show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2425 | using path_component_in_topspace by fastforce | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2426 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2427 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2428 | lemma path_component_of_discrete_topology_iff [simp]: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2429 | "path_component_of (discrete_topology U) x y \<longleftrightarrow> x \<in> U \<and> y=x" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2430 | by (metis empty_iff insertI1 mem_Collect_eq path_component_of_discrete_topology singletonD) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2431 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2432 | lemma path_components_of_discrete_topology [simp]: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2433 |    "path_components_of (discrete_topology U) = (\<lambda>x. {x}) ` U"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2434 | by (auto simp: path_components_of_def image_def fun_eq_iff) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2435 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2436 | lemma homeomorphic_map_path_component_of: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2437 | assumes f: "homeomorphic_map X Y f" and x: "x \<in> topspace X" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2438 | shows "Collect (path_component_of Y (f x)) = f ` Collect(path_component_of X x)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2439 | proof - | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2440 | obtain g where g: "homeomorphic_maps X Y f g" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2441 | using f homeomorphic_map_maps by blast | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2442 | show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2443 | proof | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2444 | have "Collect (path_component_of Y (f x)) \<subseteq> topspace Y" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2445 | by (simp add: path_component_of_subset_topspace) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2446 | moreover have "g ` Collect(path_component_of Y (f x)) \<subseteq> Collect (path_component_of X (g (f x)))" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2447 | using g x unfolding homeomorphic_maps_def | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2448 | by (metis f homeomorphic_imp_surjective_map imageI mem_Collect_eq path_component_of_maximal path_component_of_refl path_connectedin_continuous_map_image path_connectedin_path_component_of) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2449 | ultimately show "Collect (path_component_of Y (f x)) \<subseteq> f ` Collect (path_component_of X x)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2450 | using g x unfolding homeomorphic_maps_def continuous_map_def image_iff subset_iff | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2451 | by metis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2452 | show "f ` Collect (path_component_of X x) \<subseteq> Collect (path_component_of Y (f x))" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2453 | proof (rule path_component_of_maximal) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2454 | show "path_connectedin Y (f ` Collect (path_component_of X x))" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2455 | by (meson f homeomorphic_map_path_connectedness_eq path_connectedin_path_component_of) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2456 | qed (simp add: path_component_of_refl x) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2457 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2458 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2459 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2460 | lemma homeomorphic_map_path_components_of: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2461 | assumes "homeomorphic_map X Y f" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2462 | shows "path_components_of Y = (image f) ` (path_components_of X)" | 
| 72256 | 2463 | (is "?lhs = ?rhs") | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2464 | unfolding path_components_of_def homeomorphic_imp_surjective_map [OF assms, symmetric] | 
| 72256 | 2465 | using assms homeomorphic_map_path_component_of by fastforce | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2466 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2467 | |
| 60420 | 2468 | subsection \<open>Sphere is path-connected\<close> | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36583diff
changeset | 2469 | |
| 36583 | 2470 | lemma path_connected_punctured_universe: | 
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 2471 |   assumes "2 \<le> DIM('a::euclidean_space)"
 | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2472 |   shows "path_connected (- {a::'a})"
 | 
| 49653 | 2473 | proof - | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49654diff
changeset | 2474 |   let ?A = "{x::'a. \<exists>i\<in>Basis. x \<bullet> i < a \<bullet> i}"
 | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49654diff
changeset | 2475 |   let ?B = "{x::'a. \<exists>i\<in>Basis. a \<bullet> i < x \<bullet> i}"
 | 
| 36583 | 2476 | |
| 49653 | 2477 | have A: "path_connected ?A" | 
| 2478 | unfolding Collect_bex_eq | |
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 2479 | proof (rule path_connected_UNION) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49654diff
changeset | 2480 | fix i :: 'a | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49654diff
changeset | 2481 | assume "i \<in> Basis" | 
| 53640 | 2482 |     then show "(\<Sum>i\<in>Basis. (a \<bullet> i - 1)*\<^sub>R i) \<in> {x::'a. x \<bullet> i < a \<bullet> i}"
 | 
| 2483 | by simp | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49654diff
changeset | 2484 |     show "path_connected {x. x \<bullet> i < a \<bullet> i}"
 | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49654diff
changeset | 2485 | using convex_imp_path_connected [OF convex_halfspace_lt, of i "a \<bullet> i"] | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49654diff
changeset | 2486 | by (simp add: inner_commute) | 
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 2487 | qed | 
| 53640 | 2488 | have B: "path_connected ?B" | 
| 2489 | unfolding Collect_bex_eq | |
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 2490 | proof (rule path_connected_UNION) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49654diff
changeset | 2491 | fix i :: 'a | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49654diff
changeset | 2492 | assume "i \<in> Basis" | 
| 53640 | 2493 |     then show "(\<Sum>i\<in>Basis. (a \<bullet> i + 1) *\<^sub>R i) \<in> {x::'a. a \<bullet> i < x \<bullet> i}"
 | 
| 2494 | by simp | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49654diff
changeset | 2495 |     show "path_connected {x. a \<bullet> i < x \<bullet> i}"
 | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49654diff
changeset | 2496 | using convex_imp_path_connected [OF convex_halfspace_gt, of "a \<bullet> i" i] | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49654diff
changeset | 2497 | by (simp add: inner_commute) | 
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 2498 | qed | 
| 53640 | 2499 | obtain S :: "'a set" where "S \<subseteq> Basis" and "card S = Suc (Suc 0)" | 
| 2500 | using ex_card[OF assms] | |
| 2501 | by auto | |
| 2502 | then obtain b0 b1 :: 'a where "b0 \<in> Basis" and "b1 \<in> Basis" and "b0 \<noteq> b1" | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49654diff
changeset | 2503 | unfolding card_Suc_eq by auto | 
| 53640 | 2504 | then have "a + b0 - b1 \<in> ?A \<inter> ?B" | 
| 2505 | by (auto simp: inner_simps inner_Basis) | |
| 2506 |   then have "?A \<inter> ?B \<noteq> {}"
 | |
| 2507 | by fast | |
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 2508 | with A B have "path_connected (?A \<union> ?B)" | 
| 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 2509 | by (rule path_connected_Un) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49654diff
changeset | 2510 |   also have "?A \<union> ?B = {x. \<exists>i\<in>Basis. x \<bullet> i \<noteq> a \<bullet> i}"
 | 
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 2511 | unfolding neq_iff bex_disj_distrib Collect_disj_eq .. | 
| 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 2512 |   also have "\<dots> = {x. x \<noteq> a}"
 | 
| 53640 | 2513 | unfolding euclidean_eq_iff [where 'a='a] | 
| 2514 | by (simp add: Bex_def) | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2515 |   also have "\<dots> = - {a}"
 | 
| 53640 | 2516 | by auto | 
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 2517 | finally show ?thesis . | 
| 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 2518 | qed | 
| 36583 | 2519 | |
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2520 | corollary connected_punctured_universe: | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2521 |   "2 \<le> DIM('N::euclidean_space) \<Longrightarrow> connected(- {a::'N})"
 | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2522 | by (simp add: path_connected_punctured_universe path_connected_imp_connected) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2523 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68532diff
changeset | 2524 | proposition path_connected_sphere: | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2525 | fixes a :: "'a :: euclidean_space" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2526 |   assumes "2 \<le> DIM('a)"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2527 | shows "path_connected(sphere a r)" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68532diff
changeset | 2528 | proof (cases r "0::real" rule: linorder_cases) | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2529 | case less | 
| 53640 | 2530 | then show ?thesis | 
| 71172 | 2531 | by (simp) | 
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 2532 | next | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2533 | case equal | 
| 53640 | 2534 | then show ?thesis | 
| 71172 | 2535 | by (simp) | 
| 37674 
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
 huffman parents: 
37489diff
changeset | 2536 | next | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2537 | case greater | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2538 |   then have eq: "(sphere (0::'a) r) = (\<lambda>x. (r / norm x) *\<^sub>R x) ` (- {0::'a})"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2539 | by (force simp: image_iff split: if_split_asm) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2540 |   have "continuous_on (- {0::'a}) (\<lambda>x. (r / norm x) *\<^sub>R x)"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2541 | by (intro continuous_intros) auto | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2542 |   then have "path_connected ((\<lambda>x. (r / norm x) *\<^sub>R x) ` (- {0::'a}))"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2543 | by (intro path_connected_continuous_image path_connected_punctured_universe assms) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2544 | with eq have "path_connected (sphere (0::'a) r)" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2545 | by auto | 
| 67399 | 2546 | then have "path_connected((+) a ` (sphere (0::'a) r))" | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2547 | by (simp add: path_connected_translation) | 
| 53640 | 2548 | then show ?thesis | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2549 | by (metis add.right_neutral sphere_translation) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2550 | qed | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2551 | |
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2552 | lemma connected_sphere: | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2553 | fixes a :: "'a :: euclidean_space" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2554 |     assumes "2 \<le> DIM('a)"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2555 | shows "connected(sphere a r)" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2556 | using path_connected_sphere [OF assms] | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2557 | by (simp add: path_connected_imp_connected) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2558 | |
| 36583 | 2559 | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2560 | corollary path_connected_complement_bounded_convex: | 
| 72256 | 2561 | fixes S :: "'a :: euclidean_space set" | 
| 2562 |     assumes "bounded S" "convex S" and 2: "2 \<le> DIM('a)"
 | |
| 2563 | shows "path_connected (- S)" | |
| 2564 | proof (cases "S = {}")
 | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2565 | case True then show ?thesis | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2566 | using convex_imp_path_connected by auto | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2567 | next | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2568 | case False | 
| 72256 | 2569 | then obtain a where "a \<in> S" by auto | 
| 2570 | have \<section> [rule_format]: "\<forall>y\<in>S. \<forall>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> (1 - u) *\<^sub>R a + u *\<^sub>R y \<in> S" | |
| 2571 | using \<open>convex S\<close> \<open>a \<in> S\<close> by (simp add: convex_alt) | |
| 2572 |   { fix x y assume "x \<notin> S" "y \<notin> S"
 | |
| 2573 | then have "x \<noteq> a" "y \<noteq> a" using \<open>a \<in> S\<close> by auto | |
| 2574 | then have bxy: "bounded(insert x (insert y S))" | |
| 2575 | by (simp add: \<open>bounded S\<close>) | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2576 | then obtain B::real where B: "0 < B" and Bx: "norm (a - x) < B" and By: "norm (a - y) < B" | 
| 72256 | 2577 | and "S \<subseteq> ball a B" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2578 | using bounded_subset_ballD [OF bxy, of a] by (auto simp: dist_norm) | 
| 63040 | 2579 | define C where "C = B / norm(x - a)" | 
| 72256 | 2580 | let ?Cxa = "a + C *\<^sub>R (x - a)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2581 |     { fix u
 | 
| 72256 | 2582 | assume u: "(1 - u) *\<^sub>R x + u *\<^sub>R ?Cxa \<in> S" and "0 \<le> u" "u \<le> 1" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2583 | have CC: "1 \<le> 1 + (C - 1) * u" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 2584 | using \<open>x \<noteq> a\<close> \<open>0 \<le> u\<close> Bx | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 2585 | by (auto simp add: C_def norm_minus_commute) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2586 | have *: "\<And>v. (1 - u) *\<^sub>R x + u *\<^sub>R (a + v *\<^sub>R (x - a)) = a + (1 + (v - 1) * u) *\<^sub>R (x - a)" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2587 | by (simp add: algebra_simps) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2588 | have "a + ((1 / (1 + C * u - u)) *\<^sub>R x + ((u / (1 + C * u - u)) *\<^sub>R a + (C * u / (1 + C * u - u)) *\<^sub>R x)) = | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2589 | (1 + (u / (1 + C * u - u))) *\<^sub>R a + ((1 / (1 + C * u - u)) + (C * u / (1 + C * u - u))) *\<^sub>R x" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2590 | by (simp add: algebra_simps) | 
| 68096 | 2591 | also have "\<dots> = (1 + (u / (1 + C * u - u))) *\<^sub>R a + (1 + (u / (1 + C * u - u))) *\<^sub>R x" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2592 | using CC by (simp add: field_simps) | 
| 68096 | 2593 | also have "\<dots> = x + (1 + (u / (1 + C * u - u))) *\<^sub>R a + (u / (1 + C * u - u)) *\<^sub>R x" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2594 | by (simp add: algebra_simps) | 
| 68096 | 2595 | also have "\<dots> = x + ((1 / (1 + C * u - u)) *\<^sub>R a + | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2596 | ((u / (1 + C * u - u)) *\<^sub>R x + (C * u / (1 + C * u - u)) *\<^sub>R a))" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2597 | using CC by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2598 | finally have xeq: "(1 - 1 / (1 + (C - 1) * u)) *\<^sub>R a + (1 / (1 + (C - 1) * u)) *\<^sub>R (a + (1 + (C - 1) * u) *\<^sub>R (x - a)) = x" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2599 | by (simp add: algebra_simps) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2600 | have False | 
| 72256 | 2601 | using \<section> [of "a + (1 + (C - 1) * u) *\<^sub>R (x - a)" "1 / (1 + (C - 1) * u)"] | 
| 2602 | using u \<open>x \<noteq> a\<close> \<open>x \<notin> S\<close> \<open>0 \<le> u\<close> CC | |
| 2603 | by (auto simp: xeq *) | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2604 | } | 
| 72256 | 2605 | then have pcx: "path_component (- S) x ?Cxa" | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2606 | by (force simp: closed_segment_def intro!: path_component_linepath) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67399diff
changeset | 2607 | define D where "D = B / norm(y - a)" \<comment> \<open>massive duplication with the proof above\<close> | 
| 72256 | 2608 | let ?Dya = "a + D *\<^sub>R (y - a)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2609 |     { fix u
 | 
| 72256 | 2610 | assume u: "(1 - u) *\<^sub>R y + u *\<^sub>R ?Dya \<in> S" and "0 \<le> u" "u \<le> 1" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2611 | have DD: "1 \<le> 1 + (D - 1) * u" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 2612 | using \<open>y \<noteq> a\<close> \<open>0 \<le> u\<close> By | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 2613 | by (auto simp add: D_def norm_minus_commute) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2614 | have *: "\<And>v. (1 - u) *\<^sub>R y + u *\<^sub>R (a + v *\<^sub>R (y - a)) = a + (1 + (v - 1) * u) *\<^sub>R (y - a)" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2615 | by (simp add: algebra_simps) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2616 | have "a + ((1 / (1 + D * u - u)) *\<^sub>R y + ((u / (1 + D * u - u)) *\<^sub>R a + (D * u / (1 + D * u - u)) *\<^sub>R y)) = | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2617 | (1 + (u / (1 + D * u - u))) *\<^sub>R a + ((1 / (1 + D * u - u)) + (D * u / (1 + D * u - u))) *\<^sub>R y" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2618 | by (simp add: algebra_simps) | 
| 68096 | 2619 | also have "\<dots> = (1 + (u / (1 + D * u - u))) *\<^sub>R a + (1 + (u / (1 + D * u - u))) *\<^sub>R y" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2620 | using DD by (simp add: field_simps) | 
| 68096 | 2621 | also have "\<dots> = y + (1 + (u / (1 + D * u - u))) *\<^sub>R a + (u / (1 + D * u - u)) *\<^sub>R y" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2622 | by (simp add: algebra_simps) | 
| 68096 | 2623 | also have "\<dots> = y + ((1 / (1 + D * u - u)) *\<^sub>R a + | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2624 | ((u / (1 + D * u - u)) *\<^sub>R y + (D * u / (1 + D * u - u)) *\<^sub>R a))" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2625 | using DD by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2626 | finally have xeq: "(1 - 1 / (1 + (D - 1) * u)) *\<^sub>R a + (1 / (1 + (D - 1) * u)) *\<^sub>R (a + (1 + (D - 1) * u) *\<^sub>R (y - a)) = y" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2627 | by (simp add: algebra_simps) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2628 | have False | 
| 72256 | 2629 | using \<section> [of "a + (1 + (D - 1) * u) *\<^sub>R (y - a)" "1 / (1 + (D - 1) * u)"] | 
| 2630 | using u \<open>y \<noteq> a\<close> \<open>y \<notin> S\<close> \<open>0 \<le> u\<close> DD | |
| 2631 | by (auto simp: xeq *) | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2632 | } | 
| 72256 | 2633 | then have pdy: "path_component (- S) y ?Dya" | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69712diff
changeset | 2634 | by (force simp: closed_segment_def intro!: path_component_linepath) | 
| 72256 | 2635 | have pyx: "path_component (- S) ?Dya ?Cxa" | 
| 2636 | proof (rule path_component_of_subset) | |
| 2637 | show "sphere a B \<subseteq> - S" | |
| 2638 | using \<open>S \<subseteq> ball a B\<close> by (force simp: ball_def dist_norm norm_minus_commute) | |
| 2639 | have aB: "?Dya \<in> sphere a B" "?Cxa \<in> sphere a B" | |
| 2640 | using \<open>x \<noteq> a\<close> using \<open>y \<noteq> a\<close> B by (auto simp: dist_norm C_def D_def) | |
| 2641 | then show "path_component (sphere a B) ?Dya ?Cxa" | |
| 2642 | using path_connected_sphere [OF 2] path_connected_component by blast | |
| 2643 | qed | |
| 2644 | have "path_component (- S) x y" | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2645 | by (metis path_component_trans path_component_sym pcx pdy pyx) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2646 | } | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2647 | then show ?thesis | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2648 | by (auto simp: path_connected_component) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2649 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2650 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2651 | lemma connected_complement_bounded_convex: | 
| 72256 | 2652 | fixes S :: "'a :: euclidean_space set" | 
| 2653 |     assumes "bounded S" "convex S" "2 \<le> DIM('a)"
 | |
| 2654 | shows "connected (- S)" | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2655 | using path_connected_complement_bounded_convex [OF assms] path_connected_imp_connected by blast | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2656 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2657 | lemma connected_diff_ball: | 
| 72256 | 2658 | fixes S :: "'a :: euclidean_space set" | 
| 2659 |     assumes "connected S" "cball a r \<subseteq> S" "2 \<le> DIM('a)"
 | |
| 2660 | shows "connected (S - ball a r)" | |
| 2661 | proof (rule connected_diff_open_from_closed [OF ball_subset_cball]) | |
| 2662 | show "connected (cball a r - ball a r)" | |
| 2663 | using assms connected_sphere by (auto simp: cball_diff_eq_sphere) | |
| 2664 | qed (auto simp: assms dist_norm) | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2665 | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2666 | proposition connected_open_delete: | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2667 |   assumes "open S" "connected S" and 2: "2 \<le> DIM('N::euclidean_space)"
 | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2668 |     shows "connected(S - {a::'N})"
 | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2669 | proof (cases "a \<in> S") | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2670 | case True | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2671 | with \<open>open S\<close> obtain \<epsilon> where "\<epsilon> > 0" and \<epsilon>: "cball a \<epsilon> \<subseteq> S" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2672 | using open_contains_cball_eq by blast | 
| 72256 | 2673 | define b where "b \<equiv> a + \<epsilon> *\<^sub>R (SOME i. i \<in> Basis)" | 
| 2674 | have "dist a b = \<epsilon>" | |
| 2675 | by (simp add: b_def dist_norm SOME_Basis \<open>0 < \<epsilon>\<close> less_imp_le) | |
| 2676 |   with \<epsilon> have "b \<in> \<Inter>{S - ball a r |r. 0 < r \<and> r < \<epsilon>}"
 | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2677 | by auto | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2678 |   then have nonemp: "(\<Inter>{S - ball a r |r. 0 < r \<and> r < \<epsilon>}) = {} \<Longrightarrow> False"
 | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2679 | by auto | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2680 | have con: "\<And>r. r < \<epsilon> \<Longrightarrow> connected (S - ball a r)" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2681 | using \<epsilon> by (force intro: connected_diff_ball [OF \<open>connected S\<close> _ 2]) | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2682 |   have "x \<in> \<Union>{S - ball a r |r. 0 < r \<and> r < \<epsilon>}" if "x \<in> S - {a}" for x
 | 
| 72256 | 2683 | using that \<open>0 < \<epsilon>\<close> | 
| 2684 | by (intro UnionI [of "S - ball a (min \<epsilon> (dist a x) / 2)"]) auto | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2685 |   then have "S - {a} = \<Union>{S - ball a r | r. 0 < r \<and> r < \<epsilon>}"
 | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2686 | by auto | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2687 | then show ?thesis | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2688 | by (auto intro: connected_Union con dest!: nonemp) | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2689 | next | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2690 | case False then show ?thesis | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2691 | by (simp add: \<open>connected S\<close>) | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2692 | qed | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2693 | |
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2694 | corollary path_connected_open_delete: | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2695 |   assumes "open S" "connected S" and 2: "2 \<le> DIM('N::euclidean_space)"
 | 
| 72256 | 2696 |   shows "path_connected(S - {a::'N})"
 | 
| 2697 | by (simp add: assms connected_open_delete connected_open_path_connected open_delete) | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2698 | |
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2699 | corollary path_connected_punctured_ball: | 
| 72256 | 2700 |   "2 \<le> DIM('N::euclidean_space) \<Longrightarrow> path_connected(ball a r - {a::'N})"
 | 
| 2701 | by (simp add: path_connected_open_delete) | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2702 | |
| 63151 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2703 | corollary connected_punctured_ball: | 
| 72256 | 2704 |   "2 \<le> DIM('N::euclidean_space) \<Longrightarrow> connected(ball a r - {a::'N})"
 | 
| 2705 | by (simp add: connected_open_delete) | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 2706 | |
| 63151 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2707 | corollary connected_open_delete_finite: | 
| 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2708 | fixes S T::"'a::euclidean_space set" | 
| 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2709 |   assumes S: "open S" "connected S" and 2: "2 \<le> DIM('a)" and "finite T"
 | 
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 2710 | shows "connected(S - T)" | 
| 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63540diff
changeset | 2711 | using \<open>finite T\<close> S | 
| 63151 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2712 | proof (induct T) | 
| 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2713 | case empty | 
| 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2714 | show ?case using \<open>connected S\<close> by simp | 
| 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2715 | next | 
| 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2716 | case (insert x F) | 
| 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2717 | then have "connected (S-F)" by auto | 
| 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2718 | moreover have "open (S - F)" using finite_imp_closed[OF \<open>finite F\<close>] \<open>open S\<close> by auto | 
| 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2719 |   ultimately have "connected (S - F - {x})" using connected_open_delete[OF _ _ 2] by auto
 | 
| 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2720 | thus ?case by (metis Diff_insert) | 
| 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2721 | qed | 
| 
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
 paulson <lp15@cam.ac.uk> parents: 
63126diff
changeset | 2722 | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2723 | lemma sphere_1D_doubleton_zero: | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2724 |   assumes 1: "DIM('a) = 1" and "r > 0"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2725 | obtains x y::"'a::euclidean_space" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2726 |     where "sphere 0 r = {x,y} \<and> dist x y = 2*r"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2727 | proof - | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2728 |   obtain b::'a where b: "Basis = {b}"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2729 | using 1 card_1_singletonE by blast | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2730 | show ?thesis | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2731 | proof (intro that conjI) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2732 | have "x = norm x *\<^sub>R b \<or> x = - norm x *\<^sub>R b" if "r = norm x" for x | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2733 | proof - | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2734 | have xb: "(x \<bullet> b) *\<^sub>R b = x" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2735 | using euclidean_representation [of x, unfolded b] by force | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2736 | then have "norm ((x \<bullet> b) *\<^sub>R b) = norm x" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2737 | by simp | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2738 | with b have "\<bar>x \<bullet> b\<bar> = norm x" | 
| 68310 | 2739 | using norm_Basis by (simp add: b) | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2740 | with xb show ?thesis | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73795diff
changeset | 2741 | by (metis (mono_tags, opaque_lifting) abs_eq_iff abs_norm_cancel) | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2742 | qed | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2743 |     with \<open>r > 0\<close> b show "sphere 0 r = {r *\<^sub>R b, - r *\<^sub>R b}"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2744 | by (force simp: sphere_def dist_norm) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2745 | have "dist (r *\<^sub>R b) (- r *\<^sub>R b) = norm (r *\<^sub>R b + r *\<^sub>R b)" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2746 | by (simp add: dist_norm) | 
| 68096 | 2747 | also have "\<dots> = norm ((2*r) *\<^sub>R b)" | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2748 | by (metis mult_2 scaleR_add_left) | 
| 68096 | 2749 | also have "\<dots> = 2*r" | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2750 | using \<open>r > 0\<close> b norm_Basis by fastforce | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2751 | finally show "dist (r *\<^sub>R b) (- r *\<^sub>R b) = 2*r" . | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2752 | qed | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2753 | qed | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2754 | |
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2755 | lemma sphere_1D_doubleton: | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2756 | fixes a :: "'a :: euclidean_space" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2757 |   assumes "DIM('a) = 1" and "r > 0"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2758 |   obtains x y where "sphere a r = {x,y} \<and> dist x y = 2*r"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2759 | proof - | 
| 67399 | 2760 | have "sphere a r = (+) a ` sphere 0 r" | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2761 | by (metis add.right_neutral sphere_translation) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2762 | then show ?thesis | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2763 | using sphere_1D_doubleton_zero [OF assms] | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2764 | by (metis (mono_tags, lifting) dist_add_cancel image_empty image_insert that) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2765 | qed | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2766 | |
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2767 | lemma psubset_sphere_Compl_connected: | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2768 | fixes S :: "'a::euclidean_space set" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2769 |   assumes S: "S \<subset> sphere a r" and "0 < r" and 2: "2 \<le> DIM('a)"
 | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2770 | shows "connected(- S)" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2771 | proof - | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2772 | have "S \<subseteq> sphere a r" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2773 | using S by blast | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2774 | obtain b where "dist a b = r" and "b \<notin> S" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2775 | using S mem_sphere by blast | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2776 |   have CS: "- S = {x. dist a x \<le> r \<and> (x \<notin> S)} \<union> {x. r \<le> dist a x \<and> (x \<notin> S)}"
 | 
| 68096 | 2777 | by auto | 
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2778 |   have "{x. dist a x \<le> r \<and> x \<notin> S} \<inter> {x. r \<le> dist a x \<and> x \<notin> S} \<noteq> {}"
 | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2779 | using \<open>b \<notin> S\<close> \<open>dist a b = r\<close> by blast | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2780 |   moreover have "connected {x. dist a x \<le> r \<and> x \<notin> S}"
 | 
| 72256 | 2781 | using assms | 
| 2782 | by (force intro: connected_intermediate_closure [of "ball a r"]) | |
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2783 | moreover | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2784 |   have "connected {x. r \<le> dist a x \<and> x \<notin> S}"
 | 
| 72256 | 2785 | proof (rule connected_intermediate_closure [of "- cball a r"]) | 
| 2786 |     show "{x. r \<le> dist a x \<and> x \<notin> S} \<subseteq> closure (- cball a r)"
 | |
| 2787 | using interior_closure by (force intro: connected_complement_bounded_convex) | |
| 2788 | qed (use assms connected_complement_bounded_convex in auto) | |
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2789 | ultimately show ?thesis | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2790 | by (simp add: CS connected_Un) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2791 | qed | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 2792 | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2793 | |
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2794 | subsection\<open>Every annulus is a connected set\<close> | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2795 | |
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2796 | lemma path_connected_2DIM_I: | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2797 | fixes a :: "'N::euclidean_space" | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2798 |   assumes 2: "2 \<le> DIM('N)" and pc: "path_connected {r. 0 \<le> r \<and> P r}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2799 |   shows "path_connected {x. P(norm(x - a))}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2800 | proof - | 
| 67399 | 2801 |   have "{x. P(norm(x - a))} = (+) a ` {x. P(norm x)}"
 | 
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2802 | by force | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2803 |   moreover have "path_connected {x::'N. P(norm x)}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2804 | proof - | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2805 |     let ?D = "{x. 0 \<le> x \<and> P x} \<times> sphere (0::'N) 1"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2806 | have "x \<in> (\<lambda>z. fst z *\<^sub>R snd z) ` ?D" | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2807 | if "P (norm x)" for x::'N | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2808 | proof (cases "x=0") | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2809 | case True | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2810 | with that show ?thesis | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2811 | apply (simp add: image_iff) | 
| 72256 | 2812 | by (metis (no_types) mem_sphere_0 order_refl vector_choose_size zero_le_one) | 
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2813 | next | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2814 | case False | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2815 | with that show ?thesis | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2816 | by (rule_tac x="(norm x, x /\<^sub>R norm x)" in image_eqI) auto | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2817 | qed | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2818 |     then have *: "{x::'N. P(norm x)} =  (\<lambda>z. fst z *\<^sub>R snd z) ` ?D"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2819 | by auto | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2820 | have "continuous_on ?D (\<lambda>z:: real\<times>'N. fst z *\<^sub>R snd z)" | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2821 | by (intro continuous_intros) | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2822 | moreover have "path_connected ?D" | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2823 | by (metis path_connected_Times [OF pc] path_connected_sphere 2) | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2824 | ultimately show ?thesis | 
| 72256 | 2825 | by (simp add: "*" path_connected_continuous_image) | 
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2826 | qed | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2827 | ultimately show ?thesis | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2828 | using path_connected_translation by metis | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2829 | qed | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2830 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68532diff
changeset | 2831 | proposition path_connected_annulus: | 
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2832 | fixes a :: "'N::euclidean_space" | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2833 |   assumes "2 \<le> DIM('N)"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2834 |   shows "path_connected {x. r1 < norm(x - a) \<and> norm(x - a) < r2}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2835 |         "path_connected {x. r1 < norm(x - a) \<and> norm(x - a) \<le> r2}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2836 |         "path_connected {x. r1 \<le> norm(x - a) \<and> norm(x - a) < r2}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2837 |         "path_connected {x. r1 \<le> norm(x - a) \<and> norm(x - a) \<le> r2}"
 | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68532diff
changeset | 2838 | by (auto simp: is_interval_def intro!: is_interval_convex convex_imp_path_connected path_connected_2DIM_I [OF assms]) | 
| 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68532diff
changeset | 2839 | |
| 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68532diff
changeset | 2840 | proposition connected_annulus: | 
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2841 | fixes a :: "'N::euclidean_space" | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2842 |   assumes "2 \<le> DIM('N::euclidean_space)"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2843 |   shows "connected {x. r1 < norm(x - a) \<and> norm(x - a) < r2}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2844 |         "connected {x. r1 < norm(x - a) \<and> norm(x - a) \<le> r2}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2845 |         "connected {x. r1 \<le> norm(x - a) \<and> norm(x - a) < r2}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 2846 |         "connected {x. r1 \<le> norm(x - a) \<and> norm(x - a) \<le> r2}"
 | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68532diff
changeset | 2847 | by (auto simp: path_connected_annulus [OF assms] path_connected_imp_connected) | 
| 67962 | 2848 | |
| 2849 | ||
| 70136 | 2850 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Relations between components and path components\<close> | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2851 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2852 | lemma open_connected_component: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2853 | fixes S :: "'a::real_normed_vector set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2854 | assumes "open S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2855 | shows "open (connected_component_set S x)" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2856 | proof (clarsimp simp: open_contains_ball) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2857 | fix y | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2858 | assume xy: "connected_component S x y" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2859 | then obtain e where "e>0" "ball y e \<subseteq> S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2860 | using assms connected_component_in openE by blast | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2861 | then show "\<exists>e>0. ball y e \<subseteq> connected_component_set S x" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2862 | by (metis xy centre_in_ball connected_ball connected_component_eq_eq connected_component_in connected_component_maximal) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2863 | qed | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2864 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2865 | corollary open_components: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2866 | fixes S :: "'a::real_normed_vector set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2867 | shows "\<lbrakk>open u; S \<in> components u\<rbrakk> \<Longrightarrow> open S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2868 | by (simp add: components_iff) (metis open_connected_component) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2869 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2870 | lemma in_closure_connected_component: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2871 | fixes S :: "'a::real_normed_vector set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2872 | assumes x: "x \<in> S" and S: "open S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2873 | shows "x \<in> closure (connected_component_set S y) \<longleftrightarrow> x \<in> connected_component_set S y" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2874 | proof - | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2875 |   { assume "x \<in> closure (connected_component_set S y)"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2876 | moreover have "x \<in> connected_component_set S x" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2877 | using x by simp | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2878 | ultimately have "x \<in> connected_component_set S y" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2879 | using S by (meson Compl_disjoint closure_iff_nhds_not_empty connected_component_disjoint disjoint_eq_subset_Compl open_connected_component) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2880 | } | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2881 | then show ?thesis | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2882 | by (auto simp: closure_def) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2883 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2884 | |
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2885 | lemma connected_disjoint_Union_open_pick: | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2886 | assumes "pairwise disjnt B" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2887 |           "\<And>S. S \<in> A \<Longrightarrow> connected S \<and> S \<noteq> {}"
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2888 | "\<And>S. S \<in> B \<Longrightarrow> open S" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2889 | "\<Union>A \<subseteq> \<Union>B" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2890 | "S \<in> A" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2891 |   obtains T where "T \<in> B" "S \<subseteq> T" "S \<inter> \<Union>(B - {T}) = {}"
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2892 | proof - | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2893 |   have "S \<subseteq> \<Union>B" "connected S" "S \<noteq> {}"
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2894 | using assms \<open>S \<in> A\<close> by blast+ | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2895 |   then obtain T where "T \<in> B" "S \<inter> T \<noteq> {}"
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2896 | by (metis Sup_inf_eq_bot_iff inf.absorb_iff2 inf_commute) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2897 | have 1: "open T" by (simp add: \<open>T \<in> B\<close> assms) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2898 |   have 2: "open (\<Union>(B-{T}))" using assms by blast
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2899 |   have 3: "S \<subseteq> T \<union> \<Union>(B - {T})" using \<open>S \<subseteq> \<Union>B\<close> by blast
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2900 |   have "T \<inter> \<Union>(B - {T}) = {}" using \<open>T \<in> B\<close> \<open>pairwise disjnt B\<close>
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2901 | by (auto simp: pairwise_def disjnt_def) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2902 |   then have 4: "T \<inter> \<Union>(B - {T}) \<inter> S = {}" by auto
 | 
| 71244 | 2903 | from connectedD [OF \<open>connected S\<close> 1 2 4 3] | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2904 |   have "S \<inter> \<Union>(B-{T}) = {}"
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2905 |     by (auto simp: Int_commute \<open>S \<inter> T \<noteq> {}\<close>)
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2906 | with \<open>T \<in> B\<close> have "S \<subseteq> T" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2907 | using "3" by auto | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2908 | show ?thesis | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2909 |     using \<open>S \<inter> \<Union>(B - {T}) = {}\<close> \<open>S \<subseteq> T\<close> \<open>T \<in> B\<close> that by auto
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2910 | qed | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2911 | |
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2912 | lemma connected_disjoint_Union_open_subset: | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2913 | assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2914 |       and SA: "\<And>S. S \<in> A \<Longrightarrow> open S \<and> connected S \<and> S \<noteq> {}"
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2915 |       and SB: "\<And>S. S \<in> B \<Longrightarrow> open S \<and> connected S \<and> S \<noteq> {}"
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2916 | and eq [simp]: "\<Union>A = \<Union>B" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2917 | shows "A \<subseteq> B" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2918 | proof | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2919 | fix S | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2920 | assume "S \<in> A" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2921 |   obtain T where "T \<in> B" "S \<subseteq> T" "S \<inter> \<Union>(B - {T}) = {}"
 | 
| 72256 | 2922 | using SA SB \<open>S \<in> A\<close> connected_disjoint_Union_open_pick [OF B, of A] eq order_refl by blast | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2923 |   moreover obtain S' where "S' \<in> A" "T \<subseteq> S'" "T \<inter> \<Union>(A - {S'}) = {}"
 | 
| 72256 | 2924 | using SA SB \<open>T \<in> B\<close> connected_disjoint_Union_open_pick [OF A, of B] eq order_refl by blast | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2925 | ultimately have "S' = S" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2926 | by (metis A Int_subset_iff SA \<open>S \<in> A\<close> disjnt_def inf.orderE pairwise_def) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2927 | with \<open>T \<subseteq> S'\<close> have "T \<subseteq> S" by simp | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2928 | with \<open>S \<subseteq> T\<close> have "S = T" by blast | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2929 | with \<open>T \<in> B\<close> show "S \<in> B" by simp | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2930 | qed | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2931 | |
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2932 | lemma connected_disjoint_Union_open_unique: | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2933 | assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2934 |       and SA: "\<And>S. S \<in> A \<Longrightarrow> open S \<and> connected S \<and> S \<noteq> {}"
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2935 |       and SB: "\<And>S. S \<in> B \<Longrightarrow> open S \<and> connected S \<and> S \<noteq> {}"
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2936 | and eq [simp]: "\<Union>A = \<Union>B" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2937 | shows "A = B" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2938 | by (rule subset_antisym; metis connected_disjoint_Union_open_subset assms) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2939 | |
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2940 | proposition components_open_unique: | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2941 | fixes S :: "'a::real_normed_vector set" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2942 | assumes "pairwise disjnt A" "\<Union>A = S" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2943 |           "\<And>X. X \<in> A \<Longrightarrow> open X \<and> connected X \<and> X \<noteq> {}"
 | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2944 | shows "components S = A" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2945 | proof - | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2946 | have "open S" using assms by blast | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2947 | show ?thesis | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2948 | proof (rule connected_disjoint_Union_open_unique) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2949 | show "disjoint (components S)" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2950 | by (simp add: components_eq disjnt_def pairwise_def) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2951 | qed (use \<open>open S\<close> in \<open>simp_all add: assms open_components in_components_connected in_components_nonempty\<close>) | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2952 | qed | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2953 | |
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63092diff
changeset | 2954 | |
| 70136 | 2955 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Existence of unbounded components\<close> | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2956 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2957 | lemma cobounded_unbounded_component: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2958 | fixes S :: "'a :: euclidean_space set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2959 | assumes "bounded (-S)" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2960 | shows "\<exists>x. x \<in> S \<and> \<not> bounded (connected_component_set S x)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2961 | proof - | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2962 | obtain i::'a where i: "i \<in> Basis" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2963 | using nonempty_Basis by blast | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2964 | obtain B where B: "B>0" "-S \<subseteq> ball 0 B" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2965 | using bounded_subset_ballD [OF assms, of 0] by auto | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2966 | then have *: "\<And>x. B \<le> norm x \<Longrightarrow> x \<in> S" | 
| 68096 | 2967 | by (force simp: ball_def dist_norm) | 
| 69508 | 2968 |   have unbounded_inner: "\<not> bounded {x. inner i x \<ge> B}"
 | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2969 | proof (clarsimp simp: bounded_def dist_norm) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2970 | fix e x | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2971 | show "\<exists>y. B \<le> i \<bullet> y \<and> \<not> norm (x - y) \<le> e" | 
| 72256 | 2972 | using i | 
| 2973 | by (rule_tac x="x + (max B e + 1 + \<bar>i \<bullet> x\<bar>) *\<^sub>R i" in exI) (auto simp: inner_right_distrib) | |
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2974 | qed | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2975 | have \<section>: "\<And>x. B \<le> i \<bullet> x \<Longrightarrow> x \<in> S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2976 | using * Basis_le_norm [OF i] by (metis abs_ge_self inner_commute order_trans) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2977 |   have "{x. B \<le> i \<bullet> x} \<subseteq> connected_component_set S (B *\<^sub>R i)"
 | 
| 72256 | 2978 | by (intro connected_component_maximal) (auto simp: i intro: convex_connected convex_halfspace_ge [of B] \<section>) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2979 | then have "\<not> bounded (connected_component_set S (B *\<^sub>R i))" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2980 | using bounded_subset unbounded_inner by blast | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2981 | moreover have "B *\<^sub>R i \<in> S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2982 | by (rule *) (simp add: norm_Basis [OF i]) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2983 | ultimately show ?thesis | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2984 | by blast | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2985 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2986 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2987 | lemma cobounded_unique_unbounded_component: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2988 | fixes S :: "'a :: euclidean_space set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2989 |     assumes bs: "bounded (-S)" and "2 \<le> DIM('a)"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2990 | and bo: "\<not> bounded(connected_component_set S x)" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2991 | "\<not> bounded(connected_component_set S y)" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2992 | shows "connected_component_set S x = connected_component_set S y" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2993 | proof - | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2994 | obtain i::'a where i: "i \<in> Basis" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2995 | using nonempty_Basis by blast | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2996 | obtain B where B: "B>0" "-S \<subseteq> ball 0 B" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 2997 | using bounded_subset_ballD [OF bs, of 0] by auto | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 2998 | then have *: "\<And>x. B \<le> norm x \<Longrightarrow> x \<in> S" | 
| 68096 | 2999 | by (force simp: ball_def dist_norm) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3000 | obtain x' where x': "connected_component S x x'" "norm x' > B" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3001 | using bo [unfolded bounded_def dist_norm, simplified, rule_format] | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3002 | by (metis diff_zero norm_minus_commute not_less) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3003 | obtain y' where y': "connected_component S y y'" "norm y' > B" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3004 | using bo [unfolded bounded_def dist_norm, simplified, rule_format] | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3005 | by (metis diff_zero norm_minus_commute not_less) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3006 | have x'y': "connected_component S x' y'" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3007 | unfolding connected_component_def | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3008 | proof (intro exI conjI) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3009 | show "connected (- ball 0 B :: 'a set)" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3010 | using assms by (auto intro: connected_complement_bounded_convex) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3011 | qed (use x' y' dist_norm * in auto) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3012 | show ?thesis | 
| 72256 | 3013 | proof (rule connected_component_eq) | 
| 3014 | show "x \<in> connected_component_set S y" | |
| 3015 | using x' y' x'y' | |
| 3016 | by (metis (no_types) connected_component_eq_eq connected_component_in mem_Collect_eq) | |
| 3017 | qed | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3018 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3019 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3020 | lemma cobounded_unbounded_components: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3021 | fixes S :: "'a :: euclidean_space set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3022 | shows "bounded (-S) \<Longrightarrow> \<exists>c. c \<in> components S \<and> \<not>bounded c" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3023 | by (metis cobounded_unbounded_component components_def imageI) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3024 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3025 | lemma cobounded_unique_unbounded_components: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3026 | fixes S :: "'a :: euclidean_space set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3027 |     shows  "\<lbrakk>bounded (- S); c \<in> components S; \<not> bounded c; c' \<in> components S; \<not> bounded c'; 2 \<le> DIM('a)\<rbrakk> \<Longrightarrow> c' = c"
 | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3028 | unfolding components_iff | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3029 | by (metis cobounded_unique_unbounded_component) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3030 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3031 | lemma cobounded_has_bounded_component: | 
| 64122 | 3032 | fixes S :: "'a :: euclidean_space set" | 
| 3033 |   assumes "bounded (- S)" "\<not> connected S" "2 \<le> DIM('a)"
 | |
| 3034 | obtains C where "C \<in> components S" "bounded C" | |
| 3035 | by (meson cobounded_unique_unbounded_components connected_eq_connected_components_eq assms) | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3036 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3037 | |
| 69620 | 3038 | subsection\<open>The \<open>inside\<close> and \<open>outside\<close> of a Set\<close> | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3039 | |
| 70136 | 3040 | text\<^marker>\<open>tag important\<close>\<open>The inside comprises the points in a bounded connected component of the set's complement. | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3041 | The outside comprises the points in unbounded connected component of the complement.\<close> | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3042 | |
| 70136 | 3043 | definition\<^marker>\<open>tag important\<close> inside where | 
| 68096 | 3044 |   "inside S \<equiv> {x. (x \<notin> S) \<and> bounded(connected_component_set ( - S) x)}"
 | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3045 | |
| 70136 | 3046 | definition\<^marker>\<open>tag important\<close> outside where | 
| 69508 | 3047 |   "outside S \<equiv> -S \<inter> {x. \<not> bounded(connected_component_set (- S) x)}"
 | 
| 3048 | ||
| 3049 | lemma outside: "outside S = {x. \<not> bounded(connected_component_set (- S) x)}"
 | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3050 | by (auto simp: outside_def) (metis Compl_iff bounded_empty connected_component_eq_empty) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3051 | |
| 68096 | 3052 | lemma inside_no_overlap [simp]: "inside S \<inter> S = {}"
 | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3053 | by (auto simp: inside_def) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3054 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3055 | lemma outside_no_overlap [simp]: | 
| 68096 | 3056 |    "outside S \<inter> S = {}"
 | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3057 | by (auto simp: outside_def) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3058 | |
| 68096 | 3059 | lemma inside_Int_outside [simp]: "inside S \<inter> outside S = {}"
 | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3060 | by (auto simp: inside_def outside_def) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3061 | |
| 68096 | 3062 | lemma inside_Un_outside [simp]: "inside S \<union> outside S = (- S)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3063 | by (auto simp: inside_def outside_def) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3064 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3065 | lemma inside_eq_outside: | 
| 68096 | 3066 | "inside S = outside S \<longleftrightarrow> S = UNIV" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3067 | by (auto simp: inside_def outside_def) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3068 | |
| 68096 | 3069 | lemma inside_outside: "inside S = (- (S \<union> outside S))" | 
| 3070 | by (force simp: inside_def outside) | |
| 3071 | ||
| 3072 | lemma outside_inside: "outside S = (- (S \<union> inside S))" | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3073 | by (auto simp: inside_outside) (metis IntI equals0D outside_no_overlap) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3074 | |
| 68096 | 3075 | lemma union_with_inside: "S \<union> inside S = - outside S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3076 | by (auto simp: inside_outside) (simp add: outside_inside) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3077 | |
| 68096 | 3078 | lemma union_with_outside: "S \<union> outside S = - inside S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3079 | by (simp add: inside_outside) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3080 | |
| 68096 | 3081 | lemma outside_mono: "S \<subseteq> T \<Longrightarrow> outside T \<subseteq> outside S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3082 | by (auto simp: outside bounded_subset connected_component_mono) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3083 | |
| 68096 | 3084 | lemma inside_mono: "S \<subseteq> T \<Longrightarrow> inside S - T \<subseteq> inside T" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3085 | by (auto simp: inside_def bounded_subset connected_component_mono) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3086 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3087 | lemma segment_bound_lemma: | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3088 | fixes u::real | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3089 | assumes "x \<ge> B" "y \<ge> B" "0 \<le> u" "u \<le> 1" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3090 | shows "(1 - u) * x + u * y \<ge> B" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3091 | proof - | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3092 | obtain dx dy where "dx \<ge> 0" "dy \<ge> 0" "x = B + dx" "y = B + dy" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3093 | using assms by auto (metis add.commute diff_add_cancel) | 
| 61808 | 3094 | with \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> show ?thesis | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3095 | by (simp add: add_increasing2 mult_left_le field_simps) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3096 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3097 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3098 | lemma cobounded_outside: | 
| 68096 | 3099 | fixes S :: "'a :: real_normed_vector set" | 
| 3100 | assumes "bounded S" shows "bounded (- outside S)" | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3101 | proof - | 
| 68096 | 3102 | obtain B where B: "B>0" "S \<subseteq> ball 0 B" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3103 | using bounded_subset_ballD [OF assms, of 0] by auto | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3104 |   { fix x::'a and C::real
 | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3105 | assume Bno: "B \<le> norm x" and C: "0 < C" | 
| 68096 | 3106 | have "\<exists>y. connected_component (- S) x y \<and> norm y > C" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3107 | proof (cases "x = 0") | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3108 | case True with B Bno show ?thesis by force | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3109 | next | 
| 68096 | 3110 | case False | 
| 3111 | have "closed_segment x (((B + C) / norm x) *\<^sub>R x) \<subseteq> - ball 0 B" | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3112 | proof | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3113 | fix w | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3114 | assume "w \<in> closed_segment x (((B + C) / norm x) *\<^sub>R x)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3115 | then obtain u where | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3116 | w: "w = (1 - u + u * (B + C) / norm x) *\<^sub>R x" "0 \<le> u" "u \<le> 1" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3117 | by (auto simp add: closed_segment_def real_vector_class.scaleR_add_left [symmetric]) | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3118 | with False B C have "B \<le> (1 - u) * norm x + u * (B + C)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3119 | using segment_bound_lemma [of B "norm x" "B + C" u] Bno | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3120 | by simp | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3121 | with False B C show "w \<in> - ball 0 B" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3122 | using distrib_right [of _ _ "norm x"] | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3123 | by (simp add: ball_def w not_less) | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3124 | qed | 
| 68096 | 3125 | also have "... \<subseteq> -S" | 
| 3126 | by (simp add: B) | |
| 3127 | finally have "\<exists>T. connected T \<and> T \<subseteq> - S \<and> x \<in> T \<and> ((B + C) / norm x) *\<^sub>R x \<in> T" | |
| 3128 | by (rule_tac x="closed_segment x (((B+C)/norm x) *\<^sub>R x)" in exI) simp | |
| 3129 | with False B | |
| 3130 | show ?thesis | |
| 3131 | by (rule_tac x="((B+C)/norm x) *\<^sub>R x" in exI) (simp add: connected_component_def) | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3132 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3133 | } | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3134 | then show ?thesis | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3135 | apply (simp add: outside_def assms) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3136 | apply (rule bounded_subset [OF bounded_ball [of 0 B]]) | 
| 68096 | 3137 | apply (force simp: dist_norm not_less bounded_pos) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3138 | done | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3139 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3140 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3141 | lemma unbounded_outside: | 
| 68096 | 3142 |     fixes S :: "'a::{real_normed_vector, perfect_space} set"
 | 
| 69508 | 3143 | shows "bounded S \<Longrightarrow> \<not> bounded(outside S)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3144 | using cobounded_imp_unbounded cobounded_outside by blast | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3145 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3146 | lemma bounded_inside: | 
| 68096 | 3147 |     fixes S :: "'a::{real_normed_vector, perfect_space} set"
 | 
| 3148 | shows "bounded S \<Longrightarrow> bounded(inside S)" | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3149 | by (simp add: bounded_Int cobounded_outside inside_outside) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3150 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3151 | lemma connected_outside: | 
| 68096 | 3152 | fixes S :: "'a::euclidean_space set" | 
| 3153 |     assumes "bounded S" "2 \<le> DIM('a)"
 | |
| 3154 | shows "connected(outside S)" | |
| 3155 | apply (clarsimp simp add: connected_iff_connected_component outside) | |
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3156 | apply (rule_tac S="connected_component_set (- S) x" in connected_component_of_subset) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3157 | apply (metis (no_types) assms cobounded_unbounded_component cobounded_unique_unbounded_component connected_component_eq_eq connected_component_idemp double_complement mem_Collect_eq) | 
| 72256 | 3158 | by (simp add: Collect_mono connected_component_eq) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3159 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3160 | lemma outside_connected_component_lt: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3161 |   "outside S = {x. \<forall>B. \<exists>y. B < norm(y) \<and> connected_component (- S) x y}"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3162 | apply (auto simp: outside bounded_def dist_norm) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3163 | apply (metis diff_0 norm_minus_cancel not_less) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3164 | by (metis less_diff_eq norm_minus_commute norm_triangle_ineq2 order.trans pinf(6)) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3165 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3166 | lemma outside_connected_component_le: | 
| 72256 | 3167 |   "outside S = {x. \<forall>B. \<exists>y. B \<le> norm(y) \<and> connected_component (- S) x y}"
 | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3168 | apply (simp add: outside_connected_component_lt Set.set_eq_iff) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3169 | by (meson gt_ex leD le_less_linear less_imp_le order.trans) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3170 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3171 | lemma not_outside_connected_component_lt: | 
| 68096 | 3172 | fixes S :: "'a::euclidean_space set" | 
| 3173 |     assumes S: "bounded S" and "2 \<le> DIM('a)"
 | |
| 69508 | 3174 |       shows "- (outside S) = {x. \<forall>B. \<exists>y. B < norm(y) \<and> \<not> connected_component (- S) x y}"
 | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3175 | proof - | 
| 68096 | 3176 | obtain B::real where B: "0 < B" and Bno: "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> B" | 
| 3177 | using S [simplified bounded_pos] by auto | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3178 |   { fix y::'a and z::'a
 | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3179 | assume yz: "B < norm z" "B < norm y" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3180 | have "connected_component (- cball 0 B) y z" | 
| 72256 | 3181 | using assms yz | 
| 3182 | by (force simp: dist_norm intro: connected_componentI [OF _ subset_refl] connected_complement_bounded_convex) | |
| 68096 | 3183 | then have "connected_component (- S) y z" | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3184 | by (metis connected_component_of_subset Bno Compl_anti_mono mem_cball_0 subset_iff) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3185 | } note cyz = this | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3186 | show ?thesis | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3187 | apply (auto simp: outside bounded_pos) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3188 | apply (metis Compl_iff bounded_iff cobounded_imp_unbounded mem_Collect_eq not_le) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3189 | by (metis B connected_component_trans cyz not_le) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3190 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3191 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3192 | lemma not_outside_connected_component_le: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3193 | fixes S :: "'a::euclidean_space set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3194 |   assumes S: "bounded S"  "2 \<le> DIM('a)"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3195 |   shows "- (outside S) = {x. \<forall>B. \<exists>y. B \<le> norm(y) \<and> \<not> connected_component (- S) x y}"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3196 | apply (auto intro: less_imp_le simp: not_outside_connected_component_lt [OF assms]) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3197 | by (meson gt_ex less_le_trans) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3198 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3199 | lemma inside_connected_component_lt: | 
| 68096 | 3200 | fixes S :: "'a::euclidean_space set" | 
| 3201 |     assumes S: "bounded S"  "2 \<le> DIM('a)"
 | |
| 69508 | 3202 |       shows "inside S = {x. (x \<notin> S) \<and> (\<forall>B. \<exists>y. B < norm(y) \<and> \<not> connected_component (- S) x y)}"
 | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3203 | by (auto simp: inside_outside not_outside_connected_component_lt [OF assms]) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3204 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3205 | lemma inside_connected_component_le: | 
| 68096 | 3206 | fixes S :: "'a::euclidean_space set" | 
| 3207 |     assumes S: "bounded S"  "2 \<le> DIM('a)"
 | |
| 69508 | 3208 |       shows "inside S = {x. (x \<notin> S) \<and> (\<forall>B. \<exists>y. B \<le> norm(y) \<and> \<not> connected_component (- S) x y)}"
 | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3209 | by (auto simp: inside_outside not_outside_connected_component_le [OF assms]) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3210 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3211 | lemma inside_subset: | 
| 69508 | 3212 | assumes "connected U" and "\<not> bounded U" and "T \<union> U = - S" | 
| 68096 | 3213 | shows "inside S \<subseteq> T" | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3214 | apply (auto simp: inside_def) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3215 | by (metis bounded_subset [of "connected_component_set (- S) _"] connected_component_maximal | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3216 | Compl_iff Un_iff assms subsetI) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3217 | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3218 | lemma frontier_not_empty: | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3219 | fixes S :: "'a :: real_normed_vector set" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3220 |   shows "\<lbrakk>S \<noteq> {}; S \<noteq> UNIV\<rbrakk> \<Longrightarrow> frontier S \<noteq> {}"
 | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3221 | using connected_Int_frontier [of UNIV S] by auto | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3222 | |
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3223 | lemma frontier_eq_empty: | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3224 | fixes S :: "'a :: real_normed_vector set" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3225 |   shows "frontier S = {} \<longleftrightarrow> S = {} \<or> S = UNIV"
 | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3226 | using frontier_UNIV frontier_empty frontier_not_empty by blast | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3227 | |
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3228 | lemma frontier_of_connected_component_subset: | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3229 | fixes S :: "'a::real_normed_vector set" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3230 | shows "frontier(connected_component_set S x) \<subseteq> frontier S" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3231 | proof - | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3232 |   { fix y
 | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3233 | assume y1: "y \<in> closure (connected_component_set S x)" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3234 | and y2: "y \<notin> interior (connected_component_set S x)" | 
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3235 | have "y \<in> closure S" | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3236 | using y1 closure_mono connected_component_subset by blast | 
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3237 | moreover have "z \<in> interior (connected_component_set S x)" | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3238 | if "0 < e" "ball y e \<subseteq> interior S" "dist y z < e" for e z | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3239 | proof - | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3240 | have "ball y e \<subseteq> connected_component_set S y" | 
| 72256 | 3241 | using connected_component_maximal that interior_subset | 
| 3242 | by (metis centre_in_ball connected_ball subset_trans) | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3243 | then show ?thesis | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3244 | using y1 apply (simp add: closure_approachable open_contains_ball_eq [OF open_interior]) | 
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3245 | by (metis connected_component_eq dist_commute mem_Collect_eq mem_ball mem_interior subsetD \<open>0 < e\<close> y2) | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3246 | qed | 
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3247 | then have "y \<notin> interior S" | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3248 | using y2 by (force simp: open_contains_ball_eq [OF open_interior]) | 
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3249 | ultimately have "y \<in> frontier S" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3250 | by (auto simp: frontier_def) | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3251 | } | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3252 | then show ?thesis by (auto simp: frontier_def) | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3253 | qed | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62087diff
changeset | 3254 | |
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3255 | lemma frontier_Union_subset_closure: | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3256 | fixes F :: "'a::real_normed_vector set set" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3257 | shows "frontier(\<Union>F) \<subseteq> closure(\<Union>t \<in> F. frontier t)" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3258 | proof - | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3259 | have "\<exists>y\<in>F. \<exists>y\<in>frontier y. dist y x < e" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3260 | if "T \<in> F" "y \<in> T" "dist y x < e" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3261 | "x \<notin> interior (\<Union>F)" "0 < e" for x y e T | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3262 | proof (cases "x \<in> T") | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3263 | case True with that show ?thesis | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3264 | by (metis Diff_iff Sup_upper closure_subset contra_subsetD dist_self frontier_def interior_mono) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3265 | next | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3266 | case False | 
| 72256 | 3267 |     have 1: "closed_segment x y \<inter> T \<noteq> {}" 
 | 
| 3268 | using \<open>y \<in> T\<close> by blast | |
| 62620 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3269 |     have 2: "closed_segment x y - T \<noteq> {}"
 | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3270 | using False by blast | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3271 | obtain c where "c \<in> closed_segment x y" "c \<in> frontier T" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3272 | using False connected_Int_frontier [OF connected_segment 1 2] by auto | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3273 | then show ?thesis | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3274 | proof - | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3275 | have "norm (y - x) < e" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3276 | by (metis dist_norm \<open>dist y x < e\<close>) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3277 | moreover have "norm (c - x) \<le> norm (y - x)" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3278 | by (simp add: \<open>c \<in> closed_segment x y\<close> segment_bound(1)) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3279 | ultimately have "norm (c - x) < e" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3280 | by linarith | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3281 | then show ?thesis | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3282 | by (metis (no_types) \<open>c \<in> frontier T\<close> dist_norm that(1)) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3283 | qed | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3284 | qed | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3285 | then show ?thesis | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3286 | by (fastforce simp add: frontier_def closure_approachable) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3287 | qed | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3288 | |
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3289 | lemma frontier_Union_subset: | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3290 | fixes F :: "'a::real_normed_vector set set" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3291 | shows "finite F \<Longrightarrow> frontier(\<Union>F) \<subseteq> (\<Union>t \<in> F. frontier t)" | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3292 | by (rule order_trans [OF frontier_Union_subset_closure]) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3293 | (auto simp: closure_subset_eq) | 
| 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 3294 | |
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3295 | lemma frontier_of_components_subset: | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3296 | fixes S :: "'a::real_normed_vector set" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3297 | shows "C \<in> components S \<Longrightarrow> frontier C \<subseteq> frontier S" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3298 | by (metis Path_Connected.frontier_of_connected_component_subset components_iff) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3299 | |
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3300 | lemma frontier_of_components_closed_complement: | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3301 | fixes S :: "'a::real_normed_vector set" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3302 | shows "\<lbrakk>closed S; C \<in> components (- S)\<rbrakk> \<Longrightarrow> frontier C \<subseteq> S" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3303 | using frontier_complement frontier_of_components_subset frontier_subset_eq by blast | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3304 | |
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3305 | lemma frontier_minimal_separating_closed: | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3306 | fixes S :: "'a::real_normed_vector set" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3307 | assumes "closed S" | 
| 69508 | 3308 | and nconn: "\<not> connected(- S)" | 
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3309 | and C: "C \<in> components (- S)" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3310 | and conn: "\<And>T. \<lbrakk>closed T; T \<subset> S\<rbrakk> \<Longrightarrow> connected(- T)" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3311 | shows "frontier C = S" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3312 | proof (rule ccontr) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3313 | assume "frontier C \<noteq> S" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3314 | then have "frontier C \<subset> S" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3315 | using frontier_of_components_closed_complement [OF \<open>closed S\<close> C] by blast | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3316 | then have "connected(- (frontier C))" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3317 | by (simp add: conn) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3318 | have "\<not> connected(- (frontier C))" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3319 | unfolding connected_def not_not | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3320 | proof (intro exI conjI) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3321 | show "open C" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3322 | using C \<open>closed S\<close> open_components by blast | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3323 | show "open (- closure C)" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3324 | by blast | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3325 |     show "C \<inter> - closure C \<inter> - frontier C = {}"
 | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3326 | using closure_subset by blast | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3327 |     show "C \<inter> - frontier C \<noteq> {}"
 | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3328 | using C \<open>open C\<close> components_eq frontier_disjoint_eq by fastforce | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3329 | show "- frontier C \<subseteq> C \<union> - closure C" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3330 | by (simp add: \<open>open C\<close> closed_Compl frontier_closures) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3331 |     then show "- closure C \<inter> - frontier C \<noteq> {}"
 | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3332 | by (metis (no_types, lifting) C Compl_subset_Compl_iff \<open>frontier C \<subset> S\<close> compl_sup frontier_closures in_components_subset psubsetE sup.absorb_iff2 sup.boundedE sup_bot.right_neutral sup_inf_absorb) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3333 | qed | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3334 | then show False | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3335 | using \<open>connected (- frontier C)\<close> by blast | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3336 | qed | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63978diff
changeset | 3337 | |
| 62843 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62626diff
changeset | 3338 | lemma connected_component_UNIV [simp]: | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3339 | fixes x :: "'a::real_normed_vector" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3340 | shows "connected_component_set UNIV x = UNIV" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3341 | using connected_iff_eq_connected_component_set [of "UNIV::'a set"] connected_UNIV | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3342 | by auto | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3343 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3344 | lemma connected_component_eq_UNIV: | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3345 | fixes x :: "'a::real_normed_vector" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3346 | shows "connected_component_set s x = UNIV \<longleftrightarrow> s = UNIV" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3347 | using connected_component_in connected_component_UNIV by blast | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3348 | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3349 | lemma components_UNIV [simp]: "components UNIV = {UNIV :: 'a::real_normed_vector set}"
 | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3350 | by (auto simp: components_eq_sing_iff) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3351 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3352 | lemma interior_inside_frontier: | 
| 72256 | 3353 | fixes S :: "'a::real_normed_vector set" | 
| 3354 | assumes "bounded S" | |
| 3355 | shows "interior S \<subseteq> inside (frontier S)" | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3356 | proof - | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3357 |   { fix x y
 | 
| 72256 | 3358 | assume x: "x \<in> interior S" and y: "y \<notin> S" | 
| 3359 | and cc: "connected_component (- frontier S) x y" | |
| 3360 |     have "connected_component_set (- frontier S) x \<inter> frontier S \<noteq> {}"
 | |
| 3361 | proof (rule connected_Int_frontier; simp add: set_eq_iff) | |
| 3362 | show "\<exists>u. connected_component (- frontier S) x u \<and> u \<in> S" | |
| 3363 | by (meson cc connected_component_in connected_component_refl_eq interior_subset subsetD x) | |
| 3364 | show "\<exists>u. connected_component (- frontier S) x u \<and> u \<notin> S" | |
| 3365 | using y cc by blast | |
| 3366 | qed | |
| 3367 | then have "bounded (connected_component_set (- frontier S) x)" | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3368 | using connected_component_in by auto | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3369 | } | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3370 | then show ?thesis | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3371 | apply (auto simp: inside_def frontier_def) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3372 | apply (rule classical) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3373 | apply (rule bounded_subset [OF assms], blast) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3374 | done | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3375 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3376 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3377 | lemma inside_empty [simp]: "inside {} = ({} :: 'a :: {real_normed_vector, perfect_space} set)"
 | 
| 71172 | 3378 | by (simp add: inside_def) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3379 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3380 | lemma outside_empty [simp]: "outside {} = (UNIV :: 'a :: {real_normed_vector, perfect_space} set)"
 | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3381 | using inside_empty inside_Un_outside by blast | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3382 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3383 | lemma inside_same_component: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3384 | "\<lbrakk>connected_component (- S) x y; x \<in> inside S\<rbrakk> \<Longrightarrow> y \<in> inside S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3385 | using connected_component_eq connected_component_in | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3386 | by (fastforce simp add: inside_def) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3387 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3388 | lemma outside_same_component: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3389 | "\<lbrakk>connected_component (- S) x y; x \<in> outside S\<rbrakk> \<Longrightarrow> y \<in> outside S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3390 | using connected_component_eq connected_component_in | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3391 | by (fastforce simp add: outside_def) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3392 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3393 | lemma convex_in_outside: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3394 |   fixes S :: "'a :: {real_normed_vector, perfect_space} set"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3395 | assumes S: "convex S" and z: "z \<notin> S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3396 | shows "z \<in> outside S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3397 | proof (cases "S={}")
 | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3398 | case True then show ?thesis by simp | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3399 | next | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3400 | case False then obtain a where "a \<in> S" by blast | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3401 | with z have zna: "z \<noteq> a" by auto | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3402 |   { assume "bounded (connected_component_set (- S) z)"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3403 | with bounded_pos_less obtain B where "B>0" and B: "\<And>x. connected_component (- S) z x \<Longrightarrow> norm x < B" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3404 | by (metis mem_Collect_eq) | 
| 63040 | 3405 | define C where "C = (B + 1 + norm z) / norm (z-a)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3406 | have "C > 0" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3407 | using \<open>0 < B\<close> zna by (simp add: C_def field_split_simps add_strict_increasing) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3408 | have "\<bar>norm (z + C *\<^sub>R (z-a)) - norm (C *\<^sub>R (z-a))\<bar> \<le> norm z" | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3409 | by (metis add_diff_cancel norm_triangle_ineq3) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3410 | moreover have "norm (C *\<^sub>R (z-a)) > norm z + B" | 
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70196diff
changeset | 3411 | using zna \<open>B>0\<close> by (simp add: C_def le_max_iff_disj) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3412 | ultimately have C: "norm (z + C *\<^sub>R (z-a)) > B" by linarith | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3413 |     { fix u::real
 | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3414 | assume u: "0\<le>u" "u\<le>1" and ins: "(1 - u) *\<^sub>R z + u *\<^sub>R (z + C *\<^sub>R (z - a)) \<in> S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3415 | then have Cpos: "1 + u * C > 0" | 
| 61808 | 3416 | by (meson \<open>0 < C\<close> add_pos_nonneg less_eq_real_def zero_le_mult_iff zero_less_one) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3417 | then have *: "(1 / (1 + u * C)) *\<^sub>R z + (u * C / (1 + u * C)) *\<^sub>R z = z" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3418 | by (simp add: scaleR_add_left [symmetric] field_split_simps) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3419 | then have False | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3420 | using convexD_alt [OF S \<open>a \<in> S\<close> ins, of "1/(u*C + 1)"] \<open>C>0\<close> \<open>z \<notin> S\<close> Cpos u | 
| 71172 | 3421 | by (simp add: * field_split_simps) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3422 | } note contra = this | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3423 | have "connected_component (- S) z (z + C *\<^sub>R (z-a))" | 
| 72256 | 3424 | proof (rule connected_componentI [OF connected_segment]) | 
| 3425 | show "closed_segment z (z + C *\<^sub>R (z - a)) \<subseteq> - S" | |
| 3426 | using contra by (force simp add: closed_segment_def) | |
| 3427 | qed auto | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3428 | then have False | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3429 | using zna B [of "z + C *\<^sub>R (z-a)"] C | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 3430 | by (auto simp: field_split_simps max_mult_distrib_right) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3431 | } | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3432 | then show ?thesis | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3433 | by (auto simp: outside_def z) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3434 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3435 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3436 | lemma outside_convex: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3437 |   fixes S :: "'a :: {real_normed_vector, perfect_space} set"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3438 | assumes "convex S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3439 | shows "outside S = - S" | 
| 63955 | 3440 | by (metis ComplD assms convex_in_outside equalityI inside_Un_outside subsetI sup.cobounded2) | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3441 | |
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 3442 | lemma outside_singleton [simp]: | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 3443 |   fixes x :: "'a :: {real_normed_vector, perfect_space}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 3444 |   shows "outside {x} = -{x}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 3445 | by (auto simp: outside_convex) | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 3446 | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3447 | lemma inside_convex: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3448 |   fixes S :: "'a :: {real_normed_vector, perfect_space} set"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3449 |   shows "convex S \<Longrightarrow> inside S = {}"
 | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3450 | by (simp add: inside_outside outside_convex) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3451 | |
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 3452 | lemma inside_singleton [simp]: | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 3453 |   fixes x :: "'a :: {real_normed_vector, perfect_space}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 3454 |   shows "inside {x} = {}"
 | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 3455 | by (auto simp: inside_convex) | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66456diff
changeset | 3456 | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3457 | lemma outside_subset_convex: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3458 |   fixes S :: "'a :: {real_normed_vector, perfect_space} set"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3459 | shows "\<lbrakk>convex T; S \<subseteq> T\<rbrakk> \<Longrightarrow> - T \<subseteq> outside S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3460 | using outside_convex outside_mono by blast | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3461 | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3462 | lemma outside_Un_outside_Un: | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3463 | fixes S :: "'a::real_normed_vector set" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3464 |   assumes "S \<inter> outside(T \<union> U) = {}"
 | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3465 | shows "outside(T \<union> U) \<subseteq> outside(T \<union> S)" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3466 | proof | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3467 | fix x | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3468 | assume x: "x \<in> outside (T \<union> U)" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3469 | have "Y \<subseteq> - S" if "connected Y" "Y \<subseteq> - T" "Y \<subseteq> - U" "x \<in> Y" "u \<in> Y" for u Y | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3470 | proof - | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3471 | have "Y \<subseteq> connected_component_set (- (T \<union> U)) x" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3472 | by (simp add: connected_component_maximal that) | 
| 68096 | 3473 | also have "\<dots> \<subseteq> outside(T \<union> U)" | 
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3474 | by (metis (mono_tags, lifting) Collect_mono mem_Collect_eq outside outside_same_component x) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3475 | finally have "Y \<subseteq> outside(T \<union> U)" . | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3476 | with assms show ?thesis by auto | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3477 | qed | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3478 | with x show "x \<in> outside (T \<union> S)" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3479 | by (simp add: outside_connected_component_lt connected_component_def) meson | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3480 | qed | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3481 | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3482 | lemma outside_frontier_misses_closure: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3483 | fixes S :: "'a::real_normed_vector set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3484 | assumes "bounded S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3485 | shows "outside(frontier S) \<subseteq> - closure S" | 
| 74123 
7c5842b06114
clarified abstract and concrete boolean algebras
 haftmann parents: 
74007diff
changeset | 3486 | unfolding outside_inside boolean_algebra_class.compl_le_compl_iff | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3487 | proof - | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3488 |   { assume "interior S \<subseteq> inside (frontier S)"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3489 | hence "interior S \<union> inside (frontier S) = inside (frontier S)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3490 | by (simp add: subset_Un_eq) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3491 | then have "closure S \<subseteq> frontier S \<union> inside (frontier S)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3492 | using frontier_def by auto | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3493 | } | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3494 | then show "closure S \<subseteq> frontier S \<union> inside (frontier S)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3495 | using interior_inside_frontier [OF assms] by blast | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3496 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3497 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3498 | lemma outside_frontier_eq_complement_closure: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3499 |   fixes S :: "'a :: {real_normed_vector, perfect_space} set"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3500 | assumes "bounded S" "convex S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3501 | shows "outside(frontier S) = - closure S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3502 | by (metis Diff_subset assms convex_closure frontier_def outside_frontier_misses_closure | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3503 | outside_subset_convex subset_antisym) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3504 | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3505 | lemma inside_frontier_eq_interior: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3506 |      fixes S :: "'a :: {real_normed_vector, perfect_space} set"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3507 | shows "\<lbrakk>bounded S; convex S\<rbrakk> \<Longrightarrow> inside(frontier S) = interior S" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3508 | apply (simp add: inside_outside outside_frontier_eq_complement_closure) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3509 | using closure_subset interior_subset | 
| 68096 | 3510 | apply (auto simp: frontier_def) | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3511 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3512 | |
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3513 | lemma open_inside: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3514 | fixes S :: "'a::real_normed_vector set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3515 | assumes "closed S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3516 | shows "open (inside S)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3517 | proof - | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3518 |   { fix x assume x: "x \<in> inside S"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3519 | have "open (connected_component_set (- S) x)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3520 | using assms open_connected_component by blast | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3521 | then obtain e where e: "e>0" and e: "\<And>y. dist y x < e \<longrightarrow> connected_component (- S) x y" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3522 | using dist_not_less_zero | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3523 | apply (simp add: open_dist) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3524 | by (metis (no_types, lifting) Compl_iff connected_component_refl_eq inside_def mem_Collect_eq x) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3525 | then have "\<exists>e>0. ball x e \<subseteq> inside S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3526 | by (metis e dist_commute inside_same_component mem_ball subsetI x) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3527 | } | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3528 | then show ?thesis | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3529 | by (simp add: open_contains_ball) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3530 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3531 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3532 | lemma open_outside: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3533 | fixes S :: "'a::real_normed_vector set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3534 | assumes "closed S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3535 | shows "open (outside S)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3536 | proof - | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3537 |   { fix x assume x: "x \<in> outside S"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3538 | have "open (connected_component_set (- S) x)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3539 | using assms open_connected_component by blast | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3540 | then obtain e where e: "e>0" and e: "\<And>y. dist y x < e \<longrightarrow> connected_component (- S) x y" | 
| 72256 | 3541 | using dist_not_less_zero x | 
| 3542 | by (auto simp add: open_dist outside_def intro: connected_component_refl) | |
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3543 | then have "\<exists>e>0. ball x e \<subseteq> outside S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3544 | by (metis e dist_commute outside_same_component mem_ball subsetI x) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3545 | } | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3546 | then show ?thesis | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3547 | by (simp add: open_contains_ball) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3548 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3549 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3550 | lemma closure_inside_subset: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3551 | fixes S :: "'a::real_normed_vector set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3552 | assumes "closed S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3553 | shows "closure(inside S) \<subseteq> S \<union> inside S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3554 | by (metis assms closure_minimal open_closed open_outside sup.cobounded2 union_with_inside) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3555 | |
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3556 | lemma frontier_inside_subset: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3557 | fixes S :: "'a::real_normed_vector set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3558 | assumes "closed S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3559 | shows "frontier(inside S) \<subseteq> S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3560 | proof - | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3561 | have "closure (inside S) \<inter> - inside S = closure (inside S) - interior (inside S)" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3562 | by (metis (no_types) Diff_Compl assms closure_closed interior_closure open_closed open_inside) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3563 | moreover have "- inside S \<inter> - outside S = S" | 
| 63955 | 3564 | by (metis (no_types) compl_sup double_compl inside_Un_outside) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3565 | moreover have "closure (inside S) \<subseteq> - outside S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3566 | by (metis (no_types) assms closure_inside_subset union_with_inside) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3567 | ultimately have "closure (inside S) - interior (inside S) \<subseteq> S" | 
| 61426 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3568 | by blast | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3569 | then show ?thesis | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3570 | by (simp add: frontier_def open_inside interior_open) | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3571 | qed | 
| 
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
 paulson <lp15@cam.ac.uk> parents: 
61204diff
changeset | 3572 | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3573 | lemma closure_outside_subset: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3574 | fixes S :: "'a::real_normed_vector set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3575 | assumes "closed S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3576 | shows "closure(outside S) \<subseteq> S \<union> outside S" | 
| 72256 | 3577 | by (metis assms closed_open closure_minimal inside_outside open_inside sup_ge2) | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3578 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3579 | lemma frontier_outside_subset: | 
| 72256 | 3580 | fixes S :: "'a::real_normed_vector set" | 
| 3581 | assumes "closed S" | |
| 3582 | shows "frontier(outside S) \<subseteq> S" | |
| 3583 | unfolding frontier_def | |
| 3584 | by (metis Diff_subset_conv assms closure_outside_subset interior_eq open_outside sup_aci(1)) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3585 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3586 | lemma inside_complement_unbounded_connected_empty: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3587 |      "\<lbrakk>connected (- S); \<not> bounded (- S)\<rbrakk> \<Longrightarrow> inside S = {}"
 | 
| 72256 | 3588 | using inside_subset by blast | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3589 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3590 | lemma inside_bounded_complement_connected_empty: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3591 |     fixes S :: "'a::{real_normed_vector, perfect_space} set"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3592 |     shows "\<lbrakk>connected (- S); bounded S\<rbrakk> \<Longrightarrow> inside S = {}"
 | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3593 | by (metis inside_complement_unbounded_connected_empty cobounded_imp_unbounded) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3594 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3595 | lemma inside_inside: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3596 | assumes "S \<subseteq> inside T" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3597 | shows "inside S - T \<subseteq> inside T" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3598 | unfolding inside_def | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3599 | proof clarify | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3600 | fix x | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3601 | assume x: "x \<notin> T" "x \<notin> S" and bo: "bounded (connected_component_set (- S) x)" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3602 | show "bounded (connected_component_set (- T) x)" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3603 |   proof (cases "S \<inter> connected_component_set (- T) x = {}")
 | 
| 72256 | 3604 | case True then show ?thesis | 
| 3605 | by (metis bounded_subset [OF bo] compl_le_compl_iff connected_component_idemp connected_component_mono disjoint_eq_subset_Compl double_compl) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3606 | next | 
| 72256 | 3607 | case False | 
| 3608 | then obtain y where y: "y \<in> S" "y \<in> connected_component_set (- T) x" | |
| 3609 | by (meson disjoint_iff) | |
| 3610 | then have "bounded (connected_component_set (- T) y)" | |
| 3611 | using assms [unfolded inside_def] by blast | |
| 3612 | with y show ?thesis | |
| 3613 | by (metis connected_component_eq) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3614 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3615 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3616 | |
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3617 | lemma inside_inside_subset: "inside(inside S) \<subseteq> S" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3618 | using inside_inside union_with_outside by fastforce | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3619 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3620 | lemma inside_outside_intersect_connected: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3621 |       "\<lbrakk>connected T; inside S \<inter> T \<noteq> {}; outside S \<inter> T \<noteq> {}\<rbrakk> \<Longrightarrow> S \<inter> T \<noteq> {}"
 | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3622 | apply (simp add: inside_def outside_def ex_in_conv [symmetric] disjoint_eq_subset_Compl, clarify) | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73795diff
changeset | 3623 | by (metis (no_types, opaque_lifting) Compl_anti_mono connected_component_eq connected_component_maximal contra_subsetD double_compl) | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3624 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3625 | lemma outside_bounded_nonempty: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3626 |   fixes S :: "'a :: {real_normed_vector, perfect_space} set"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3627 |     assumes "bounded S" shows "outside S \<noteq> {}"
 | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3628 | by (metis (no_types, lifting) Collect_empty_eq Collect_mem_eq Compl_eq_Diff_UNIV Diff_cancel | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3629 | Diff_disjoint UNIV_I assms ball_eq_empty bounded_diff cobounded_outside convex_ball | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3630 | double_complement order_refl outside_convex outside_def) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3631 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3632 | lemma outside_compact_in_open: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3633 |     fixes S :: "'a :: {real_normed_vector,perfect_space} set"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3634 |     assumes S: "compact S" and T: "open T" and "S \<subseteq> T" "T \<noteq> {}"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3635 |       shows "outside S \<inter> T \<noteq> {}"
 | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3636 | proof - | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3637 |   have "outside S \<noteq> {}"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3638 | by (simp add: compact_imp_bounded outside_bounded_nonempty S) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3639 | with assms obtain a b where a: "a \<in> outside S" and b: "b \<in> T" by auto | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3640 | show ?thesis | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3641 | proof (cases "a \<in> T") | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3642 | case True with a show ?thesis by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3643 | next | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3644 | case False | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3645 | have front: "frontier T \<subseteq> - S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3646 | using \<open>S \<subseteq> T\<close> frontier_disjoint_eq T by auto | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3647 |       { fix \<gamma>
 | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3648 |         assume "path \<gamma>" and pimg_sbs: "path_image \<gamma> - {pathfinish \<gamma>} \<subseteq> interior (- T)"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3649 | and pf: "pathfinish \<gamma> \<in> frontier T" and ps: "pathstart \<gamma> = a" | 
| 63040 | 3650 | define c where "c = pathfinish \<gamma>" | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3651 | have "c \<in> -S" unfolding c_def using front pf by blast | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3652 | moreover have "open (-S)" using S compact_imp_closed by blast | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3653 | ultimately obtain \<epsilon>::real where "\<epsilon> > 0" and \<epsilon>: "cball c \<epsilon> \<subseteq> -S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3654 | using open_contains_cball[of "-S"] S by blast | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3655 | then obtain d where "d \<in> T" and d: "dist d c < \<epsilon>" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3656 | using closure_approachable [of c T] pf unfolding c_def | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3657 | by (metis Diff_iff frontier_def) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3658 | then have "d \<in> -S" using \<epsilon> | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3659 | using dist_commute by (metis contra_subsetD mem_cball not_le not_less_iff_gr_or_eq) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3660 | have pimg_sbs_cos: "path_image \<gamma> \<subseteq> -S" | 
| 72256 | 3661 | using \<open>c \<in> - S\<close> \<open>S \<subseteq> T\<close> c_def interior_subset pimg_sbs by fastforce | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3662 | have "closed_segment c d \<le> cball c \<epsilon>" | 
| 72256 | 3663 | by (metis \<open>0 < \<epsilon>\<close> centre_in_cball closed_segment_subset convex_cball d dist_commute less_eq_real_def mem_cball) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3664 | with \<epsilon> have "closed_segment c d \<subseteq> -S" by blast | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3665 | moreover have con_gcd: "connected (path_image \<gamma> \<union> closed_segment c d)" | 
| 61808 | 3666 | by (rule connected_Un) (auto simp: c_def \<open>path \<gamma>\<close> connected_path_image) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3667 | ultimately have "connected_component (- S) a d" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3668 | unfolding connected_component_def using pimg_sbs_cos ps by blast | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3669 |         then have "outside S \<inter> T \<noteq> {}"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3670 | using outside_same_component [OF _ a] by (metis IntI \<open>d \<in> T\<close> empty_iff) | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3671 | } note * = this | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3672 | have pal: "pathstart (linepath a b) \<in> closure (- T)" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3673 | by (auto simp: False closure_def) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3674 | show ?thesis | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3675 | by (rule exists_path_subpath_to_frontier [OF path_linepath pal _ *]) (auto simp: b) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3676 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3677 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3678 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3679 | lemma inside_inside_compact_connected: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3680 | fixes S :: "'a :: euclidean_space set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3681 | assumes S: "closed S" and T: "compact T" and "connected T" "S \<subseteq> inside T" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3682 | shows "inside S \<subseteq> inside T" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3683 | proof (cases "inside T = {}")
 | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3684 | case True with assms show ?thesis by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3685 | next | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3686 | case False | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3687 |   consider "DIM('a) = 1" | "DIM('a) \<ge> 2"
 | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3688 | using antisym not_less_eq_eq by fastforce | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3689 | then show ?thesis | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3690 | proof cases | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3691 | case 1 then show ?thesis | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3692 | using connected_convex_1_gen assms False inside_convex by blast | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3693 | next | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3694 | case 2 | 
| 72256 | 3695 | have "bounded S" | 
| 3696 | using assms by (meson bounded_inside bounded_subset compact_imp_bounded) | |
| 3697 | then have coms: "compact S" | |
| 3698 | by (simp add: S compact_eq_bounded_closed) | |
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3699 | then have bst: "bounded (S \<union> T)" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3700 | by (simp add: compact_imp_bounded T) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3701 | then obtain r where "0 < r" and r: "S \<union> T \<subseteq> ball 0 r" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3702 | using bounded_subset_ballD by blast | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3703 |     have outst: "outside S \<inter> outside T \<noteq> {}"
 | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3704 | proof - | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3705 | have "- ball 0 r \<subseteq> outside S" | 
| 72256 | 3706 | by (meson convex_ball le_supE outside_subset_convex r) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3707 | moreover have "- ball 0 r \<subseteq> outside T" | 
| 72256 | 3708 | by (meson convex_ball le_supE outside_subset_convex r) | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3709 | ultimately show ?thesis | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3710 | by (metis Compl_subset_Compl_iff Int_subset_iff bounded_ball inf.orderE outside_bounded_nonempty outside_no_overlap) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3711 | qed | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3712 |     have "S \<inter> T = {}" using assms
 | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3713 | by (metis disjoint_iff_not_equal inside_no_overlap subsetCE) | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3714 |     moreover have "outside S \<inter> inside T \<noteq> {}"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3715 | by (meson False assms(4) compact_eq_bounded_closed coms open_inside outside_compact_in_open T) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3716 |     ultimately have "inside S \<inter> T = {}"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3717 | using inside_outside_intersect_connected [OF \<open>connected T\<close>, of S] | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3718 | by (metis "2" compact_eq_bounded_closed coms connected_outside inf.commute inside_outside_intersect_connected outst) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3719 | then show ?thesis | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3720 | using inside_inside [OF \<open>S \<subseteq> inside T\<close>] by blast | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3721 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3722 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3723 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3724 | lemma connected_with_inside: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3725 | fixes S :: "'a :: real_normed_vector set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3726 | assumes S: "closed S" and cons: "connected S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3727 | shows "connected(S \<union> inside S)" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3728 | proof (cases "S \<union> inside S = UNIV") | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3729 | case True with assms show ?thesis by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3730 | next | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3731 | case False | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3732 | then obtain b where b: "b \<notin> S" "b \<notin> inside S" by blast | 
| 72256 | 3733 | have *: "\<exists>y T. y \<in> S \<and> connected T \<and> a \<in> T \<and> y \<in> T \<and> T \<subseteq> (S \<union> inside S)" | 
| 3734 | if "a \<in> S \<union> inside S" for a | |
| 3735 | using that | |
| 3736 | proof | |
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3737 | assume "a \<in> S" then show ?thesis | 
| 72256 | 3738 |       by (rule_tac x=a in exI, rule_tac x="{a}" in exI, simp)
 | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3739 | next | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3740 | assume a: "a \<in> inside S" | 
| 72256 | 3741 | then have ain: "a \<in> closure (inside S)" | 
| 3742 | by (simp add: closure_def) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3743 | show ?thesis | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3744 | apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "inside S"]) | 
| 72256 | 3745 | apply (simp_all add: ain b) | 
| 3746 | subgoal for h | |
| 3747 | apply (rule_tac x="pathfinish h" in exI) | |
| 3748 | apply (simp add: subsetD [OF frontier_inside_subset[OF S]]) | |
| 3749 | apply (rule_tac x="path_image h" in exI) | |
| 3750 | apply (simp add: pathfinish_in_path_image connected_path_image, auto) | |
| 3751 | by (metis Diff_single_insert S frontier_inside_subset insert_iff interior_subset subsetD) | |
| 3752 | done | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3753 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3754 | show ?thesis | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3755 | apply (simp add: connected_iff_connected_component) | 
| 72256 | 3756 | apply (clarsimp simp add: connected_component_def dest!: *) | 
| 3757 | subgoal for x y u u' T t' | |
| 3758 | by (rule_tac x="(S \<union> T \<union> t')" in exI) (auto intro!: connected_Un cons) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3759 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3760 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3761 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3762 | text\<open>The proof is virtually the same as that above.\<close> | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3763 | lemma connected_with_outside: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3764 | fixes S :: "'a :: real_normed_vector set" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3765 | assumes S: "closed S" and cons: "connected S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3766 | shows "connected(S \<union> outside S)" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3767 | proof (cases "S \<union> outside S = UNIV") | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3768 | case True with assms show ?thesis by auto | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3769 | next | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3770 | case False | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3771 | then obtain b where b: "b \<notin> S" "b \<notin> outside S" by blast | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3772 | have *: "\<exists>y T. y \<in> S \<and> connected T \<and> a \<in> T \<and> y \<in> T \<and> T \<subseteq> (S \<union> outside S)" if "a \<in> (S \<union> outside S)" for a | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3773 | using that proof | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3774 | assume "a \<in> S" then show ?thesis | 
| 72256 | 3775 |       by (rule_tac x=a in exI, rule_tac x="{a}" in exI, simp)
 | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3776 | next | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3777 | assume a: "a \<in> outside S" | 
| 72256 | 3778 | then have ain: "a \<in> closure (outside S)" | 
| 3779 | by (simp add: closure_def) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3780 | show ?thesis | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3781 | apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "outside S"]) | 
| 72256 | 3782 | apply (simp_all add: ain b) | 
| 3783 | subgoal for h | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3784 | apply (rule_tac x="pathfinish h" in exI) | 
| 72256 | 3785 | apply (simp add: subsetD [OF frontier_outside_subset[OF S]]) | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3786 | apply (rule_tac x="path_image h" in exI) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3787 | apply (simp add: pathfinish_in_path_image connected_path_image, auto) | 
| 72256 | 3788 | by (metis (no_types, lifting) frontier_outside_subset insertE insert_Diff interior_eq open_outside pathfinish_in_path_image S subsetCE) | 
| 3789 | done | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3790 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3791 | show ?thesis | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3792 | apply (simp add: connected_iff_connected_component) | 
| 72256 | 3793 | apply (clarsimp simp add: connected_component_def dest!: *) | 
| 3794 | subgoal for x y u u' T t' | |
| 3795 | by (rule_tac x="(S \<union> T \<union> t')" in exI) (auto intro!: connected_Un cons) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3796 | done | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3797 | qed | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3798 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3799 | lemma inside_inside_eq_empty [simp]: | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3800 |     fixes S :: "'a :: {real_normed_vector, perfect_space} set"
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3801 | assumes S: "closed S" and cons: "connected S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3802 |       shows "inside (inside S) = {}"
 | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3803 | by (metis (no_types) unbounded_outside connected_with_outside [OF assms] bounded_Un | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3804 | inside_complement_unbounded_connected_empty unbounded_outside union_with_outside) | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3805 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3806 | lemma inside_in_components: | 
| 72256 | 3807 |      "inside S \<in> components (- S) \<longleftrightarrow> connected(inside S) \<and> inside S \<noteq> {}" (is "?lhs = ?rhs")
 | 
| 3808 | proof | |
| 3809 | assume R: ?rhs | |
| 3810 | then have "\<And>x. \<lbrakk>x \<in> S; x \<in> inside S\<rbrakk> \<Longrightarrow> \<not> connected (inside S)" | |
| 3811 | by (simp add: inside_outside) | |
| 3812 | with R show ?lhs | |
| 3813 | unfolding in_components_maximal | |
| 3814 | by (auto intro: inside_same_component connected_componentI) | |
| 3815 | qed (simp add: in_components_maximal) | |
| 3816 | ||
| 3817 | text\<open>The proof is like that above.\<close> | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3818 | lemma outside_in_components: | 
| 72256 | 3819 |      "outside S \<in> components (- S) \<longleftrightarrow> connected(outside S) \<and> outside S \<noteq> {}" (is "?lhs = ?rhs")
 | 
| 3820 | proof | |
| 3821 | assume R: ?rhs | |
| 3822 | then have "\<And>x. \<lbrakk>x \<in> S; x \<in> outside S\<rbrakk> \<Longrightarrow> \<not> connected (outside S)" | |
| 3823 | by (meson disjoint_iff outside_no_overlap) | |
| 3824 | with R show ?lhs | |
| 3825 | unfolding in_components_maximal | |
| 3826 | by (auto intro: outside_same_component connected_componentI) | |
| 3827 | qed (simp add: in_components_maximal) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3828 | |
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3829 | lemma bounded_unique_outside: | 
| 72256 | 3830 | fixes S :: "'a :: euclidean_space set" | 
| 3831 |   assumes "bounded S" "DIM('a) \<ge> 2"
 | |
| 3832 | shows "(c \<in> components (- S) \<and> \<not> bounded c \<longleftrightarrow> c = outside S)" | |
| 3833 | using assms | |
| 3834 | by (metis cobounded_unique_unbounded_components connected_outside double_compl outside_bounded_nonempty outside_in_components unbounded_outside) | |
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61426diff
changeset | 3835 | |
| 69514 | 3836 | |
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3837 | subsection\<open>Condition for an open map's image to contain a ball\<close> | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3838 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68532diff
changeset | 3839 | proposition ball_subset_open_map_image: | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3840 |   fixes f :: "'a::heine_borel \<Rightarrow> 'b :: {real_normed_vector,heine_borel}"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3841 | assumes contf: "continuous_on (closure S) f" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3842 | and oint: "open (f ` interior S)" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3843 | and le_no: "\<And>z. z \<in> frontier S \<Longrightarrow> r \<le> norm(f z - f a)" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3844 | and "bounded S" "a \<in> S" "0 < r" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3845 | shows "ball (f a) r \<subseteq> f ` S" | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68532diff
changeset | 3846 | proof (cases "f ` S = UNIV") | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3847 | case True then show ?thesis by simp | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3848 | next | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3849 | case False | 
| 72256 | 3850 |   then have "closed (frontier (f ` S))" "frontier (f ` S) \<noteq> {}"
 | 
| 3851 | using \<open>a \<in> S\<close> by (auto simp: frontier_eq_empty) | |
| 3852 | then obtain w where w: "w \<in> frontier (f ` S)" | |
| 3853 | and dw_le: "\<And>y. y \<in> frontier (f ` S) \<Longrightarrow> norm (f a - w) \<le> norm (f a - y)" | |
| 3854 | by (auto simp add: dist_norm intro: distance_attains_inf [of "frontier(f ` S)" "f a"]) | |
| 3855 | then obtain \<xi> where \<xi>: "\<And>n. \<xi> n \<in> f ` S" and tendsw: "\<xi> \<longlonglongrightarrow> w" | |
| 3856 | by (metis Diff_iff frontier_def closure_sequential) | |
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3857 | then have "\<And>n. \<exists>x \<in> S. \<xi> n = f x" by force | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3858 | then obtain z where zs: "\<And>n. z n \<in> S" and fz: "\<And>n. \<xi> n = f (z n)" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3859 | by metis | 
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
65038diff
changeset | 3860 | then obtain y K where y: "y \<in> closure S" and "strict_mono (K :: nat \<Rightarrow> nat)" | 
| 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
65038diff
changeset | 3861 | and Klim: "(z \<circ> K) \<longlonglongrightarrow> y" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3862 | using \<open>bounded S\<close> | 
| 72256 | 3863 | unfolding compact_closure [symmetric] compact_def by (meson closure_subset subset_iff) | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3864 | then have ftendsw: "((\<lambda>n. f (z n)) \<circ> K) \<longlonglongrightarrow> w" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3865 | by (metis LIMSEQ_subseq_LIMSEQ fun.map_cong0 fz tendsw) | 
| 68096 | 3866 | have zKs: "\<And>n. (z \<circ> K) n \<in> S" by (simp add: zs) | 
| 63540 | 3867 | have fz: "f \<circ> z = \<xi>" "(\<lambda>n. f (z n)) = \<xi>" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3868 | using fz by auto | 
| 63540 | 3869 | then have "(\<xi> \<circ> K) \<longlonglongrightarrow> f y" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3870 | by (metis (no_types) Klim zKs y contf comp_assoc continuous_on_closure_sequentially) | 
| 63540 | 3871 | with fz have wy: "w = f y" using fz LIMSEQ_unique ftendsw by auto | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3872 | have rle: "r \<le> norm (f y - f a)" | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3873 | proof (rule le_no) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3874 | show "y \<in> frontier S" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3875 | using w wy oint by (force simp: imageI image_mono interiorI interior_subset frontier_def y) | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3876 | qed | 
| 69508 | 3877 |     have **: "(b \<inter> (- S) \<noteq> {} \<and> b - (- S) \<noteq> {} \<Longrightarrow> b \<inter> f \<noteq> {})
 | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3878 |               \<Longrightarrow> (b \<inter> S \<noteq> {}) \<Longrightarrow> b \<inter> f = {} \<Longrightarrow> b \<subseteq> S" 
 | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3879 | for b f and S :: "'b set" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3880 | by blast | 
| 72228 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3881 | have \<section>: "\<And>y. \<lbrakk>norm (f a - y) < r; y \<in> frontier (f ` S)\<rbrakk> \<Longrightarrow> False" | 
| 
aa7cb84983e9
minor tidying, also s->S and t->T
 paulson <lp15@cam.ac.uk> parents: 
71633diff
changeset | 3882 | by (metis dw_le norm_minus_commute not_less order_trans rle wy) | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3883 | show ?thesis | 
| 72256 | 3884 | apply (rule ** [OF connected_Int_frontier [where t = "f`S", OF connected_ball]]) | 
| 3885 | (*such a horrible mess*) | |
| 3886 | using \<open>a \<in> S\<close> \<open>0 < r\<close> by (auto simp: disjoint_iff_not_equal dist_norm dest: \<section>) | |
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3887 | qed | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62398diff
changeset | 3888 | |
| 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: 
70178diff
changeset | 3889 | |
| 
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: 
70178diff
changeset | 3890 | subsubsection\<open>Special characterizations of classes of functions into and out of R.\<close> | 
| 
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: 
70178diff
changeset | 3891 | |
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3892 | lemma Hausdorff_space_euclidean [simp]: "Hausdorff_space (euclidean :: 'a::metric_space topology)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3893 | proof - | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3894 | have "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> disjnt U V" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3895 | if "x \<noteq> y" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3896 | for x y :: 'a | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3897 | proof (intro exI conjI) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3898 | let ?r = "dist x y / 2" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3899 | have [simp]: "?r > 0" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3900 | by (simp add: that) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3901 | show "open (ball x ?r)" "open (ball y ?r)" "x \<in> (ball x ?r)" "y \<in> (ball y ?r)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3902 | by (auto simp add: that) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3903 | show "disjnt (ball x ?r) (ball y ?r)" | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3904 | unfolding disjnt_def by (simp add: disjoint_ballI) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3905 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3906 | then show ?thesis | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3907 | by (simp add: Hausdorff_space_def) | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3908 | qed | 
| 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
71191diff
changeset | 3909 | |
| 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: 
70178diff
changeset | 3910 | proposition embedding_map_into_euclideanreal: | 
| 
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: 
70178diff
changeset | 3911 | assumes "path_connected_space X" | 
| 
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: 
70178diff
changeset | 3912 | shows "embedding_map X euclideanreal f \<longleftrightarrow> | 
| 
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: 
70178diff
changeset | 3913 | continuous_map X euclideanreal f \<and> inj_on f (topspace X)" | 
| 
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: 
70178diff
changeset | 3914 | proof safe | 
| 
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: 
70178diff
changeset | 3915 | show "continuous_map X euclideanreal 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: 
70178diff
changeset | 3916 | if "embedding_map X euclideanreal 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: 
70178diff
changeset | 3917 | using continuous_map_in_subtopology homeomorphic_imp_continuous_map that | 
| 
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: 
70178diff
changeset | 3918 | unfolding embedding_map_def by blast | 
| 
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: 
70178diff
changeset | 3919 | show "inj_on f (topspace X)" | 
| 
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: 
70178diff
changeset | 3920 | if "embedding_map X euclideanreal 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: 
70178diff
changeset | 3921 | using that homeomorphic_imp_injective_map | 
| 
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: 
70178diff
changeset | 3922 | unfolding embedding_map_def by blast | 
| 
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: 
70178diff
changeset | 3923 | show "embedding_map X euclideanreal 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: 
70178diff
changeset | 3924 | if cont: "continuous_map X euclideanreal f" and inj: "inj_on f (topspace X)" | 
| 
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: 
70178diff
changeset | 3925 | proof - | 
| 
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: 
70178diff
changeset | 3926 | obtain g where gf: "\<And>x. x \<in> topspace X \<Longrightarrow> g (f x) = x" | 
| 
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: 
70178diff
changeset | 3927 | using inv_into_f_f [OF inj] by auto | 
| 
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: 
70178diff
changeset | 3928 | show ?thesis | 
| 
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: 
70178diff
changeset | 3929 | unfolding embedding_map_def homeomorphic_map_maps homeomorphic_maps_def | 
| 
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: 
70178diff
changeset | 3930 | proof (intro exI conjI) | 
| 
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: 
70178diff
changeset | 3931 | show "continuous_map X (top_of_set (f ` topspace X)) 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: 
70178diff
changeset | 3932 | by (simp add: cont continuous_map_in_subtopology) | 
| 
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: 
70178diff
changeset | 3933 | let ?S = "f ` topspace X" | 
| 
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: 
70178diff
changeset | 3934 |       have eq: "{x \<in> ?S. g x \<in> U} = f ` U" if "openin X U" for U
 | 
| 
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: 
70178diff
changeset | 3935 | using openin_subset [OF that] by (auto simp: gf) | 
| 
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: 
70178diff
changeset | 3936 | have 1: "g ` ?S \<subseteq> topspace X" | 
| 
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: 
70178diff
changeset | 3937 | using eq by blast | 
| 
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: 
70178diff
changeset | 3938 |       have "openin (top_of_set ?S) {x \<in> ?S. g x \<in> T}"
 | 
| 
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: 
70178diff
changeset | 3939 | if "openin X T" for T | 
| 
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: 
70178diff
changeset | 3940 | proof - | 
| 
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: 
70178diff
changeset | 3941 | have "T \<subseteq> topspace X" | 
| 
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: 
70178diff
changeset | 3942 | by (simp add: openin_subset that) | 
| 
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: 
70178diff
changeset | 3943 | have RR: "\<forall>x \<in> ?S \<inter> g -` T. \<exists>d>0. \<forall>x' \<in> ?S \<inter> ball x d. g x' \<in> T" | 
| 
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: 
70178diff
changeset | 3944 | proof (clarsimp simp add: gf) | 
| 
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: 
70178diff
changeset | 3945 | have pcS: "path_connectedin euclidean ?S" | 
| 
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: 
70178diff
changeset | 3946 | using assms cont path_connectedin_continuous_map_image path_connectedin_topspace by blast | 
| 
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: 
70178diff
changeset | 3947 | show "\<exists>d>0. \<forall>x'\<in>f ` topspace X \<inter> ball (f x) d. g x' \<in> T" | 
| 
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: 
70178diff
changeset | 3948 | if "x \<in> T" for x | 
| 
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: 
70178diff
changeset | 3949 | proof - | 
| 
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: 
70178diff
changeset | 3950 | have x: "x \<in> topspace X" | 
| 
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: 
70178diff
changeset | 3951 | using \<open>T \<subseteq> topspace X\<close> \<open>x \<in> T\<close> by blast | 
| 
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: 
70178diff
changeset | 3952 | obtain u v d where "0 < d" "u \<in> topspace X" "v \<in> topspace X" | 
| 
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: 
70178diff
changeset | 3953 |                          and sub_fuv: "?S \<inter> {f x - d .. f x + d} \<subseteq> {f u..f v}"
 | 
| 
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: 
70178diff
changeset | 3954 | proof (cases "\<exists>u \<in> topspace X. f u < f x") | 
| 
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: 
70178diff
changeset | 3955 | case True | 
| 
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: 
70178diff
changeset | 3956 | then obtain u where u: "u \<in> topspace X" "f u < f x" .. | 
| 
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: 
70178diff
changeset | 3957 | show ?thesis | 
| 
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: 
70178diff
changeset | 3958 | proof (cases "\<exists>v \<in> topspace X. f x < f v") | 
| 
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: 
70178diff
changeset | 3959 | case True | 
| 
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: 
70178diff
changeset | 3960 | then obtain v where v: "v \<in> topspace X" "f x < f v" .. | 
| 
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: 
70178diff
changeset | 3961 | show ?thesis | 
| 
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: 
70178diff
changeset | 3962 | proof | 
| 
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: 
70178diff
changeset | 3963 | let ?d = "min (f x - f u) (f v - f x)" | 
| 
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: 
70178diff
changeset | 3964 | show "0 < ?d" | 
| 
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: 
70178diff
changeset | 3965 | by (simp add: \<open>f u < f x\<close> \<open>f x < f v\<close>) | 
| 
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: 
70178diff
changeset | 3966 |                   show "f ` topspace X \<inter> {f x - ?d..f x + ?d} \<subseteq> {f u..f v}"
 | 
| 
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: 
70178diff
changeset | 3967 | by fastforce | 
| 
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: 
70178diff
changeset | 3968 | qed (auto simp: u v) | 
| 
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: 
70178diff
changeset | 3969 | next | 
| 
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: 
70178diff
changeset | 3970 | case False | 
| 
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: 
70178diff
changeset | 3971 | show ?thesis | 
| 
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: 
70178diff
changeset | 3972 | proof | 
| 
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: 
70178diff
changeset | 3973 | let ?d = "f x - f u" | 
| 
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: 
70178diff
changeset | 3974 | show "0 < ?d" | 
| 
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: 
70178diff
changeset | 3975 | by (simp add: u) | 
| 
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: 
70178diff
changeset | 3976 |                   show "f ` topspace X \<inter> {f x - ?d..f x + ?d} \<subseteq> {f u..f x}"
 | 
| 
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: 
70178diff
changeset | 3977 | using x u False by auto | 
| 
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: 
70178diff
changeset | 3978 | qed (auto simp: x u) | 
| 
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: 
70178diff
changeset | 3979 | qed | 
| 
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: 
70178diff
changeset | 3980 | next | 
| 
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: 
70178diff
changeset | 3981 | case False | 
| 
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: 
70178diff
changeset | 3982 | note no_u = False | 
| 
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: 
70178diff
changeset | 3983 | show ?thesis | 
| 
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: 
70178diff
changeset | 3984 | proof (cases "\<exists>v \<in> topspace X. f x < f v") | 
| 
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: 
70178diff
changeset | 3985 | case True | 
| 
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: 
70178diff
changeset | 3986 | then obtain v where v: "v \<in> topspace X" "f x < f v" .. | 
| 
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: 
70178diff
changeset | 3987 | show ?thesis | 
| 
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: 
70178diff
changeset | 3988 | proof | 
| 
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: 
70178diff
changeset | 3989 | let ?d = "f v - f x" | 
| 
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: 
70178diff
changeset | 3990 | show "0 < ?d" | 
| 
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: 
70178diff
changeset | 3991 | by (simp add: v) | 
| 
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: 
70178diff
changeset | 3992 |                   show "f ` topspace X \<inter> {f x - ?d..f x + ?d} \<subseteq> {f x..f v}"
 | 
| 
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: 
70178diff
changeset | 3993 | using False by auto | 
| 
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: 
70178diff
changeset | 3994 | qed (auto simp: x v) | 
| 
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: 
70178diff
changeset | 3995 | next | 
| 
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: 
70178diff
changeset | 3996 | case False | 
| 
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: 
70178diff
changeset | 3997 | show ?thesis | 
| 
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: 
70178diff
changeset | 3998 | proof | 
| 
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: 
70178diff
changeset | 3999 |                   show "f ` topspace X \<inter> {f x - 1..f x + 1} \<subseteq> {f x..f x}"
 | 
| 
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: 
70178diff
changeset | 4000 | using False no_u by fastforce | 
| 
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: 
70178diff
changeset | 4001 | qed (auto simp: x) | 
| 
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: 
70178diff
changeset | 4002 | qed | 
| 
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: 
70178diff
changeset | 4003 | qed | 
| 
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: 
70178diff
changeset | 4004 | then obtain h where "pathin X h" "h 0 = u" "h 1 = v" | 
| 
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: 
70178diff
changeset | 4005 | using assms unfolding path_connected_space_def by blast | 
| 
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: 
70178diff
changeset | 4006 | obtain C where "compactin X C" "connectedin X C" "u \<in> C" "v \<in> C" | 
| 
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: 
70178diff
changeset | 4007 | proof | 
| 
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: 
70178diff
changeset | 4008 |               show "compactin X (h ` {0..1})"
 | 
| 
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: 
70178diff
changeset | 4009 | using that by (simp add: \<open>pathin X h\<close> compactin_path_image) | 
| 
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: 
70178diff
changeset | 4010 |               show "connectedin X (h ` {0..1})"
 | 
| 
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: 
70178diff
changeset | 4011 | using \<open>pathin X h\<close> connectedin_path_image by blast | 
| 
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: 
70178diff
changeset | 4012 | qed (use \<open>h 0 = u\<close> \<open>h 1 = v\<close> in auto) | 
| 
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: 
70178diff
changeset | 4013 |             have "continuous_map (subtopology euclideanreal (?S \<inter> {f x - d .. f x + d})) (subtopology X C) g"
 | 
| 
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: 
70178diff
changeset | 4014 | proof (rule continuous_inverse_map) | 
| 
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: 
70178diff
changeset | 4015 | show "compact_space (subtopology X C)" | 
| 
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: 
70178diff
changeset | 4016 | using \<open>compactin X C\<close> compactin_subspace by blast | 
| 
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: 
70178diff
changeset | 4017 | show "continuous_map (subtopology X C) euclideanreal 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: 
70178diff
changeset | 4018 | by (simp add: cont continuous_map_from_subtopology) | 
| 
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: 
70178diff
changeset | 4019 |               have "{f u .. f v} \<subseteq> f ` topspace (subtopology X C)"
 | 
| 
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: 
70178diff
changeset | 4020 | proof (rule connected_contains_Icc) | 
| 
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: 
70178diff
changeset | 4021 | show "connected (f ` topspace (subtopology X C))" | 
| 
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: 
70178diff
changeset | 4022 | using connectedin_continuous_map_image [OF cont] | 
| 
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: 
70178diff
changeset | 4023 | by (simp add: \<open>compactin X C\<close> \<open>connectedin X C\<close> compactin_subset_topspace inf_absorb2) | 
| 
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: 
70178diff
changeset | 4024 | show "f u \<in> f ` topspace (subtopology X C)" | 
| 
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: 
70178diff
changeset | 4025 | by (simp add: \<open>u \<in> C\<close> \<open>u \<in> topspace X\<close>) | 
| 
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: 
70178diff
changeset | 4026 | show "f v \<in> f ` topspace (subtopology X C)" | 
| 
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: 
70178diff
changeset | 4027 | by (simp add: \<open>v \<in> C\<close> \<open>v \<in> topspace X\<close>) | 
| 
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: 
70178diff
changeset | 4028 | qed | 
| 
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: 
70178diff
changeset | 4029 |               then show "f ` topspace X \<inter> {f x - d..f x + d} \<subseteq> f ` topspace (subtopology X C)"
 | 
| 
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: 
70178diff
changeset | 4030 | using sub_fuv by blast | 
| 
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: 
70178diff
changeset | 4031 | qed (auto simp: gf) | 
| 
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: 
70178diff
changeset | 4032 |             then have contg: "continuous_map (subtopology euclideanreal (?S \<inter> {f x - d .. f x + d})) X g"
 | 
| 
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: 
70178diff
changeset | 4033 | using continuous_map_in_subtopology by blast | 
| 
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: 
70178diff
changeset | 4034 |             have "\<exists>e>0. \<forall>x \<in> ?S \<inter> {f x - d .. f x + d} \<inter> ball (f x) e. g x \<in> T"
 | 
| 
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: 
70178diff
changeset | 4035 | using openin_continuous_map_preimage [OF contg \<open>openin X T\<close>] x \<open>x \<in> T\<close> \<open>0 < d\<close> | 
| 
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: 
70178diff
changeset | 4036 | unfolding openin_euclidean_subtopology_iff | 
| 
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: 
70178diff
changeset | 4037 | by (force simp: gf dist_commute) | 
| 
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: 
70178diff
changeset | 4038 |             then obtain e where "e > 0 \<and> (\<forall>x\<in>f ` topspace X \<inter> {f x - d..f x + d} \<inter> ball (f x) e. g x \<in> T)"
 | 
| 
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: 
70178diff
changeset | 4039 | by metis | 
| 
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: 
70178diff
changeset | 4040 | with \<open>0 < d\<close> have "min d e > 0" "\<forall>u. u \<in> topspace X \<longrightarrow> \<bar>f x - f u\<bar> < min d e \<longrightarrow> u \<in> T" | 
| 
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: 
70178diff
changeset | 4041 | using dist_real_def gf by force+ | 
| 
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: 
70178diff
changeset | 4042 | then show ?thesis | 
| 
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: 
70178diff
changeset | 4043 | by (metis (full_types) Int_iff dist_real_def image_iff mem_ball gf) | 
| 
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: 
70178diff
changeset | 4044 | qed | 
| 
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: 
70178diff
changeset | 4045 | qed | 
| 
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: 
70178diff
changeset | 4046 | then obtain d where d: "\<And>r. r \<in> ?S \<inter> g -` T \<Longrightarrow> | 
| 
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: 
70178diff
changeset | 4047 | d r > 0 \<and> (\<forall>x \<in> ?S \<inter> ball r (d r). g x \<in> T)" | 
| 
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: 
70178diff
changeset | 4048 | by metis | 
| 
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: 
70178diff
changeset | 4049 | show ?thesis | 
| 
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: 
70178diff
changeset | 4050 | unfolding openin_subtopology | 
| 
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: 
70178diff
changeset | 4051 | proof (intro exI conjI) | 
| 
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: 
70178diff
changeset | 4052 |           show "{x \<in> ?S. g x \<in> T} = (\<Union>r \<in> ?S \<inter> g -` T. ball r (d r)) \<inter> f ` topspace X"
 | 
| 
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: 
70178diff
changeset | 4053 | using d by (auto simp: gf) | 
| 
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: 
70178diff
changeset | 4054 | qed auto | 
| 
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: 
70178diff
changeset | 4055 | qed | 
| 
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: 
70178diff
changeset | 4056 | then show "continuous_map (top_of_set ?S) X g" | 
| 
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: 
70178diff
changeset | 4057 | by (simp add: continuous_map_def gf) | 
| 
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: 
70178diff
changeset | 4058 | qed (auto simp: gf) | 
| 
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: 
70178diff
changeset | 4059 | qed | 
| 
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: 
70178diff
changeset | 4060 | qed | 
| 
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: 
70178diff
changeset | 4061 | |
| 
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: 
70178diff
changeset | 4062 | subsubsection \<open>An injective function into R is a homeomorphism and so an open map.\<close> | 
| 
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: 
70178diff
changeset | 4063 | |
| 
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: 
70178diff
changeset | 4064 | lemma injective_into_1d_eq_homeomorphism: | 
| 
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: 
70178diff
changeset | 4065 | fixes f :: "'a::topological_space \<Rightarrow> real" | 
| 
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: 
70178diff
changeset | 4066 | assumes f: "continuous_on S f" and S: "path_connected S" | 
| 
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: 
70178diff
changeset | 4067 | shows "inj_on f S \<longleftrightarrow> (\<exists>g. homeomorphism S (f ` S) f g)" | 
| 
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: 
70178diff
changeset | 4068 | proof | 
| 
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: 
70178diff
changeset | 4069 | show "\<exists>g. homeomorphism S (f ` S) f g" | 
| 
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: 
70178diff
changeset | 4070 | if "inj_on f S" | 
| 
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: 
70178diff
changeset | 4071 | proof - | 
| 
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: 
70178diff
changeset | 4072 | have "embedding_map (top_of_set S) euclideanreal 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: 
70178diff
changeset | 4073 | using that embedding_map_into_euclideanreal [of "top_of_set S" f] assms by auto | 
| 
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: 
70178diff
changeset | 4074 | then show ?thesis | 
| 
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: 
70178diff
changeset | 4075 | by (simp add: embedding_map_def) (metis all_closedin_homeomorphic_image f homeomorphism_injective_closed_map that) | 
| 
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: 
70178diff
changeset | 4076 | qed | 
| 
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: 
70178diff
changeset | 4077 | qed (metis homeomorphism_def inj_onI) | 
| 
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: 
70178diff
changeset | 4078 | |
| 
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: 
70178diff
changeset | 4079 | lemma injective_into_1d_imp_open_map: | 
| 
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: 
70178diff
changeset | 4080 | fixes f :: "'a::topological_space \<Rightarrow> real" | 
| 
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: 
70178diff
changeset | 4081 | assumes "continuous_on S f" "path_connected S" "inj_on f S" "openin (subtopology euclidean S) T" | 
| 
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: 
70178diff
changeset | 4082 | shows "openin (subtopology euclidean (f ` S)) (f ` T)" | 
| 
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: 
70178diff
changeset | 4083 | using assms homeomorphism_imp_open_map injective_into_1d_eq_homeomorphism by blast | 
| 
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: 
70178diff
changeset | 4084 | |
| 
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: 
70178diff
changeset | 4085 | lemma homeomorphism_into_1d: | 
| 
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: 
70178diff
changeset | 4086 | fixes f :: "'a::topological_space \<Rightarrow> real" | 
| 
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: 
70178diff
changeset | 4087 | assumes "path_connected S" "continuous_on S f" "f ` S = T" "inj_on f S" | 
| 
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: 
70178diff
changeset | 4088 | shows "\<exists>g. homeomorphism S T f g" | 
| 
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: 
70178diff
changeset | 4089 | using assms injective_into_1d_eq_homeomorphism by blast | 
| 
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: 
70178diff
changeset | 4090 | |
| 71189 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4091 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Rectangular paths\<close> | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4092 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4093 | definition\<^marker>\<open>tag unimportant\<close> rectpath where | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4094 | "rectpath a1 a3 = (let a2 = Complex (Re a3) (Im a1); a4 = Complex (Re a1) (Im a3) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4095 | in linepath a1 a2 +++ linepath a2 a3 +++ linepath a3 a4 +++ linepath a4 a1)" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4096 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4097 | lemma path_rectpath [simp, intro]: "path (rectpath a b)" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4098 | by (simp add: Let_def rectpath_def) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4099 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4100 | lemma pathstart_rectpath [simp]: "pathstart (rectpath a1 a3) = a1" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4101 | by (simp add: rectpath_def Let_def) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4102 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4103 | lemma pathfinish_rectpath [simp]: "pathfinish (rectpath a1 a3) = a1" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4104 | by (simp add: rectpath_def Let_def) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4105 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4106 | lemma simple_path_rectpath [simp, intro]: | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4107 | assumes "Re a1 \<noteq> Re a3" "Im a1 \<noteq> Im a3" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4108 | shows "simple_path (rectpath a1 a3)" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4109 | unfolding rectpath_def Let_def using assms | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4110 | by (intro simple_path_join_loop arc_join arc_linepath) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4111 | (auto simp: complex_eq_iff path_image_join closed_segment_same_Re closed_segment_same_Im) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4112 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4113 | lemma path_image_rectpath: | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4114 | assumes "Re a1 \<le> Re a3" "Im a1 \<le> Im a3" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4115 | shows "path_image (rectpath a1 a3) = | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4116 |            {z. Re z \<in> {Re a1, Re a3} \<and> Im z \<in> {Im a1..Im a3}} \<union>
 | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4117 |            {z. Im z \<in> {Im a1, Im a3} \<and> Re z \<in> {Re a1..Re a3}}" (is "?lhs = ?rhs")
 | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4118 | proof - | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4119 | define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4120 | have "?lhs = closed_segment a1 a2 \<union> closed_segment a2 a3 \<union> | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4121 | closed_segment a4 a3 \<union> closed_segment a1 a4" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4122 | by (simp_all add: rectpath_def Let_def path_image_join closed_segment_commute | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4123 | a2_def a4_def Un_assoc) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4124 | also have "\<dots> = ?rhs" using assms | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4125 | by (auto simp: rectpath_def Let_def path_image_join a2_def a4_def | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4126 | closed_segment_same_Re closed_segment_same_Im closed_segment_eq_real_ivl) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4127 | finally show ?thesis . | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4128 | qed | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4129 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4130 | lemma path_image_rectpath_subset_cbox: | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4131 | assumes "Re a \<le> Re b" "Im a \<le> Im b" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4132 | shows "path_image (rectpath a b) \<subseteq> cbox a b" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4133 | using assms by (auto simp: path_image_rectpath in_cbox_complex_iff) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4134 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4135 | lemma path_image_rectpath_inter_box: | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4136 | assumes "Re a \<le> Re b" "Im a \<le> Im b" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4137 |   shows   "path_image (rectpath a b) \<inter> box a b = {}"
 | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4138 | using assms by (auto simp: path_image_rectpath in_box_complex_iff) | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4139 | |
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4140 | lemma path_image_rectpath_cbox_minus_box: | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4141 | assumes "Re a \<le> Re b" "Im a \<le> Im b" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4142 | shows "path_image (rectpath a b) = cbox a b - box a b" | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4143 | using assms by (auto simp: path_image_rectpath in_cbox_complex_iff | 
| 
954ee5acaae0
Split off new HOL-Complex_Analysis session from HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71172diff
changeset | 4144 | in_box_complex_iff) | 
| 71184 
d62fdaafdafc
renamed Analysis/Winding_Numbers to Winding_Numbers_2; reorganised Analysis/Cauchy_Integral_Theorem by splitting it into Contour_Integration, Winding_Numbers,Cauchy_Integral_Theorem and Cauchy_Integral_Formula.
 Wenda Li <wl302@cam.ac.uk> parents: 
71172diff
changeset | 4145 | |
| 36583 | 4146 | end |