| author | wenzelm | 
| Sun, 04 Aug 2024 12:21:13 +0200 | |
| changeset 80632 | 3a196e63a80d | 
| parent 78480 | b22f39c54e8c | 
| permissions | -rw-r--r-- | 
| 64846 | 1 | (* Title: HOL/Analysis/Jordan_Curve.thy | 
| 2 | Authors: LC Paulson, based on material from HOL Light | |
| 3 | *) | |
| 4 | ||
| 5 | section \<open>The Jordan Curve Theorem and Applications\<close> | |
| 6 | ||
| 7 | theory Jordan_Curve | |
| 8 | imports Arcwise_Connected Further_Topology | |
| 9 | begin | |
| 10 | ||
| 67968 | 11 | subsection\<open>Janiszewski's theorem\<close> | 
| 64846 | 12 | |
| 13 | lemma Janiszewski_weak: | |
| 14 | fixes a b::complex | |
| 15 | assumes "compact S" "compact T" and conST: "connected(S \<inter> T)" | |
| 16 | and ccS: "connected_component (- S) a b" and ccT: "connected_component (- T) a b" | |
| 17 | shows "connected_component (- (S \<union> T)) a b" | |
| 18 | proof - | |
| 19 | have [simp]: "a \<notin> S" "a \<notin> T" "b \<notin> S" "b \<notin> T" | |
| 20 | by (meson ComplD ccS ccT connected_component_in)+ | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69508diff
changeset | 21 | have clo: "closedin (top_of_set (S \<union> T)) S" "closedin (top_of_set (S \<union> T)) T" | 
| 64846 | 22 | by (simp_all add: assms closed_subset compact_imp_closed) | 
| 23 | obtain g where contg: "continuous_on S g" | |
| 24 | and g: "\<And>x. x \<in> S \<Longrightarrow> exp (\<i>* of_real (g x)) = (x - a) /\<^sub>R cmod (x - a) / ((x - b) /\<^sub>R cmod (x - b))" | |
| 25 | using ccS \<open>compact S\<close> | |
| 26 | apply (simp add: Borsuk_maps_homotopic_in_connected_component_eq [symmetric]) | |
| 27 | apply (subst (asm) homotopic_circlemaps_divide) | |
| 28 | apply (auto simp: inessential_eq_continuous_logarithm_circle) | |
| 29 | done | |
| 30 | obtain h where conth: "continuous_on T h" | |
| 31 | and h: "\<And>x. x \<in> T \<Longrightarrow> exp (\<i>* of_real (h x)) = (x - a) /\<^sub>R cmod (x - a) / ((x - b) /\<^sub>R cmod (x - b))" | |
| 32 | using ccT \<open>compact T\<close> | |
| 33 | apply (simp add: Borsuk_maps_homotopic_in_connected_component_eq [symmetric]) | |
| 34 | apply (subst (asm) homotopic_circlemaps_divide) | |
| 35 | apply (auto simp: inessential_eq_continuous_logarithm_circle) | |
| 36 | done | |
| 37 | have "continuous_on (S \<union> T) (\<lambda>x. (x - a) /\<^sub>R cmod (x - a))" "continuous_on (S \<union> T) (\<lambda>x. (x - b) /\<^sub>R cmod (x - b))" | |
| 38 | by (intro continuous_intros; force)+ | |
| 39 | moreover have "(\<lambda>x. (x - a) /\<^sub>R cmod (x - a)) ` (S \<union> T) \<subseteq> sphere 0 1" "(\<lambda>x. (x - b) /\<^sub>R cmod (x - b)) ` (S \<union> T) \<subseteq> sphere 0 1" | |
| 40 | by (auto simp: divide_simps) | |
| 41 | moreover have "\<exists>g. continuous_on (S \<union> T) g \<and> | |
| 42 | (\<forall>x\<in>S \<union> T. (x - a) /\<^sub>R cmod (x - a) / ((x - b) /\<^sub>R cmod (x - b)) = exp (\<i>*complex_of_real (g x)))" | |
| 43 |   proof (cases "S \<inter> T = {}")
 | |
| 44 | case True | |
| 78480 | 45 | then have "continuous_on (S \<union> T) (\<lambda>x. if x \<in> S then g x else h x)" | 
| 46 | using continuous_on_cases_local [OF clo contg conth] | |
| 47 | by (meson disjoint_iff) | |
| 64846 | 48 | then show ?thesis | 
| 49 | by (rule_tac x="(\<lambda>x. if x \<in> S then g x else h x)" in exI) (auto simp: g h) | |
| 50 | next | |
| 51 | case False | |
| 52 | have diffpi: "\<exists>n. g x = h x + 2* of_int n*pi" if "x \<in> S \<inter> T" for x | |
| 53 | proof - | |
| 54 | have "exp (\<i>* of_real (g x)) = exp (\<i>* of_real (h x))" | |
| 55 | using that by (simp add: g h) | |
| 56 | then obtain n where "complex_of_real (g x) = complex_of_real (h x) + 2* of_int n*complex_of_real pi" | |
| 78480 | 57 | apply (simp add: exp_eq) | 
| 64846 | 58 | by (metis complex_i_not_zero distrib_left mult.commute mult_cancel_left) | 
| 59 | then show ?thesis | |
| 78480 | 60 | using of_real_eq_iff by (fastforce intro!: exI [where x=n]) | 
| 64846 | 61 | qed | 
| 62 | have contgh: "continuous_on (S \<inter> T) (\<lambda>x. g x - h x)" | |
| 63 | by (intro continuous_intros continuous_on_subset [OF contg] continuous_on_subset [OF conth]) auto | |
| 64 | moreover have disc: | |
| 65 | "\<exists>e>0. \<forall>y. y \<in> S \<inter> T \<and> g y - h y \<noteq> g x - h x \<longrightarrow> e \<le> norm ((g y - h y) - (g x - h x))" | |
| 66 | if "x \<in> S \<inter> T" for x | |
| 67 | proof - | |
| 68 | obtain nx where nx: "g x = h x + 2* of_int nx*pi" | |
| 69 | using \<open>x \<in> S \<inter> T\<close> diffpi by blast | |
| 70 | have "2*pi \<le> norm (g y - h y - (g x - h x))" if y: "y \<in> S \<inter> T" and neq: "g y - h y \<noteq> g x - h x" for y | |
| 71 | proof - | |
| 72 | obtain ny where ny: "g y = h y + 2* of_int ny*pi" | |
| 73 | using \<open>y \<in> S \<inter> T\<close> diffpi by blast | |
| 74 |         { assume "nx \<noteq> ny"
 | |
| 75 | then have "1 \<le> \<bar>real_of_int ny - real_of_int nx\<bar>" | |
| 76 | by linarith | |
| 77 | then have "(2*pi)*1 \<le> (2*pi)*\<bar>real_of_int ny - real_of_int nx\<bar>" | |
| 78 | by simp | |
| 79 | also have "... = \<bar>2*real_of_int ny*pi - 2*real_of_int nx*pi\<bar>" | |
| 80 | by (simp add: algebra_simps abs_if) | |
| 81 | finally have "2*pi \<le> \<bar>2*real_of_int ny*pi - 2*real_of_int nx*pi\<bar>" by simp | |
| 82 | } | |
| 83 | with neq show ?thesis | |
| 84 | by (simp add: nx ny) | |
| 85 | qed | |
| 86 | then show ?thesis | |
| 87 | by (rule_tac x="2*pi" in exI) auto | |
| 88 | qed | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65064diff
changeset | 89 | ultimately have "(\<lambda>x. g x - h x) constant_on S \<inter> T" | 
| 64846 | 90 | using continuous_discrete_range_constant [OF conST contgh] by blast | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65064diff
changeset | 91 | then obtain z where z: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> g x - h x = z" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65064diff
changeset | 92 | by (auto simp: constant_on_def) | 
| 65064 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 paulson <lp15@cam.ac.uk> parents: 
64846diff
changeset | 93 | obtain w where "exp(\<i> * of_real(h w)) = exp (\<i> * of_real(z + h w))" | 
| 64846 | 94 | using disc z False | 
| 95 | by auto (metis diff_add_cancel g h of_real_add) | |
| 96 | then have [simp]: "exp (\<i>* of_real z) = 1" | |
| 97 | by (metis cis_conv_exp cis_mult exp_not_eq_zero mult_cancel_right1) | |
| 98 | show ?thesis | |
| 99 | proof (intro exI conjI) | |
| 100 | show "continuous_on (S \<union> T) (\<lambda>x. if x \<in> S then g x else z + h x)" | |
| 78480 | 101 | by (intro continuous_intros continuous_on_cases_local [OF clo contg] conth) (use z in force) | 
| 64846 | 102 | qed (auto simp: g h algebra_simps exp_add) | 
| 103 | qed | |
| 78480 | 104 | ultimately have "homotopic_with_canon (\<lambda>x. True) (S \<union> T) (sphere 0 1) | 
| 64846 | 105 | (\<lambda>x. (x - a) /\<^sub>R cmod (x - a)) (\<lambda>x. (x - b) /\<^sub>R cmod (x - b))" | 
| 106 | by (subst homotopic_circlemaps_divide) (auto simp: inessential_eq_continuous_logarithm_circle) | |
| 78480 | 107 | moreover have "compact (S \<union> T)" | 
| 108 | using assms by blast | |
| 109 | ultimately show ?thesis | |
| 110 | using assms Borsuk_maps_homotopic_in_connected_component_eq by fastforce | |
| 64846 | 111 | qed | 
| 112 | ||
| 113 | ||
| 114 | theorem Janiszewski: | |
| 68651 | 115 | fixes a b :: complex | 
| 116 | assumes "compact S" "closed T" and conST: "connected (S \<inter> T)" | |
| 64846 | 117 | and ccS: "connected_component (- S) a b" and ccT: "connected_component (- T) a b" | 
| 118 | shows "connected_component (- (S \<union> T)) a b" | |
| 119 | proof - | |
| 120 | have "path_component(- T) a b" | |
| 121 | by (simp add: \<open>closed T\<close> ccT open_Compl open_path_connected_component) | |
| 122 | then obtain g where g: "path g" "path_image g \<subseteq> - T" "pathstart g = a" "pathfinish g = b" | |
| 123 | by (auto simp: path_component_def) | |
| 78480 | 124 |   then obtain C where C: "compact C" "connected C" "a \<in> C" "b \<in> C" "C \<inter> T = {}"
 | 
| 125 | by fastforce | |
| 64846 | 126 | obtain r where "0 < r" and r: "C \<union> S \<subseteq> ball 0 r" | 
| 127 | by (metis \<open>compact C\<close> \<open>compact S\<close> bounded_Un compact_imp_bounded bounded_subset_ballD) | |
| 128 | have "connected_component (- (S \<union> (T \<inter> cball 0 r \<union> sphere 0 r))) a b" | |
| 129 | proof (rule Janiszewski_weak [OF \<open>compact S\<close>]) | |
| 130 | show comT': "compact ((T \<inter> cball 0 r) \<union> sphere 0 r)" | |
| 131 | by (simp add: \<open>closed T\<close> closed_Int_compact compact_Un) | |
| 132 | have "S \<inter> (T \<inter> cball 0 r \<union> sphere 0 r) = S \<inter> T" | |
| 133 | using r by auto | |
| 134 | with conST show "connected (S \<inter> (T \<inter> cball 0 r \<union> sphere 0 r))" | |
| 135 | by simp | |
| 136 | show "connected_component (- (T \<inter> cball 0 r \<union> sphere 0 r)) a b" | |
| 137 | using conST C r | |
| 138 | apply (simp add: connected_component_def) | |
| 139 | apply (rule_tac x=C in exI) | |
| 140 | by auto | |
| 141 | qed (simp add: ccS) | |
| 142 | then obtain U where U: "connected U" "U \<subseteq> - S" "U \<subseteq> - T \<union> - cball 0 r" "U \<subseteq> - sphere 0 r" "a \<in> U" "b \<in> U" | |
| 143 | by (auto simp: connected_component_def) | |
| 144 | show ?thesis | |
| 145 | unfolding connected_component_def | |
| 146 | proof (intro exI conjI) | |
| 147 | show "U \<subseteq> - (S \<union> T)" | |
| 148 | using U r \<open>0 < r\<close> \<open>a \<in> C\<close> connected_Int_frontier [of U "cball 0 r"] | |
| 149 | apply simp | |
| 150 | by (metis ball_subset_cball compl_inf disjoint_eq_subset_Compl disjoint_iff_not_equal inf.orderE inf_sup_aci(3) subsetCE) | |
| 151 | qed (auto simp: U) | |
| 152 | qed | |
| 153 | ||
| 154 | lemma Janiszewski_connected: | |
| 155 | fixes S :: "complex set" | |
| 156 | assumes ST: "compact S" "closed T" "connected(S \<inter> T)" | |
| 78480 | 157 | and notST: "connected (- S)" "connected (- T)" | 
| 158 | shows "connected(- (S \<union> T))" | |
| 159 | using Janiszewski [OF ST] by (metis IntD1 IntD2 notST compl_sup connected_iff_connected_component) | |
| 64846 | 160 | |
| 68651 | 161 | |
| 64846 | 162 | subsection\<open>The Jordan Curve theorem\<close> | 
| 163 | ||
| 164 | lemma exists_double_arc: | |
| 165 | fixes g :: "real \<Rightarrow> 'a::real_normed_vector" | |
| 166 | assumes "simple_path g" "pathfinish g = pathstart g" "a \<in> path_image g" "b \<in> path_image g" "a \<noteq> b" | |
| 167 | obtains u d where "arc u" "arc d" "pathstart u = a" "pathfinish u = b" | |
| 168 | "pathstart d = b" "pathfinish d = a" | |
| 169 |                     "(path_image u) \<inter> (path_image d) = {a,b}"
 | |
| 170 | "(path_image u) \<union> (path_image d) = path_image g" | |
| 171 | proof - | |
| 172 | obtain u where u: "0 \<le> u" "u \<le> 1" "g u = a" | |
| 173 | using assms by (auto simp: path_image_def) | |
| 174 | define h where "h \<equiv> shiftpath u g" | |
| 175 | have "simple_path h" | |
| 176 | using \<open>simple_path g\<close> simple_path_shiftpath \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> assms(2) h_def by blast | |
| 177 | have "pathstart h = g u" | |
| 178 | by (simp add: \<open>u \<le> 1\<close> h_def pathstart_shiftpath) | |
| 179 | have "pathfinish h = g u" | |
| 180 | by (simp add: \<open>0 \<le> u\<close> assms h_def pathfinish_shiftpath) | |
| 181 | have pihg: "path_image h = path_image g" | |
| 182 | by (simp add: \<open>0 \<le> u\<close> \<open>u \<le> 1\<close> assms h_def path_image_shiftpath) | |
| 183 | then obtain v where v: "0 \<le> v" "v \<le> 1" "h v = b" | |
| 184 | using assms by (metis (mono_tags, lifting) atLeastAtMost_iff imageE path_image_def) | |
| 185 | show ?thesis | |
| 186 | proof | |
| 187 | show "arc (subpath 0 v h)" | |
| 188 | by (metis (no_types) \<open>pathstart h = g u\<close> \<open>simple_path h\<close> arc_simple_path_subpath \<open>a \<noteq> b\<close> atLeastAtMost_iff zero_le_one order_refl pathstart_def u(3) v) | |
| 189 | show "arc (subpath v 1 h)" | |
| 190 | by (metis (no_types) \<open>pathfinish h = g u\<close> \<open>simple_path h\<close> arc_simple_path_subpath \<open>a \<noteq> b\<close> atLeastAtMost_iff zero_le_one order_refl pathfinish_def u(3) v) | |
| 191 | show "pathstart (subpath 0 v h) = a" | |
| 192 | by (metis \<open>pathstart h = g u\<close> pathstart_def pathstart_subpath u(3)) | |
| 193 | show "pathfinish (subpath 0 v h) = b" "pathstart (subpath v 1 h) = b" | |
| 194 | by (simp_all add: v(3)) | |
| 195 | show "pathfinish (subpath v 1 h) = a" | |
| 196 | by (metis \<open>pathfinish h = g u\<close> pathfinish_def pathfinish_subpath u(3)) | |
| 197 |     show "path_image (subpath 0 v h) \<inter> path_image (subpath v 1 h) = {a, b}"
 | |
| 198 | proof | |
| 76874 
d6b02d54dbf8
Tidying up of paths, introducing "loop_free" as a separate predicate in the definition of "simple_path"
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 199 | have "loop_free h" | 
| 
d6b02d54dbf8
Tidying up of paths, introducing "loop_free" as a separate predicate in the definition of "simple_path"
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 200 | using \<open>simple_path h\<close> simple_path_def by blast | 
| 
d6b02d54dbf8
Tidying up of paths, introducing "loop_free" as a separate predicate in the definition of "simple_path"
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 201 |       then show "path_image (subpath 0 v h) \<inter> path_image (subpath v 1 h) \<subseteq> {a, b}"
 | 
| 
d6b02d54dbf8
Tidying up of paths, introducing "loop_free" as a separate predicate in the definition of "simple_path"
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 202 | using v \<open>pathfinish (subpath v 1 h) = a\<close> | 
| 
d6b02d54dbf8
Tidying up of paths, introducing "loop_free" as a separate predicate in the definition of "simple_path"
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 203 | apply (clarsimp simp add: loop_free_def path_image_subpath Ball_def) | 
| 
d6b02d54dbf8
Tidying up of paths, introducing "loop_free" as a separate predicate in the definition of "simple_path"
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 204 | by (smt (verit)) | 
| 64846 | 205 |       show "{a, b} \<subseteq> path_image (subpath 0 v h) \<inter> path_image (subpath v 1 h)"
 | 
| 206 | using v \<open>pathstart (subpath 0 v h) = a\<close> \<open>pathfinish (subpath v 1 h) = a\<close> | |
| 76874 
d6b02d54dbf8
Tidying up of paths, introducing "loop_free" as a separate predicate in the definition of "simple_path"
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 207 | by (auto simp: path_image_subpath image_iff Bex_def) | 
| 64846 | 208 | qed | 
| 209 | show "path_image (subpath 0 v h) \<union> path_image (subpath v 1 h) = path_image g" | |
| 78480 | 210 | using v path_image_subpath pihg path_image_def | 
| 211 | by (metis (full_types) image_Un ivl_disj_un_two_touch(4)) | |
| 64846 | 212 | qed | 
| 213 | qed | |
| 214 | ||
| 215 | ||
| 70136 | 216 | theorem\<^marker>\<open>tag unimportant\<close> Jordan_curve: | 
| 64846 | 217 | fixes c :: "real \<Rightarrow> complex" | 
| 218 | assumes "simple_path c" and loop: "pathfinish c = pathstart c" | |
| 219 | obtains inner outer where | |
| 220 |                 "inner \<noteq> {}" "open inner" "connected inner"
 | |
| 221 |                 "outer \<noteq> {}" "open outer" "connected outer"
 | |
| 222 |                 "bounded inner" "\<not> bounded outer" "inner \<inter> outer = {}"
 | |
| 223 | "inner \<union> outer = - path_image c" | |
| 224 | "frontier inner = path_image c" | |
| 225 | "frontier outer = path_image c" | |
| 226 | proof - | |
| 227 | have "path c" | |
| 228 | by (simp add: assms simple_path_imp_path) | |
| 229 | have hom: "(path_image c) homeomorphic (sphere(0::complex) 1)" | |
| 230 | by (simp add: assms homeomorphic_simple_path_image_circle) | |
| 231 | with Jordan_Brouwer_separation have "\<not> connected (- (path_image c))" | |
| 232 | by fastforce | |
| 233 | then obtain inner where inner: "inner \<in> components (- path_image c)" and "bounded inner" | |
| 234 | using cobounded_has_bounded_component [of "- (path_image c)"] | |
| 235 | using \<open>\<not> connected (- path_image c)\<close> \<open>simple_path c\<close> bounded_simple_path_image by force | |
| 236 | obtain outer where outer: "outer \<in> components (- path_image c)" and "\<not> bounded outer" | |
| 237 | using cobounded_unbounded_components [of "- (path_image c)"] | |
| 238 | using \<open>path c\<close> bounded_path_image by auto | |
| 239 | show ?thesis | |
| 240 | proof | |
| 241 |     show "inner \<noteq> {}"
 | |
| 242 | using inner in_components_nonempty by auto | |
| 243 | show "open inner" | |
| 244 | by (meson \<open>simple_path c\<close> compact_imp_closed compact_simple_path_image inner open_Compl open_components) | |
| 245 | show "connected inner" | |
| 246 | using in_components_connected inner by blast | |
| 247 |     show "outer \<noteq> {}"
 | |
| 248 | using outer in_components_nonempty by auto | |
| 249 | show "open outer" | |
| 250 | by (meson \<open>simple_path c\<close> compact_imp_closed compact_simple_path_image outer open_Compl open_components) | |
| 251 | show "connected outer" | |
| 252 | using in_components_connected outer by blast | |
| 78480 | 253 |     show inner_outer: "inner \<inter> outer = {}"
 | 
| 64846 | 254 | by (meson \<open>\<not> bounded outer\<close> \<open>bounded inner\<close> \<open>connected outer\<close> bounded_subset components_maximal in_components_subset inner outer) | 
| 255 | show fro_inner: "frontier inner = path_image c" | |
| 256 | by (simp add: Jordan_Brouwer_frontier [OF hom inner]) | |
| 257 | show fro_outer: "frontier outer = path_image c" | |
| 258 | by (simp add: Jordan_Brouwer_frontier [OF hom outer]) | |
| 259 | have False if m: "middle \<in> components (- path_image c)" and "middle \<noteq> inner" "middle \<noteq> outer" for middle | |
| 260 | proof - | |
| 261 | have "frontier middle = path_image c" | |
| 262 | by (simp add: Jordan_Brouwer_frontier [OF hom] that) | |
| 78480 | 263 |       obtain middle: "open middle" "connected middle" "middle \<noteq> {}"
 | 
| 264 | by (metis fro_inner frontier_closed in_components_maximal m open_Compl open_components) | |
| 64846 | 265 | obtain a0 b0 where "a0 \<in> path_image c" "b0 \<in> path_image c" "a0 \<noteq> b0" | 
| 266 | using simple_path_image_uncountable [OF \<open>simple_path c\<close>] | |
| 267 | by (metis Diff_cancel countable_Diff_eq countable_empty insert_iff subsetI subset_singleton_iff) | |
| 268 | obtain a b g where ab: "a \<in> path_image c" "b \<in> path_image c" "a \<noteq> b" | |
| 78480 | 269 | and g: "arc g" "pathstart g = a" "pathfinish g = b" | 
| 64846 | 270 |                      and pag_sub: "path_image g - {a,b} \<subseteq> middle"
 | 
| 271 | proof (rule dense_accessible_frontier_point_pairs [OF \<open>open middle\<close> \<open>connected middle\<close>, of "path_image c \<inter> ball a0 (dist a0 b0)" "path_image c \<inter> ball b0 (dist a0 b0)"]) | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69508diff
changeset | 272 | show "openin (top_of_set (frontier middle)) (path_image c \<inter> ball a0 (dist a0 b0))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69508diff
changeset | 273 | "openin (top_of_set (frontier middle)) (path_image c \<inter> ball b0 (dist a0 b0))" | 
| 64846 | 274 | by (simp_all add: \<open>frontier middle = path_image c\<close> openin_open_Int) | 
| 275 | show "path_image c \<inter> ball a0 (dist a0 b0) \<noteq> path_image c \<inter> ball b0 (dist a0 b0)" | |
| 276 | using \<open>a0 \<noteq> b0\<close> \<open>b0 \<in> path_image c\<close> by auto | |
| 277 |         show "path_image c \<inter> ball a0 (dist a0 b0) \<noteq> {}"
 | |
| 278 | using \<open>a0 \<in> path_image c\<close> \<open>a0 \<noteq> b0\<close> by auto | |
| 279 |         show "path_image c \<inter> ball b0 (dist a0 b0) \<noteq> {}"
 | |
| 280 | using \<open>b0 \<in> path_image c\<close> \<open>a0 \<noteq> b0\<close> by auto | |
| 281 | qed (use arc_distinct_ends arc_imp_simple_path simple_path_endless that in fastforce) | |
| 282 | obtain u d where "arc u" "arc d" | |
| 283 | and "pathstart u = a" "pathfinish u = b" "pathstart d = b" "pathfinish d = a" | |
| 284 |                    and ud_ab: "(path_image u) \<inter> (path_image d) = {a,b}"
 | |
| 285 | and ud_Un: "(path_image u) \<union> (path_image d) = path_image c" | |
| 286 | using exists_double_arc [OF assms ab] by blast | |
| 287 | obtain x y where "x \<in> inner" "y \<in> outer" | |
| 288 |         using \<open>inner \<noteq> {}\<close> \<open>outer \<noteq> {}\<close> by auto
 | |
| 289 |       have "inner \<inter> middle = {}" "middle \<inter> outer = {}"
 | |
| 290 | using components_nonoverlap inner outer m that by blast+ | |
| 291 | have "connected_component (- (path_image u \<union> path_image g \<union> (path_image d \<union> path_image g))) x y" | |
| 292 | proof (rule Janiszewski) | |
| 293 | show "compact (path_image u \<union> path_image g)" | |
| 294 | by (simp add: \<open>arc g\<close> \<open>arc u\<close> compact_Un compact_arc_image) | |
| 295 | show "closed (path_image d \<union> path_image g)" | |
| 296 | by (simp add: \<open>arc d\<close> \<open>arc g\<close> closed_Un closed_arc_image) | |
| 297 | show "connected ((path_image u \<union> path_image g) \<inter> (path_image d \<union> path_image g))" | |
| 78480 | 298 | using ud_ab | 
| 299 | by (metis Un_insert_left g connected_arc_image insert_absorb pathfinish_in_path_image pathstart_in_path_image sup_bot_left sup_commute sup_inf_distrib1) | |
| 64846 | 300 | show "connected_component (- (path_image u \<union> path_image g)) x y" | 
| 301 | unfolding connected_component_def | |
| 302 | proof (intro exI conjI) | |
| 303 | have "connected ((inner \<union> (path_image c - path_image u)) \<union> (outer \<union> (path_image c - path_image u)))" | |
| 304 | proof (rule connected_Un) | |
| 305 | show "connected (inner \<union> (path_image c - path_image u))" | |
| 78480 | 306 | using connected_intermediate_closure [OF \<open>connected inner\<close>] | 
| 307 | by (metis Diff_subset closure_Un_frontier dual_order.refl fro_inner sup.mono sup_ge1) | |
| 64846 | 308 | show "connected (outer \<union> (path_image c - path_image u))" | 
| 78480 | 309 | using connected_intermediate_closure [OF \<open>connected outer\<close>] | 
| 310 | by (simp add: Diff_eq closure_Un_frontier fro_outer sup_inf_distrib1) | |
| 64846 | 311 |             have "(inner \<inter> outer) \<union> (path_image c - path_image u) \<noteq> {}"
 | 
| 78480 | 312 | using \<open>arc d\<close> \<open>pathfinish d = a\<close> \<open>pathstart d = b\<close> arc_imp_simple_path nonempty_simple_path_endless ud_Un ud_ab by fastforce | 
| 64846 | 313 |             then show "(inner \<union> (path_image c - path_image u)) \<inter> (outer \<union> (path_image c - path_image u)) \<noteq> {}"
 | 
| 314 | by auto | |
| 315 | qed | |
| 316 | then show "connected (inner \<union> outer \<union> (path_image c - path_image u))" | |
| 317 | by (metis sup.right_idem sup_assoc sup_commute) | |
| 318 | have "inner \<subseteq> - path_image u" "outer \<subseteq> - path_image u" | |
| 319 | using in_components_subset inner outer ud_Un by auto | |
| 320 | moreover have "inner \<subseteq> - path_image g" "outer \<subseteq> - path_image g" | |
| 321 |             using \<open>inner \<inter> middle = {}\<close> \<open>inner \<subseteq> - path_image u\<close>
 | |
| 322 |             using \<open>middle \<inter> outer = {}\<close> \<open>outer \<subseteq> - path_image u\<close> pag_sub ud_ab by fastforce+
 | |
| 323 | moreover have "path_image c - path_image u \<subseteq> - path_image g" | |
| 324 | using in_components_subset m pag_sub ud_ab by fastforce | |
| 325 | ultimately show "inner \<union> outer \<union> (path_image c - path_image u) \<subseteq> - (path_image u \<union> path_image g)" | |
| 326 | by force | |
| 327 | show "x \<in> inner \<union> outer \<union> (path_image c - path_image u)" | |
| 328 | by (auto simp: \<open>x \<in> inner\<close>) | |
| 329 | show "y \<in> inner \<union> outer \<union> (path_image c - path_image u)" | |
| 330 | by (auto simp: \<open>y \<in> outer\<close>) | |
| 331 | qed | |
| 332 | show "connected_component (- (path_image d \<union> path_image g)) x y" | |
| 333 | unfolding connected_component_def | |
| 334 | proof (intro exI conjI) | |
| 335 | have "connected ((inner \<union> (path_image c - path_image d)) \<union> (outer \<union> (path_image c - path_image d)))" | |
| 336 | proof (rule connected_Un) | |
| 337 | show "connected (inner \<union> (path_image c - path_image d))" | |
| 78480 | 338 | using connected_intermediate_closure [OF \<open>connected inner\<close>] fro_inner | 
| 339 | by (simp add: closure_Un_frontier sup.coboundedI2) | |
| 64846 | 340 | show "connected (outer \<union> (path_image c - path_image d))" | 
| 78480 | 341 | using connected_intermediate_closure [OF \<open>connected outer\<close>] | 
| 342 | by (simp add: closure_Un_frontier fro_outer sup.coboundedI2) | |
| 64846 | 343 |             have "(inner \<inter> outer) \<union> (path_image c - path_image d) \<noteq> {}"
 | 
| 344 | using \<open>arc u\<close> \<open>pathfinish u = b\<close> \<open>pathstart u = a\<close> arc_imp_simple_path nonempty_simple_path_endless ud_Un ud_ab by fastforce | |
| 345 |             then show "(inner \<union> (path_image c - path_image d)) \<inter> (outer \<union> (path_image c - path_image d)) \<noteq> {}"
 | |
| 346 | by auto | |
| 347 | qed | |
| 348 | then show "connected (inner \<union> outer \<union> (path_image c - path_image d))" | |
| 349 | by (metis sup.right_idem sup_assoc sup_commute) | |
| 350 | have "inner \<subseteq> - path_image d" "outer \<subseteq> - path_image d" | |
| 351 | using in_components_subset inner outer ud_Un by auto | |
| 352 | moreover have "inner \<subseteq> - path_image g" "outer \<subseteq> - path_image g" | |
| 353 |             using \<open>inner \<inter> middle = {}\<close> \<open>inner \<subseteq> - path_image d\<close>
 | |
| 354 |             using \<open>middle \<inter> outer = {}\<close> \<open>outer \<subseteq> - path_image d\<close> pag_sub ud_ab by fastforce+
 | |
| 355 | moreover have "path_image c - path_image d \<subseteq> - path_image g" | |
| 356 | using in_components_subset m pag_sub ud_ab by fastforce | |
| 357 | ultimately show "inner \<union> outer \<union> (path_image c - path_image d) \<subseteq> - (path_image d \<union> path_image g)" | |
| 358 | by force | |
| 359 | show "x \<in> inner \<union> outer \<union> (path_image c - path_image d)" | |
| 360 | by (auto simp: \<open>x \<in> inner\<close>) | |
| 361 | show "y \<in> inner \<union> outer \<union> (path_image c - path_image d)" | |
| 362 | by (auto simp: \<open>y \<in> outer\<close>) | |
| 363 | qed | |
| 364 | qed | |
| 365 | then have "connected_component (- (path_image u \<union> path_image d \<union> path_image g)) x y" | |
| 366 | by (simp add: Un_ac) | |
| 69508 | 367 | moreover have "\<not>(connected_component (- (path_image c)) x y)" | 
| 64846 | 368 | by (metis (no_types, lifting) \<open>\<not> bounded outer\<close> \<open>bounded inner\<close> \<open>x \<in> inner\<close> \<open>y \<in> outer\<close> componentsE connected_component_eq inner mem_Collect_eq outer) | 
| 369 | ultimately show False | |
| 370 | by (auto simp: ud_Un [symmetric] connected_component_def) | |
| 371 | qed | |
| 372 |     then have "components (- path_image c) = {inner,outer}"
 | |
| 373 | using inner outer by blast | |
| 374 | then have "Union (components (- path_image c)) = inner \<union> outer" | |
| 375 | by simp | |
| 376 | then show "inner \<union> outer = - path_image c" | |
| 377 | by auto | |
| 378 | qed (auto simp: \<open>bounded inner\<close> \<open>\<not> bounded outer\<close>) | |
| 379 | qed | |
| 380 | ||
| 381 | ||
| 70136 | 382 | corollary\<^marker>\<open>tag unimportant\<close> Jordan_disconnected: | 
| 64846 | 383 | fixes c :: "real \<Rightarrow> complex" | 
| 384 | assumes "simple_path c" "pathfinish c = pathstart c" | |
| 78480 | 385 | shows "\<not> connected(- path_image c)" | 
| 386 | using Jordan_curve [OF assms] | |
| 64846 | 387 | by (metis Jordan_Brouwer_separation assms homeomorphic_simple_path_image_circle zero_less_one) | 
| 388 | ||
| 389 | ||
| 390 | corollary Jordan_inside_outside: | |
| 391 | fixes c :: "real \<Rightarrow> complex" | |
| 392 | assumes "simple_path c" "pathfinish c = pathstart c" | |
| 393 |     shows "inside(path_image c) \<noteq> {} \<and>
 | |
| 394 | open(inside(path_image c)) \<and> | |
| 395 | connected(inside(path_image c)) \<and> | |
| 396 |           outside(path_image c) \<noteq> {} \<and>
 | |
| 397 | open(outside(path_image c)) \<and> | |
| 398 | connected(outside(path_image c)) \<and> | |
| 399 | bounded(inside(path_image c)) \<and> | |
| 400 | \<not> bounded(outside(path_image c)) \<and> | |
| 401 |           inside(path_image c) \<inter> outside(path_image c) = {} \<and>
 | |
| 402 | inside(path_image c) \<union> outside(path_image c) = | |
| 403 | - path_image c \<and> | |
| 404 | frontier(inside(path_image c)) = path_image c \<and> | |
| 405 | frontier(outside(path_image c)) = path_image c" | |
| 406 | proof - | |
| 407 | obtain inner outer | |
| 408 |     where *: "inner \<noteq> {}" "open inner" "connected inner"
 | |
| 409 |              "outer \<noteq> {}" "open outer" "connected outer"
 | |
| 410 |              "bounded inner" "\<not> bounded outer" "inner \<inter> outer = {}"
 | |
| 411 | "inner \<union> outer = - path_image c" | |
| 412 | "frontier inner = path_image c" | |
| 413 | "frontier outer = path_image c" | |
| 414 | using Jordan_curve [OF assms] by blast | |
| 415 | then have inner: "inside(path_image c) = inner" | |
| 416 | by (metis dual_order.antisym inside_subset interior_eq interior_inside_frontier) | |
| 417 | have outer: "outside(path_image c) = outer" | |
| 418 | using \<open>inner \<union> outer = - path_image c\<close> \<open>inside (path_image c) = inner\<close> | |
| 419 |           outside_inside \<open>inner \<inter> outer = {}\<close> by auto
 | |
| 420 | show ?thesis | |
| 421 | using * by (auto simp: inner outer) | |
| 422 | qed | |
| 423 | ||
| 424 | subsubsection\<open>Triple-curve or "theta-curve" theorem\<close> | |
| 425 | ||
| 426 | text\<open>Proof that there is no fourth component taken from | |
| 427 | Kuratowski's Topology vol 2, para 61, II.\<close> | |
| 428 | ||
| 429 | theorem split_inside_simple_closed_curve: | |
| 430 | fixes c :: "real \<Rightarrow> complex" | |
| 431 | assumes "simple_path c1" and c1: "pathstart c1 = a" "pathfinish c1 = b" | |
| 432 | and "simple_path c2" and c2: "pathstart c2 = a" "pathfinish c2 = b" | |
| 433 | and "simple_path c" and c: "pathstart c = a" "pathfinish c = b" | |
| 434 | and "a \<noteq> b" | |
| 435 |       and c1c2: "path_image c1 \<inter> path_image c2 = {a,b}"
 | |
| 436 |       and c1c: "path_image c1 \<inter> path_image c = {a,b}"
 | |
| 437 |       and c2c: "path_image c2 \<inter> path_image c = {a,b}"
 | |
| 438 |       and ne_12: "path_image c \<inter> inside(path_image c1 \<union> path_image c2) \<noteq> {}"
 | |
| 439 |   obtains "inside(path_image c1 \<union> path_image c) \<inter> inside(path_image c2 \<union> path_image c) = {}"
 | |
| 440 | "inside(path_image c1 \<union> path_image c) \<union> inside(path_image c2 \<union> path_image c) \<union> | |
| 441 |            (path_image c - {a,b}) = inside(path_image c1 \<union> path_image c2)"
 | |
| 442 | proof - | |
| 443 | let ?\<Theta> = "path_image c" let ?\<Theta>1 = "path_image c1" let ?\<Theta>2 = "path_image c2" | |
| 444 | have sp: "simple_path (c1 +++ reversepath c2)" "simple_path (c1 +++ reversepath c)" "simple_path (c2 +++ reversepath c)" | |
| 445 | using assms by (auto simp: simple_path_join_loop_eq arc_simple_path simple_path_reversepath) | |
| 446 | then have op_in12: "open (inside (?\<Theta>1 \<union> ?\<Theta>2))" | |
| 447 | and op_out12: "open (outside (?\<Theta>1 \<union> ?\<Theta>2))" | |
| 448 | and op_in1c: "open (inside (?\<Theta>1 \<union> ?\<Theta>))" | |
| 449 | and op_in2c: "open (inside (?\<Theta>2 \<union> ?\<Theta>))" | |
| 450 | and op_out1c: "open (outside (?\<Theta>1 \<union> ?\<Theta>))" | |
| 451 | and op_out2c: "open (outside (?\<Theta>2 \<union> ?\<Theta>))" | |
| 452 | and co_in1c: "connected (inside (?\<Theta>1 \<union> ?\<Theta>))" | |
| 453 | and co_in2c: "connected (inside (?\<Theta>2 \<union> ?\<Theta>))" | |
| 454 | and co_out12c: "connected (outside (?\<Theta>1 \<union> ?\<Theta>2))" | |
| 455 | and co_out1c: "connected (outside (?\<Theta>1 \<union> ?\<Theta>))" | |
| 456 | and co_out2c: "connected (outside (?\<Theta>2 \<union> ?\<Theta>))" | |
| 457 |      and pa_c: "?\<Theta> - {pathstart c, pathfinish c} \<subseteq> - ?\<Theta>1"
 | |
| 458 |                "?\<Theta> - {pathstart c, pathfinish c} \<subseteq> - ?\<Theta>2"
 | |
| 459 |      and pa_c1: "?\<Theta>1 - {pathstart c1, pathfinish c1} \<subseteq> - ?\<Theta>2"
 | |
| 460 |                 "?\<Theta>1 - {pathstart c1, pathfinish c1} \<subseteq> - ?\<Theta>"
 | |
| 461 |      and pa_c2: "?\<Theta>2 - {pathstart c2, pathfinish c2} \<subseteq> - ?\<Theta>1"
 | |
| 462 |                 "?\<Theta>2 - {pathstart c2, pathfinish c2} \<subseteq> - ?\<Theta>"
 | |
| 463 |      and co_c: "connected(?\<Theta> - {pathstart c,pathfinish c})"
 | |
| 464 |      and co_c1: "connected(?\<Theta>1 - {pathstart c1,pathfinish c1})"
 | |
| 465 |      and co_c2: "connected(?\<Theta>2 - {pathstart c2,pathfinish c2})"
 | |
| 466 | and fr_in: "frontier(inside(?\<Theta>1 \<union> ?\<Theta>2)) = ?\<Theta>1 \<union> ?\<Theta>2" | |
| 467 | "frontier(inside(?\<Theta>2 \<union> ?\<Theta>)) = ?\<Theta>2 \<union> ?\<Theta>" | |
| 468 | "frontier(inside(?\<Theta>1 \<union> ?\<Theta>)) = ?\<Theta>1 \<union> ?\<Theta>" | |
| 469 | and fr_out: "frontier(outside(?\<Theta>1 \<union> ?\<Theta>2)) = ?\<Theta>1 \<union> ?\<Theta>2" | |
| 470 | "frontier(outside(?\<Theta>2 \<union> ?\<Theta>)) = ?\<Theta>2 \<union> ?\<Theta>" | |
| 471 | "frontier(outside(?\<Theta>1 \<union> ?\<Theta>)) = ?\<Theta>1 \<union> ?\<Theta>" | |
| 472 | using Jordan_inside_outside [of "c1 +++ reversepath c2"] | |
| 473 | using Jordan_inside_outside [of "c1 +++ reversepath c"] | |
| 474 | using Jordan_inside_outside [of "c2 +++ reversepath c"] assms | |
| 475 | apply (simp_all add: path_image_join closed_Un closed_simple_path_image open_inside open_outside) | |
| 78480 | 476 | apply (blast | metis connected_simple_path_endless)+ | 
| 64846 | 477 | done | 
| 478 |   have inout_12: "inside (?\<Theta>1 \<union> ?\<Theta>2) \<inter> (?\<Theta> - {pathstart c, pathfinish c}) \<noteq> {}"
 | |
| 479 | by (metis (no_types, lifting) c c1c ne_12 Diff_Int_distrib Diff_empty Int_empty_right Int_left_commute inf_sup_absorb inf_sup_aci(1) inside_no_overlap) | |
| 480 |   have pi_disjoint:  "?\<Theta> \<inter> outside(?\<Theta>1 \<union> ?\<Theta>2) = {}"
 | |
| 481 | proof (rule ccontr) | |
| 482 |     assume "?\<Theta> \<inter> outside (?\<Theta>1 \<union> ?\<Theta>2) \<noteq> {}"
 | |
| 483 | then show False | |
| 484 | using connectedD [OF co_c, of "inside(?\<Theta>1 \<union> ?\<Theta>2)" "outside(?\<Theta>1 \<union> ?\<Theta>2)"] | |
| 485 | using c c1c2 pa_c op_in12 op_out12 inout_12 | |
| 78480 | 486 | apply clarsimp | 
| 487 | by (smt (verit, ccfv_threshold) Diff_Int_distrib Diff_cancel Diff_empty Int_assoc inf_sup_absorb inf_sup_aci(1) outside_no_overlap) | |
| 64846 | 488 | qed | 
| 489 | have out_sub12: "outside(?\<Theta>1 \<union> ?\<Theta>2) \<subseteq> outside(?\<Theta>1 \<union> ?\<Theta>)" "outside(?\<Theta>1 \<union> ?\<Theta>2) \<subseteq> outside(?\<Theta>2 \<union> ?\<Theta>)" | |
| 490 | by (metis Un_commute pi_disjoint outside_Un_outside_Un)+ | |
| 491 |   have pa1_disj_in2: "?\<Theta>1 \<inter> inside (?\<Theta>2 \<union> ?\<Theta>) = {}"
 | |
| 492 | proof (rule ccontr) | |
| 493 |     assume ne: "?\<Theta>1 \<inter> inside (?\<Theta>2 \<union> ?\<Theta>) \<noteq> {}"
 | |
| 494 |     have 1: "inside (?\<Theta> \<union> ?\<Theta>2) \<inter> ?\<Theta> = {}"
 | |
| 495 | by (metis (no_types) Diff_Int_distrib Diff_cancel inf_sup_absorb inf_sup_aci(3) inside_no_overlap) | |
| 496 |     have 2: "outside (?\<Theta> \<union> ?\<Theta>2) \<inter> ?\<Theta> = {}"
 | |
| 497 | by (metis (no_types) Int_empty_right Int_left_commute inf_sup_absorb outside_no_overlap) | |
| 78480 | 498 |     have "path_image c1 \<inter> outside (path_image c2 \<union> path_image c) = {}"
 | 
| 64846 | 499 | using connectedD [OF co_c1, of "inside(?\<Theta>2 \<union> ?\<Theta>)" "outside(?\<Theta>2 \<union> ?\<Theta>)"] | 
| 500 | pa_c1 op_in2c op_out2c ne c1 c2c 1 2 by (auto simp: inf_sup_aci) | |
| 78480 | 501 | then have "outside (?\<Theta>2 \<union> ?\<Theta>) \<subseteq> outside (?\<Theta>1 \<union> ?\<Theta>2)" | 
| 502 | by (metis outside_Un_outside_Un sup_commute) | |
| 64846 | 503 | with out_sub12 | 
| 504 | have "outside(?\<Theta>1 \<union> ?\<Theta>2) = outside(?\<Theta>2 \<union> ?\<Theta>)" by blast | |
| 505 | then have "frontier(outside(?\<Theta>1 \<union> ?\<Theta>2)) = frontier(outside(?\<Theta>2 \<union> ?\<Theta>))" | |
| 506 | by simp | |
| 507 | then show False | |
| 508 | using inout_12 pi_disjoint c c1c c2c fr_out by auto | |
| 509 | qed | |
| 510 |   have pa2_disj_in1: "?\<Theta>2 \<inter> inside(?\<Theta>1 \<union> ?\<Theta>) = {}"
 | |
| 511 | proof (rule ccontr) | |
| 512 |     assume ne: "?\<Theta>2 \<inter> inside (?\<Theta>1 \<union> ?\<Theta>) \<noteq> {}"
 | |
| 513 |     have 1: "inside (?\<Theta> \<union> ?\<Theta>1) \<inter> ?\<Theta> = {}"
 | |
| 514 | by (metis (no_types) Diff_Int_distrib Diff_cancel inf_sup_absorb inf_sup_aci(3) inside_no_overlap) | |
| 515 |     have 2: "outside (?\<Theta> \<union> ?\<Theta>1) \<inter> ?\<Theta> = {}"
 | |
| 516 | by (metis (no_types) Int_empty_right Int_left_commute inf_sup_absorb outside_no_overlap) | |
| 517 | have "outside (?\<Theta>1 \<union> ?\<Theta>) \<subseteq> outside (?\<Theta>1 \<union> ?\<Theta>2)" | |
| 518 | apply (rule outside_Un_outside_Un) | |
| 519 | using connectedD [OF co_c2, of "inside(?\<Theta>1 \<union> ?\<Theta>)" "outside(?\<Theta>1 \<union> ?\<Theta>)"] | |
| 520 | pa_c2 op_in1c op_out1c ne c2 c1c 1 2 by (auto simp: inf_sup_aci) | |
| 521 | with out_sub12 | |
| 522 | have "outside(?\<Theta>1 \<union> ?\<Theta>2) = outside(?\<Theta>1 \<union> ?\<Theta>)" | |
| 523 | by blast | |
| 524 | then have "frontier(outside(?\<Theta>1 \<union> ?\<Theta>2)) = frontier(outside(?\<Theta>1 \<union> ?\<Theta>))" | |
| 525 | by simp | |
| 526 | then show False | |
| 527 | using inout_12 pi_disjoint c c1c c2c fr_out by auto | |
| 528 | qed | |
| 529 | have in_sub_in1: "inside(?\<Theta>1 \<union> ?\<Theta>) \<subseteq> inside(?\<Theta>1 \<union> ?\<Theta>2)" | |
| 530 | using pa2_disj_in1 out_sub12 by (auto simp: inside_outside) | |
| 531 | have in_sub_in2: "inside(?\<Theta>2 \<union> ?\<Theta>) \<subseteq> inside(?\<Theta>1 \<union> ?\<Theta>2)" | |
| 532 | using pa1_disj_in2 out_sub12 by (auto simp: inside_outside) | |
| 533 | have in_sub_out12: "inside(?\<Theta>1 \<union> ?\<Theta>) \<subseteq> outside(?\<Theta>2 \<union> ?\<Theta>)" | |
| 534 | proof | |
| 535 | fix x | |
| 536 | assume x: "x \<in> inside (?\<Theta>1 \<union> ?\<Theta>)" | |
| 537 | then have xnot: "x \<notin> ?\<Theta>" | |
| 538 | by (simp add: inside_def) | |
| 539 | obtain z where zim: "z \<in> ?\<Theta>1" and zout: "z \<in> outside(?\<Theta>2 \<union> ?\<Theta>)" | |
| 78480 | 540 | unfolding outside_inside | 
| 541 | using nonempty_simple_path_endless [OF \<open>simple_path c1\<close>] c1 c1c c1c2 pa1_disj_in2 by auto | |
| 64846 | 542 | obtain e where "e > 0" and e: "ball z e \<subseteq> outside(?\<Theta>2 \<union> ?\<Theta>)" | 
| 543 | using zout op_out2c open_contains_ball_eq by blast | |
| 544 | have "z \<in> frontier (inside (?\<Theta>1 \<union> ?\<Theta>))" | |
| 545 | using zim by (auto simp: fr_in) | |
| 546 | then obtain w where w1: "w \<in> inside (?\<Theta>1 \<union> ?\<Theta>)" and dwz: "dist w z < e" | |
| 547 | using zim \<open>e > 0\<close> by (auto simp: frontier_def closure_approachable) | |
| 548 | then have w2: "w \<in> outside (?\<Theta>2 \<union> ?\<Theta>)" | |
| 549 | by (metis e dist_commute mem_ball subsetCE) | |
| 550 | then have "connected_component (- ?\<Theta>2 \<inter> - ?\<Theta>) z w" | |
| 78480 | 551 | unfolding connected_component_def | 
| 552 | by (metis co_out2c compl_sup inside_Un_outside sup_ge2 zout) | |
| 64846 | 553 | moreover have "connected_component (- ?\<Theta>2 \<inter> - ?\<Theta>) w x" | 
| 554 | unfolding connected_component_def | |
| 555 | using pa2_disj_in1 co_in1c x w1 union_with_outside by fastforce | |
| 556 | ultimately have eq: "connected_component_set (- ?\<Theta>2 \<inter> - ?\<Theta>) x = | |
| 557 | connected_component_set (- ?\<Theta>2 \<inter> - ?\<Theta>) z" | |
| 558 | by (metis (mono_tags, lifting) connected_component_eq mem_Collect_eq) | |
| 559 | show "x \<in> outside (?\<Theta>2 \<union> ?\<Theta>)" | |
| 560 | using zout x pa2_disj_in1 by (auto simp: outside_def eq xnot) | |
| 561 | qed | |
| 562 | have in_sub_out21: "inside(?\<Theta>2 \<union> ?\<Theta>) \<subseteq> outside(?\<Theta>1 \<union> ?\<Theta>)" | |
| 563 | proof | |
| 564 | fix x | |
| 565 | assume x: "x \<in> inside (?\<Theta>2 \<union> ?\<Theta>)" | |
| 566 | then have xnot: "x \<notin> ?\<Theta>" | |
| 567 | by (simp add: inside_def) | |
| 568 | obtain z where zim: "z \<in> ?\<Theta>2" and zout: "z \<in> outside(?\<Theta>1 \<union> ?\<Theta>)" | |
| 78480 | 569 | unfolding outside_inside | 
| 570 | using nonempty_simple_path_endless [OF \<open>simple_path c2\<close>] c1c2 c2 c2c pa2_disj_in1 by auto | |
| 64846 | 571 | obtain e where "e > 0" and e: "ball z e \<subseteq> outside(?\<Theta>1 \<union> ?\<Theta>)" | 
| 572 | using zout op_out1c open_contains_ball_eq by blast | |
| 573 | have "z \<in> frontier (inside (?\<Theta>2 \<union> ?\<Theta>))" | |
| 574 | using zim by (auto simp: fr_in) | |
| 575 | then obtain w where w2: "w \<in> inside (?\<Theta>2 \<union> ?\<Theta>)" and dwz: "dist w z < e" | |
| 576 | using zim \<open>e > 0\<close> by (auto simp: frontier_def closure_approachable) | |
| 577 | then have w1: "w \<in> outside (?\<Theta>1 \<union> ?\<Theta>)" | |
| 578 | by (metis e dist_commute mem_ball subsetCE) | |
| 579 | then have "connected_component (- ?\<Theta>1 \<inter> - ?\<Theta>) z w" | |
| 78480 | 580 | unfolding connected_component_def | 
| 581 | by (metis Compl_Un co_out1c inside_Un_outside sup_ge2 zout) | |
| 64846 | 582 | moreover have "connected_component (- ?\<Theta>1 \<inter> - ?\<Theta>) w x" | 
| 583 | unfolding connected_component_def | |
| 584 | using pa1_disj_in2 co_in2c x w2 union_with_outside by fastforce | |
| 585 | ultimately have eq: "connected_component_set (- ?\<Theta>1 \<inter> - ?\<Theta>) x = | |
| 586 | connected_component_set (- ?\<Theta>1 \<inter> - ?\<Theta>) z" | |
| 587 | by (metis (no_types, lifting) connected_component_eq mem_Collect_eq) | |
| 588 | show "x \<in> outside (?\<Theta>1 \<union> ?\<Theta>)" | |
| 589 | using zout x pa1_disj_in2 by (auto simp: outside_def eq xnot) | |
| 590 | qed | |
| 591 | show ?thesis | |
| 592 | proof | |
| 593 |     show "inside (?\<Theta>1 \<union> ?\<Theta>) \<inter> inside (?\<Theta>2 \<union> ?\<Theta>) = {}"
 | |
| 594 | by (metis Int_Un_distrib in_sub_out12 bot_eq_sup_iff disjoint_eq_subset_Compl outside_inside) | |
| 595 | have *: "outside (?\<Theta>1 \<union> ?\<Theta>) \<inter> outside (?\<Theta>2 \<union> ?\<Theta>) \<subseteq> outside (?\<Theta>1 \<union> ?\<Theta>2)" | |
| 596 | proof (rule components_maximal) | |
| 597 | show out_in: "outside (?\<Theta>1 \<union> ?\<Theta>2) \<in> components (- (?\<Theta>1 \<union> ?\<Theta>2))" | |
| 78480 | 598 | unfolding outside_in_components co_out12c | 
| 599 | using co_out12c fr_out(1) by force | |
| 64846 | 600 | have conn_U: "connected (- (closure (inside (?\<Theta>1 \<union> ?\<Theta>)) \<union> closure (inside (?\<Theta>2 \<union> ?\<Theta>))))" | 
| 601 | proof (rule Janiszewski_connected, simp_all) | |
| 602 | show "bounded (inside (?\<Theta>1 \<union> ?\<Theta>))" | |
| 603 | by (simp add: \<open>simple_path c1\<close> \<open>simple_path c\<close> bounded_inside bounded_simple_path_image) | |
| 604 | have if1: "- (inside (?\<Theta>1 \<union> ?\<Theta>) \<union> frontier (inside (?\<Theta>1 \<union> ?\<Theta>))) = - ?\<Theta>1 \<inter> - ?\<Theta> \<inter> - inside (?\<Theta>1 \<union> ?\<Theta>)" | |
| 605 | by (metis (no_types, lifting) Int_commute Jordan_inside_outside c c1 compl_sup path_image_join path_image_reversepath pathfinish_join pathfinish_reversepath pathstart_join pathstart_reversepath sp(2) closure_Un_frontier fr_out(3)) | |
| 606 | then show "connected (- closure (inside (?\<Theta>1 \<union> ?\<Theta>)))" | |
| 607 | by (metis Compl_Un outside_inside co_out1c closure_Un_frontier) | |
| 608 | have if2: "- (inside (?\<Theta>2 \<union> ?\<Theta>) \<union> frontier (inside (?\<Theta>2 \<union> ?\<Theta>))) = - ?\<Theta>2 \<inter> - ?\<Theta> \<inter> - inside (?\<Theta>2 \<union> ?\<Theta>)" | |
| 71633 | 609 | by (metis (no_types, lifting) Int_commute Jordan_inside_outside c c2 compl_sup path_image_join path_image_reversepath pathfinish_join pathfinish_reversepath pathstart_join pathstart_reversepath sp(3) closure_Un_frontier fr_out(2)) | 
| 64846 | 610 | then show "connected (- closure (inside (?\<Theta>2 \<union> ?\<Theta>)))" | 
| 611 | by (metis Compl_Un outside_inside co_out2c closure_Un_frontier) | |
| 612 | have "connected(?\<Theta>)" | |
| 613 | by (metis \<open>simple_path c\<close> connected_simple_path_image) | |
| 614 | moreover | |
| 615 | have "closure (inside (?\<Theta>1 \<union> ?\<Theta>)) \<inter> closure (inside (?\<Theta>2 \<union> ?\<Theta>)) = ?\<Theta>" | |
| 616 | (is "?lhs = ?rhs") | |
| 617 | proof | |
| 618 | show "?lhs \<subseteq> ?rhs" | |
| 619 | proof clarify | |
| 620 | fix x | |
| 621 | assume x: "x \<in> closure (inside (?\<Theta>1 \<union> ?\<Theta>))" "x \<in> closure (inside (?\<Theta>2 \<union> ?\<Theta>))" | |
| 622 | then have "x \<notin> inside (?\<Theta>1 \<union> ?\<Theta>)" | |
| 623 | by (meson closure_iff_nhds_not_empty in_sub_out12 inside_Int_outside op_in1c) | |
| 624 | with fr_in x show "x \<in> ?\<Theta>" | |
| 625 | by (metis c1c c1c2 closure_Un_frontier pa1_disj_in2 Int_iff Un_iff insert_disjoint(2) insert_subset subsetI subset_antisym) | |
| 626 | qed | |
| 627 | show "?rhs \<subseteq> ?lhs" | |
| 628 | using if1 if2 closure_Un_frontier by fastforce | |
| 629 | qed | |
| 630 | ultimately | |
| 631 | show "connected (closure (inside (?\<Theta>1 \<union> ?\<Theta>)) \<inter> closure (inside (?\<Theta>2 \<union> ?\<Theta>)))" | |
| 632 | by auto | |
| 633 | qed | |
| 634 | show "connected (outside (?\<Theta>1 \<union> ?\<Theta>) \<inter> outside (?\<Theta>2 \<union> ?\<Theta>))" | |
| 635 | using fr_in conn_U by (simp add: closure_Un_frontier outside_inside Un_commute) | |
| 636 | show "outside (?\<Theta>1 \<union> ?\<Theta>) \<inter> outside (?\<Theta>2 \<union> ?\<Theta>) \<subseteq> - (?\<Theta>1 \<union> ?\<Theta>2)" | |
| 637 | by clarify (metis Diff_Compl Diff_iff Un_iff inf_sup_absorb outside_inside) | |
| 638 | show "outside (?\<Theta>1 \<union> ?\<Theta>2) \<inter> | |
| 639 |             (outside (?\<Theta>1 \<union> ?\<Theta>) \<inter> outside (?\<Theta>2 \<union> ?\<Theta>)) \<noteq> {}"
 | |
| 640 | by (metis Int_assoc out_in inf.orderE out_sub12(1) out_sub12(2) outside_in_components) | |
| 641 | qed | |
| 642 |     show "inside (?\<Theta>1 \<union> ?\<Theta>) \<union> inside (?\<Theta>2 \<union> ?\<Theta>) \<union> (?\<Theta> - {a, b}) = inside (?\<Theta>1 \<union> ?\<Theta>2)"
 | |
| 643 | (is "?lhs = ?rhs") | |
| 644 | proof | |
| 78480 | 645 |       have " path_image c - {a, b} \<subseteq> inside (path_image c1 \<union> path_image c2)"
 | 
| 64846 | 646 | using c1c c2c inside_outside pi_disjoint by fastforce | 
| 78480 | 647 | then show "?lhs \<subseteq> ?rhs" | 
| 648 | by (simp add: in_sub_in1 in_sub_in2) | |
| 64846 | 649 | have "inside (?\<Theta>1 \<union> ?\<Theta>2) \<subseteq> inside (?\<Theta>1 \<union> ?\<Theta>) \<union> inside (?\<Theta>2 \<union> ?\<Theta>) \<union> (?\<Theta>)" | 
| 650 | using Compl_anti_mono [OF *] by (force simp: inside_outside) | |
| 651 |       moreover have "inside (?\<Theta>1 \<union> ?\<Theta>2) \<subseteq> -{a,b}"
 | |
| 652 | using c1 union_with_outside by fastforce | |
| 653 | ultimately show "?rhs \<subseteq> ?lhs" by auto | |
| 654 | qed | |
| 655 | qed | |
| 656 | qed | |
| 657 | ||
| 658 | end |