# HG changeset patch # User immler # Date 1546869465 -3600 # Node ID 19d8a59481dbe4a60c6de14aa48a696986a26332 # Parent 3f7d8e05e0f27e6d8f5f40ee71b8e48aa9437c47 split off Homotopy.thy diff -r 3f7d8e05e0f2 -r 19d8a59481db src/HOL/Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Analysis/Brouwer_Fixpoint.thy Mon Jan 07 14:06:54 2019 +0100 +++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy Mon Jan 07 14:57:45 2019 +0100 @@ -15,7 +15,10 @@ section \Brouwer's Fixed Point Theorem\ theory Brouwer_Fixpoint -imports Path_Connected Homeomorphism + imports + Path_Connected + Homeomorphism + Continuous_Extension begin (* FIXME mv topology euclidean space *) diff -r 3f7d8e05e0f2 -r 19d8a59481db src/HOL/Analysis/Homeomorphism.thy --- a/src/HOL/Analysis/Homeomorphism.thy Mon Jan 07 14:06:54 2019 +0100 +++ b/src/HOL/Analysis/Homeomorphism.thy Mon Jan 07 14:57:45 2019 +0100 @@ -5,7 +5,7 @@ section%important \Homeomorphism Theorems\ theory Homeomorphism -imports Path_Connected +imports Homotopy begin lemma%unimportant homeomorphic_spheres': diff -r 3f7d8e05e0f2 -r 19d8a59481db src/HOL/Analysis/Homotopy.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Homotopy.thy Mon Jan 07 14:57:45 2019 +0100 @@ -0,0 +1,5159 @@ +(* Title: HOL/Analysis/Path_Connected.thy + Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light +*) + +section \Homotopy of Maps\ + +theory Homotopy + imports Path_Connected Continuum_Not_Denumerable +begin + +definition%important homotopic_with :: + "[('a::topological_space \ 'b::topological_space) \ bool, 'a set, 'b set, 'a \ 'b, 'a \ 'b] \ bool" +where + "homotopic_with P X Y p q \ + (\h:: real \ 'a \ 'b. + continuous_on ({0..1} \ X) h \ + h ` ({0..1} \ X) \ Y \ + (\x. h(0, x) = p x) \ + (\x. h(1, x) = q x) \ + (\t \ {0..1}. P(\x. h(t, x))))" + +text\\p\, \q\ are functions \X \ Y\, and the property \P\ restricts all intermediate maps. +We often just want to require that \P\ fixes some subset, but to include the case of a loop homotopy, +it is convenient to have a general property \P\.\ + +text \We often want to just localize the ending function equality or whatever.\ +text%important \%whitespace\ +proposition homotopic_with: + fixes X :: "'a::topological_space set" and Y :: "'b::topological_space set" + assumes "\h k. (\x. x \ X \ h x = k x) \ (P h \ P k)" + shows "homotopic_with P X Y p q \ + (\h :: real \ 'a \ 'b. + continuous_on ({0..1} \ X) h \ + h ` ({0..1} \ X) \ Y \ + (\x \ X. h(0,x) = p x) \ + (\x \ X. h(1,x) = q x) \ + (\t \ {0..1}. P(\x. h(t, x))))" + unfolding homotopic_with_def + apply (rule iffI, blast, clarify) + apply (rule_tac x="\(u,v). if v \ X then h(u,v) else if u = 0 then p v else q v" in exI) + apply auto + apply (force elim: continuous_on_eq) + apply (drule_tac x=t in bspec, force) + apply (subst assms; simp) + done + +proposition homotopic_with_eq: + assumes h: "homotopic_with P X Y f g" + and f': "\x. x \ X \ f' x = f x" + and g': "\x. x \ X \ g' x = g x" + and P: "(\h k. (\x. x \ X \ h x = k x) \ (P h \ P k))" + shows "homotopic_with P X Y f' g'" + using h unfolding homotopic_with_def + apply safe + apply (rule_tac x="\(u,v). if v \ X then h(u,v) else if u = 0 then f' v else g' v" in exI) + apply (simp add: f' g', safe) + apply (fastforce intro: continuous_on_eq, fastforce) + apply (subst P; fastforce) + done + +proposition homotopic_with_equal: + assumes contf: "continuous_on X f" and fXY: "f ` X \ Y" + and gf: "\x. x \ X \ g x = f x" + and P: "P f" "P g" + shows "homotopic_with P X Y f g" + unfolding homotopic_with_def + apply (rule_tac x="\(u,v). if u = 1 then g v else f v" in exI) + using assms + apply (intro conjI) + apply (rule continuous_on_eq [where f = "f \ snd"]) + apply (rule continuous_intros | force)+ + apply clarify + apply (case_tac "t=1"; force) + done + + +lemma image_Pair_const: "(\x. (x, c)) ` A = A \ {c}" + by auto + +lemma homotopic_constant_maps: + "homotopic_with (\x. True) s t (\x. a) (\x. b) \ s = {} \ path_component t a b" +proof (cases "s = {} \ t = {}") + case True with continuous_on_const show ?thesis + by (auto simp: homotopic_with path_component_def) +next + case False + then obtain c where "c \ s" by blast + show ?thesis + proof + assume "homotopic_with (\x. True) s t (\x. a) (\x. b)" + then obtain h :: "real \ 'a \ 'b" + where conth: "continuous_on ({0..1} \ s) h" + and h: "h ` ({0..1} \ s) \ t" "(\x\s. h (0, x) = a)" "(\x\s. h (1, x) = b)" + by (auto simp: homotopic_with) + have "continuous_on {0..1} (h \ (\t. (t, c)))" + apply (rule continuous_intros conth | simp add: image_Pair_const)+ + apply (blast intro: \c \ s\ continuous_on_subset [OF conth]) + done + with \c \ s\ h show "s = {} \ path_component t a b" + apply (simp_all add: homotopic_with path_component_def, auto) + apply (drule_tac x="h \ (\t. (t, c))" in spec) + apply (auto simp: pathstart_def pathfinish_def path_image_def path_def) + done + next + assume "s = {} \ path_component t a b" + with False show "homotopic_with (\x. True) s t (\x. a) (\x. b)" + apply (clarsimp simp: homotopic_with path_component_def pathstart_def pathfinish_def path_image_def path_def) + apply (rule_tac x="g \ fst" in exI) + apply (rule conjI continuous_intros | force)+ + done + qed +qed + + +subsection%unimportant\Trivial properties\ + +lemma homotopic_with_imp_property: "homotopic_with P X Y f g \ P f \ P g" + unfolding homotopic_with_def Ball_def + apply clarify + apply (frule_tac x=0 in spec) + apply (drule_tac x=1 in spec, auto) + done + +lemma continuous_on_o_Pair: "\continuous_on (T \ X) h; t \ T\ \ continuous_on X (h \ Pair t)" + by (fast intro: continuous_intros elim!: continuous_on_subset) + +lemma homotopic_with_imp_continuous: + assumes "homotopic_with P X Y f g" + shows "continuous_on X f \ continuous_on X g" +proof - + obtain h :: "real \ 'a \ 'b" + where conth: "continuous_on ({0..1} \ X) h" + and h: "\x. h (0, x) = f x" "\x. h (1, x) = g x" + using assms by (auto simp: homotopic_with_def) + have *: "t \ {0..1} \ continuous_on X (h \ (\x. (t,x)))" for t + by (rule continuous_intros continuous_on_subset [OF conth] | force)+ + show ?thesis + using h *[of 0] *[of 1] by auto +qed + +proposition homotopic_with_imp_subset1: + "homotopic_with P X Y f g \ f ` X \ Y" + by (simp add: homotopic_with_def image_subset_iff) (metis atLeastAtMost_iff order_refl zero_le_one) + +proposition homotopic_with_imp_subset2: + "homotopic_with P X Y f g \ g ` X \ Y" + by (simp add: homotopic_with_def image_subset_iff) (metis atLeastAtMost_iff order_refl zero_le_one) + +proposition homotopic_with_mono: + assumes hom: "homotopic_with P X Y f g" + and Q: "\h. \continuous_on X h; image h X \ Y \ P h\ \ Q h" + shows "homotopic_with Q X Y f g" + using hom + apply (simp add: homotopic_with_def) + apply (erule ex_forward) + apply (force simp: intro!: Q dest: continuous_on_o_Pair) + done + +proposition homotopic_with_subset_left: + "\homotopic_with P X Y f g; Z \ X\ \ homotopic_with P Z Y f g" + apply (simp add: homotopic_with_def) + apply (fast elim!: continuous_on_subset ex_forward) + done + +proposition homotopic_with_subset_right: + "\homotopic_with P X Y f g; Y \ Z\ \ homotopic_with P X Z f g" + apply (simp add: homotopic_with_def) + apply (fast elim!: continuous_on_subset ex_forward) + done + +proposition homotopic_with_compose_continuous_right: + "\homotopic_with (\f. p (f \ h)) X Y f g; continuous_on W h; h ` W \ X\ + \ homotopic_with p W Y (f \ h) (g \ h)" + apply (clarsimp simp add: homotopic_with_def) + apply (rename_tac k) + apply (rule_tac x="k \ (\y. (fst y, h (snd y)))" in exI) + apply (rule conjI continuous_intros continuous_on_compose [where f=snd and g=h, unfolded o_def] | simp)+ + apply (erule continuous_on_subset) + apply (fastforce simp: o_def)+ + done + +proposition homotopic_compose_continuous_right: + "\homotopic_with (\f. True) X Y f g; continuous_on W h; h ` W \ X\ + \ homotopic_with (\f. True) W Y (f \ h) (g \ h)" + using homotopic_with_compose_continuous_right by fastforce + +proposition homotopic_with_compose_continuous_left: + "\homotopic_with (\f. p (h \ f)) X Y f g; continuous_on Y h; h ` Y \ Z\ + \ homotopic_with p X Z (h \ f) (h \ g)" + apply (clarsimp simp add: homotopic_with_def) + apply (rename_tac k) + apply (rule_tac x="h \ k" in exI) + apply (rule conjI continuous_intros continuous_on_compose [where f=snd and g=h, unfolded o_def] | simp)+ + apply (erule continuous_on_subset) + apply (fastforce simp: o_def)+ + done + +proposition homotopic_compose_continuous_left: + "\homotopic_with (\_. True) X Y f g; + continuous_on Y h; h ` Y \ Z\ + \ homotopic_with (\f. True) X Z (h \ f) (h \ g)" + using homotopic_with_compose_continuous_left by fastforce + +proposition homotopic_with_Pair: + assumes hom: "homotopic_with p s t f g" "homotopic_with p' s' t' f' g'" + and q: "\f g. \p f; p' g\ \ q(\(x,y). (f x, g y))" + shows "homotopic_with q (s \ s') (t \ t') + (\(x,y). (f x, f' y)) (\(x,y). (g x, g' y))" + using hom + apply (clarsimp simp add: homotopic_with_def) + apply (rename_tac k k') + apply (rule_tac x="\z. ((k \ (\x. (fst x, fst (snd x)))) z, (k' \ (\x. (fst x, snd (snd x)))) z)" in exI) + apply (rule conjI continuous_intros | erule continuous_on_subset | clarsimp)+ + apply (auto intro!: q [unfolded case_prod_unfold]) + done + +lemma homotopic_on_empty [simp]: "homotopic_with (\x. True) {} t f g" + by (metis continuous_on_def empty_iff homotopic_with_equal image_subset_iff) + + +text\Homotopy with P is an equivalence relation (on continuous functions mapping X into Y that satisfy P, + though this only affects reflexivity.\ + + +proposition homotopic_with_refl: + "homotopic_with P X Y f f \ continuous_on X f \ image f X \ Y \ P f" + apply (rule iffI) + using homotopic_with_imp_continuous homotopic_with_imp_property homotopic_with_imp_subset2 apply blast + apply (simp add: homotopic_with_def) + apply (rule_tac x="f \ snd" in exI) + apply (rule conjI continuous_intros | force)+ + done + +lemma homotopic_with_symD: + fixes X :: "'a::real_normed_vector set" + assumes "homotopic_with P X Y f g" + shows "homotopic_with P X Y g f" + using assms + apply (clarsimp simp add: homotopic_with_def) + apply (rename_tac h) + apply (rule_tac x="h \ (\y. (1 - fst y, snd y))" in exI) + apply (rule conjI continuous_intros | erule continuous_on_subset | force simp: image_subset_iff)+ + done + +proposition homotopic_with_sym: + fixes X :: "'a::real_normed_vector set" + shows "homotopic_with P X Y f g \ homotopic_with P X Y g f" + using homotopic_with_symD by blast + +lemma split_01: "{0..1::real} = {0..1/2} \ {1/2..1}" + by force + +lemma split_01_prod: "{0..1::real} \ X = ({0..1/2} \ X) \ ({1/2..1} \ X)" + by force + +proposition homotopic_with_trans: + fixes X :: "'a::real_normed_vector set" + assumes "homotopic_with P X Y f g" and "homotopic_with P X Y g h" + shows "homotopic_with P X Y f h" +proof - + have clo1: "closedin (subtopology euclidean ({0..1/2} \ X \ {1/2..1} \ X)) ({0..1/2::real} \ X)" + apply (simp add: closedin_closed split_01_prod [symmetric]) + apply (rule_tac x="{0..1/2} \ UNIV" in exI) + apply (force simp: closed_Times) + done + have clo2: "closedin (subtopology euclidean ({0..1/2} \ X \ {1/2..1} \ X)) ({1/2..1::real} \ X)" + apply (simp add: closedin_closed split_01_prod [symmetric]) + apply (rule_tac x="{1/2..1} \ UNIV" in exI) + apply (force simp: closed_Times) + done + { fix k1 k2:: "real \ 'a \ 'b" + assume cont: "continuous_on ({0..1} \ X) k1" "continuous_on ({0..1} \ X) k2" + and Y: "k1 ` ({0..1} \ X) \ Y" "k2 ` ({0..1} \ X) \ Y" + and geq: "\x. k1 (1, x) = g x" "\x. k2 (0, x) = g x" + and k12: "\x. k1 (0, x) = f x" "\x. k2 (1, x) = h x" + and P: "\t\{0..1}. P (\x. k1 (t, x))" "\t\{0..1}. P (\x. k2 (t, x))" + define k where "k y = + (if fst y \ 1 / 2 + then (k1 \ (\x. (2 *\<^sub>R fst x, snd x))) y + else (k2 \ (\x. (2 *\<^sub>R fst x -1, snd x))) y)" for y + have keq: "k1 (2 * u, v) = k2 (2 * u - 1, v)" if "u = 1/2" for u v + by (simp add: geq that) + have "continuous_on ({0..1} \ X) k" + using cont + apply (simp add: split_01_prod k_def) + apply (rule clo1 clo2 continuous_on_cases_local continuous_intros | erule continuous_on_subset | simp add: linear image_subset_iff)+ + apply (force simp: keq) + done + moreover have "k ` ({0..1} \ X) \ Y" + using Y by (force simp: k_def) + moreover have "\x. k (0, x) = f x" + by (simp add: k_def k12) + moreover have "(\x. k (1, x) = h x)" + by (simp add: k_def k12) + moreover have "\t\{0..1}. P (\x. k (t, x))" + using P + apply (clarsimp simp add: k_def) + apply (case_tac "t \ 1/2", auto) + done + ultimately have *: "\k :: real \ 'a \ 'b. + continuous_on ({0..1} \ X) k \ k ` ({0..1} \ X) \ Y \ + (\x. k (0, x) = f x) \ (\x. k (1, x) = h x) \ (\t\{0..1}. P (\x. k (t, x)))" + by blast + } note * = this + show ?thesis + using assms by (auto intro: * simp add: homotopic_with_def) +qed + +proposition homotopic_compose: + fixes s :: "'a::real_normed_vector set" + shows "\homotopic_with (\x. True) s t f f'; homotopic_with (\x. True) t u g g'\ + \ homotopic_with (\x. True) s u (g \ f) (g' \ f')" + apply (rule homotopic_with_trans [where g = "g \ f'"]) + apply (metis homotopic_compose_continuous_left homotopic_with_imp_continuous homotopic_with_imp_subset1) + by (metis homotopic_compose_continuous_right homotopic_with_imp_continuous homotopic_with_imp_subset2) + + +text\Homotopic triviality implicitly incorporates path-connectedness.\ +lemma homotopic_triviality: + fixes S :: "'a::real_normed_vector set" + shows "(\f g. continuous_on S f \ f ` S \ T \ + continuous_on S g \ g ` S \ T + \ homotopic_with (\x. True) S T f g) \ + (S = {} \ path_connected T) \ + (\f. continuous_on S f \ f ` S \ T \ (\c. homotopic_with (\x. True) S T f (\x. c)))" + (is "?lhs = ?rhs") +proof (cases "S = {} \ T = {}") + case True then show ?thesis by auto +next + case False show ?thesis + proof + assume LHS [rule_format]: ?lhs + have pab: "path_component T a b" if "a \ T" "b \ T" for a b + proof - + have "homotopic_with (\x. True) S T (\x. a) (\x. b)" + by (simp add: LHS continuous_on_const image_subset_iff that) + then show ?thesis + using False homotopic_constant_maps by blast + qed + moreover + have "\c. homotopic_with (\x. True) S T f (\x. c)" if "continuous_on S f" "f ` S \ T" for f + by (metis (full_types) False LHS equals0I homotopic_constant_maps homotopic_with_imp_continuous homotopic_with_imp_subset2 pab that) + ultimately show ?rhs + by (simp add: path_connected_component) + next + assume RHS: ?rhs + with False have T: "path_connected T" + by blast + show ?lhs + proof clarify + fix f g + assume "continuous_on S f" "f ` S \ T" "continuous_on S g" "g ` S \ T" + obtain c d where c: "homotopic_with (\x. True) S T f (\x. c)" and d: "homotopic_with (\x. True) S T g (\x. d)" + using False \continuous_on S f\ \f ` S \ T\ RHS \continuous_on S g\ \g ` S \ T\ by blast + then have "c \ T" "d \ T" + using False homotopic_with_imp_subset2 by fastforce+ + with T have "path_component T c d" + using path_connected_component by blast + then have "homotopic_with (\x. True) S T (\x. c) (\x. d)" + by (simp add: homotopic_constant_maps) + with c d show "homotopic_with (\x. True) S T f g" + by (meson homotopic_with_symD homotopic_with_trans) + qed + qed +qed + + +subsection\Homotopy of paths, maintaining the same endpoints\ + + +definition%important homotopic_paths :: "['a set, real \ 'a, real \ 'a::topological_space] \ bool" + where + "homotopic_paths s p q \ + homotopic_with (\r. pathstart r = pathstart p \ pathfinish r = pathfinish p) {0..1} s p q" + +lemma homotopic_paths: + "homotopic_paths s p q \ + (\h. continuous_on ({0..1} \ {0..1}) h \ + h ` ({0..1} \ {0..1}) \ s \ + (\x \ {0..1}. h(0,x) = p x) \ + (\x \ {0..1}. h(1,x) = q x) \ + (\t \ {0..1::real}. pathstart(h \ Pair t) = pathstart p \ + pathfinish(h \ Pair t) = pathfinish p))" + by (auto simp: homotopic_paths_def homotopic_with pathstart_def pathfinish_def) + +proposition homotopic_paths_imp_pathstart: + "homotopic_paths s p q \ pathstart p = pathstart q" + by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) + +proposition homotopic_paths_imp_pathfinish: + "homotopic_paths s p q \ pathfinish p = pathfinish q" + by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) + +lemma homotopic_paths_imp_path: + "homotopic_paths s p q \ path p \ path q" + using homotopic_paths_def homotopic_with_imp_continuous path_def by blast + +lemma homotopic_paths_imp_subset: + "homotopic_paths s p q \ path_image p \ s \ path_image q \ s" + by (simp add: homotopic_paths_def homotopic_with_imp_subset1 homotopic_with_imp_subset2 path_image_def) + +proposition homotopic_paths_refl [simp]: "homotopic_paths s p p \ path p \ path_image p \ s" +by (simp add: homotopic_paths_def homotopic_with_refl path_def path_image_def) + +proposition homotopic_paths_sym: "homotopic_paths s p q \ homotopic_paths s q p" + by (metis (mono_tags) homotopic_paths_def homotopic_paths_imp_pathfinish homotopic_paths_imp_pathstart homotopic_with_symD) + +proposition homotopic_paths_sym_eq: "homotopic_paths s p q \ homotopic_paths s q p" + by (metis homotopic_paths_sym) + +proposition homotopic_paths_trans [trans]: + "\homotopic_paths s p q; homotopic_paths s q r\ \ homotopic_paths s p r" + apply (simp add: homotopic_paths_def) + apply (rule homotopic_with_trans, assumption) + by (metis (mono_tags, lifting) homotopic_with_imp_property homotopic_with_mono) + +proposition homotopic_paths_eq: + "\path p; path_image p \ s; \t. t \ {0..1} \ p t = q t\ \ homotopic_paths s p q" + apply (simp add: homotopic_paths_def) + apply (rule homotopic_with_eq) + apply (auto simp: path_def homotopic_with_refl pathstart_def pathfinish_def path_image_def elim: continuous_on_eq) + done + +proposition homotopic_paths_reparametrize: + assumes "path p" + and pips: "path_image p \ s" + and contf: "continuous_on {0..1} f" + and f01:"f ` {0..1} \ {0..1}" + and [simp]: "f(0) = 0" "f(1) = 1" + and q: "\t. t \ {0..1} \ q(t) = p(f t)" + shows "homotopic_paths s p q" +proof - + have contp: "continuous_on {0..1} p" + by (metis \path p\ path_def) + then have "continuous_on {0..1} (p \ f)" + using contf continuous_on_compose continuous_on_subset f01 by blast + then have "path q" + by (simp add: path_def) (metis q continuous_on_cong) + have piqs: "path_image q \ s" + by (metis (no_types, hide_lams) pips f01 image_subset_iff path_image_def q) + have fb0: "\a b. \0 \ a; a \ 1; 0 \ b; b \ 1\ \ 0 \ (1 - a) * f b + a * b" + using f01 by force + have fb1: "\0 \ a; a \ 1; 0 \ b; b \ 1\ \ (1 - a) * f b + a * b \ 1" for a b + using f01 [THEN subsetD, of "f b"] by (simp add: convex_bound_le) + have "homotopic_paths s q p" + proof (rule homotopic_paths_trans) + show "homotopic_paths s q (p \ f)" + using q by (force intro: homotopic_paths_eq [OF \path q\ piqs]) + next + show "homotopic_paths s (p \ f) p" + apply (simp add: homotopic_paths_def homotopic_with_def) + apply (rule_tac x="p \ (\y. (1 - (fst y)) *\<^sub>R ((f \ snd) y) + (fst y) *\<^sub>R snd y)" in exI) + apply (rule conjI contf continuous_intros continuous_on_subset [OF contp] | simp)+ + using pips [unfolded path_image_def] + apply (auto simp: fb0 fb1 pathstart_def pathfinish_def) + done + qed + then show ?thesis + by (simp add: homotopic_paths_sym) +qed + +lemma homotopic_paths_subset: "\homotopic_paths s p q; s \ t\ \ homotopic_paths t p q" + using homotopic_paths_def homotopic_with_subset_right by blast + + +text\ A slightly ad-hoc but useful lemma in constructing homotopies.\ +lemma homotopic_join_lemma: + fixes q :: "[real,real] \ 'a::topological_space" + assumes p: "continuous_on ({0..1} \ {0..1}) (\y. p (fst y) (snd y))" + and q: "continuous_on ({0..1} \ {0..1}) (\y. q (fst y) (snd y))" + and pf: "\t. t \ {0..1} \ pathfinish(p t) = pathstart(q t)" + shows "continuous_on ({0..1} \ {0..1}) (\y. (p(fst y) +++ q(fst y)) (snd y))" +proof - + have 1: "(\y. p (fst y) (2 * snd y)) = (\y. p (fst y) (snd y)) \ (\y. (fst y, 2 * snd y))" + by (rule ext) (simp) + have 2: "(\y. q (fst y) (2 * snd y - 1)) = (\y. q (fst y) (snd y)) \ (\y. (fst y, 2 * snd y - 1))" + by (rule ext) (simp) + show ?thesis + apply (simp add: joinpaths_def) + apply (rule continuous_on_cases_le) + apply (simp_all only: 1 2) + apply (rule continuous_intros continuous_on_subset [OF p] continuous_on_subset [OF q] | force)+ + using pf + apply (auto simp: mult.commute pathstart_def pathfinish_def) + done +qed + +text\ Congruence properties of homotopy w.r.t. path-combining operations.\ + +lemma homotopic_paths_reversepath_D: + assumes "homotopic_paths s p q" + shows "homotopic_paths s (reversepath p) (reversepath q)" + using assms + apply (simp add: homotopic_paths_def homotopic_with_def, clarify) + apply (rule_tac x="h \ (\x. (fst x, 1 - snd x))" in exI) + apply (rule conjI continuous_intros)+ + apply (auto simp: reversepath_def pathstart_def pathfinish_def elim!: continuous_on_subset) + done + +proposition homotopic_paths_reversepath: + "homotopic_paths s (reversepath p) (reversepath q) \ homotopic_paths s p q" + using homotopic_paths_reversepath_D by force + + +proposition homotopic_paths_join: + "\homotopic_paths s p p'; homotopic_paths s q q'; pathfinish p = pathstart q\ \ homotopic_paths s (p +++ q) (p' +++ q')" + apply (simp add: homotopic_paths_def homotopic_with_def, clarify) + apply (rename_tac k1 k2) + apply (rule_tac x="(\y. ((k1 \ Pair (fst y)) +++ (k2 \ Pair (fst y))) (snd y))" in exI) + apply (rule conjI continuous_intros homotopic_join_lemma)+ + apply (auto simp: joinpaths_def pathstart_def pathfinish_def path_image_def) + done + +proposition homotopic_paths_continuous_image: + "\homotopic_paths s f g; continuous_on s h; h ` s \ t\ \ homotopic_paths t (h \ f) (h \ g)" + unfolding homotopic_paths_def + apply (rule homotopic_with_compose_continuous_left [of _ _ _ s]) + apply (auto simp: pathstart_def pathfinish_def elim!: homotopic_with_mono) + done + + +subsection\Group properties for homotopy of paths\ + +text%important\So taking equivalence classes under homotopy would give the fundamental group\ + +proposition homotopic_paths_rid: + "\path p; path_image p \ s\ \ homotopic_paths s (p +++ linepath (pathfinish p) (pathfinish p)) p" + apply (subst homotopic_paths_sym) + apply (rule homotopic_paths_reparametrize [where f = "\t. if t \ 1 / 2 then 2 *\<^sub>R t else 1"]) + apply (simp_all del: le_divide_eq_numeral1) + apply (subst split_01) + apply (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+ + done + +proposition homotopic_paths_lid: + "\path p; path_image p \ s\ \ homotopic_paths s (linepath (pathstart p) (pathstart p) +++ p) p" + using homotopic_paths_rid [of "reversepath p" s] + by (metis homotopic_paths_reversepath path_image_reversepath path_reversepath pathfinish_linepath + pathfinish_reversepath reversepath_joinpaths reversepath_linepath) + +proposition homotopic_paths_assoc: + "\path p; path_image p \ s; path q; path_image q \ s; path r; path_image r \ s; pathfinish p = pathstart q; + pathfinish q = pathstart r\ + \ homotopic_paths s (p +++ (q +++ r)) ((p +++ q) +++ r)" + apply (subst homotopic_paths_sym) + apply (rule homotopic_paths_reparametrize + [where f = "\t. if t \ 1 / 2 then inverse 2 *\<^sub>R t + else if t \ 3 / 4 then t - (1 / 4) + else 2 *\<^sub>R t - 1"]) + apply (simp_all del: le_divide_eq_numeral1) + apply (simp add: subset_path_image_join) + apply (rule continuous_on_cases_1 continuous_intros)+ + apply (auto simp: joinpaths_def) + done + +proposition homotopic_paths_rinv: + assumes "path p" "path_image p \ s" + shows "homotopic_paths s (p +++ reversepath p) (linepath (pathstart p) (pathstart p))" +proof - + have "continuous_on ({0..1} \ {0..1}) (\x. (subpath 0 (fst x) p +++ reversepath (subpath 0 (fst x) p)) (snd x))" + using assms + apply (simp add: joinpaths_def subpath_def reversepath_def path_def del: le_divide_eq_numeral1) + apply (rule continuous_on_cases_le) + apply (rule_tac [2] continuous_on_compose [of _ _ p, unfolded o_def]) + apply (rule continuous_on_compose [of _ _ p, unfolded o_def]) + apply (auto intro!: continuous_intros simp del: eq_divide_eq_numeral1) + apply (force elim!: continuous_on_subset simp add: mult_le_one)+ + done + then show ?thesis + using assms + apply (subst homotopic_paths_sym_eq) + unfolding homotopic_paths_def homotopic_with_def + apply (rule_tac x="(\y. (subpath 0 (fst y) p +++ reversepath(subpath 0 (fst y) p)) (snd y))" in exI) + apply (simp add: path_defs joinpaths_def subpath_def reversepath_def) + apply (force simp: mult_le_one) + done +qed + +proposition homotopic_paths_linv: + assumes "path p" "path_image p \ s" + shows "homotopic_paths s (reversepath p +++ p) (linepath (pathfinish p) (pathfinish p))" + using homotopic_paths_rinv [of "reversepath p" s] assms by simp + + +subsection\Homotopy of loops without requiring preservation of endpoints\ + +definition%important homotopic_loops :: "'a::topological_space set \ (real \ 'a) \ (real \ 'a) \ bool" where + "homotopic_loops s p q \ + homotopic_with (\r. pathfinish r = pathstart r) {0..1} s p q" + +lemma homotopic_loops: + "homotopic_loops s p q \ + (\h. continuous_on ({0..1::real} \ {0..1}) h \ + image h ({0..1} \ {0..1}) \ s \ + (\x \ {0..1}. h(0,x) = p x) \ + (\x \ {0..1}. h(1,x) = q x) \ + (\t \ {0..1}. pathfinish(h \ Pair t) = pathstart(h \ Pair t)))" + by (simp add: homotopic_loops_def pathstart_def pathfinish_def homotopic_with) + +proposition homotopic_loops_imp_loop: + "homotopic_loops s p q \ pathfinish p = pathstart p \ pathfinish q = pathstart q" +using homotopic_with_imp_property homotopic_loops_def by blast + +proposition homotopic_loops_imp_path: + "homotopic_loops s p q \ path p \ path q" + unfolding homotopic_loops_def path_def + using homotopic_with_imp_continuous by blast + +proposition homotopic_loops_imp_subset: + "homotopic_loops s p q \ path_image p \ s \ path_image q \ s" + unfolding homotopic_loops_def path_image_def + by (metis homotopic_with_imp_subset1 homotopic_with_imp_subset2) + +proposition homotopic_loops_refl: + "homotopic_loops s p p \ + path p \ path_image p \ s \ pathfinish p = pathstart p" + by (simp add: homotopic_loops_def homotopic_with_refl path_image_def path_def) + +proposition homotopic_loops_sym: "homotopic_loops s p q \ homotopic_loops s q p" + by (simp add: homotopic_loops_def homotopic_with_sym) + +proposition homotopic_loops_sym_eq: "homotopic_loops s p q \ homotopic_loops s q p" + by (metis homotopic_loops_sym) + +proposition homotopic_loops_trans: + "\homotopic_loops s p q; homotopic_loops s q r\ \ homotopic_loops s p r" + unfolding homotopic_loops_def by (blast intro: homotopic_with_trans) + +proposition homotopic_loops_subset: + "\homotopic_loops s p q; s \ t\ \ homotopic_loops t p q" + by (simp add: homotopic_loops_def homotopic_with_subset_right) + +proposition homotopic_loops_eq: + "\path p; path_image p \ s; pathfinish p = pathstart p; \t. t \ {0..1} \ p(t) = q(t)\ + \ homotopic_loops s p q" + unfolding homotopic_loops_def + apply (rule homotopic_with_eq) + apply (rule homotopic_with_refl [where f = p, THEN iffD2]) + apply (simp_all add: path_image_def path_def pathstart_def pathfinish_def) + done + +proposition homotopic_loops_continuous_image: + "\homotopic_loops s f g; continuous_on s h; h ` s \ t\ \ homotopic_loops t (h \ f) (h \ g)" + unfolding homotopic_loops_def + apply (rule homotopic_with_compose_continuous_left) + apply (erule homotopic_with_mono) + by (simp add: pathfinish_def pathstart_def) + + +subsection\Relations between the two variants of homotopy\ + +proposition homotopic_paths_imp_homotopic_loops: + "\homotopic_paths s p q; pathfinish p = pathstart p; pathfinish q = pathstart p\ \ homotopic_loops s p q" + by (auto simp: homotopic_paths_def homotopic_loops_def intro: homotopic_with_mono) + +proposition homotopic_loops_imp_homotopic_paths_null: + assumes "homotopic_loops s p (linepath a a)" + shows "homotopic_paths s p (linepath (pathstart p) (pathstart p))" +proof - + have "path p" by (metis assms homotopic_loops_imp_path) + have ploop: "pathfinish p = pathstart p" by (metis assms homotopic_loops_imp_loop) + have pip: "path_image p \ s" by (metis assms homotopic_loops_imp_subset) + obtain h where conth: "continuous_on ({0..1::real} \ {0..1}) h" + and hs: "h ` ({0..1} \ {0..1}) \ s" + and [simp]: "\x. x \ {0..1} \ h(0,x) = p x" + and [simp]: "\x. x \ {0..1} \ h(1,x) = a" + and ends: "\t. t \ {0..1} \ pathfinish (h \ Pair t) = pathstart (h \ Pair t)" + using assms by (auto simp: homotopic_loops homotopic_with) + have conth0: "path (\u. h (u, 0))" + unfolding path_def + apply (rule continuous_on_compose [of _ _ h, unfolded o_def]) + apply (force intro: continuous_intros continuous_on_subset [OF conth])+ + done + have pih0: "path_image (\u. h (u, 0)) \ s" + using hs by (force simp: path_image_def) + have c1: "continuous_on ({0..1} \ {0..1}) (\x. h (fst x * snd x, 0))" + apply (rule continuous_on_compose [of _ _ h, unfolded o_def]) + apply (force simp: mult_le_one intro: continuous_intros continuous_on_subset [OF conth])+ + done + have c2: "continuous_on ({0..1} \ {0..1}) (\x. h (fst x - fst x * snd x, 0))" + apply (rule continuous_on_compose [of _ _ h, unfolded o_def]) + apply (force simp: mult_left_le mult_le_one intro: continuous_intros continuous_on_subset [OF conth])+ + apply (rule continuous_on_subset [OF conth]) + apply (auto simp: algebra_simps add_increasing2 mult_left_le) + done + have [simp]: "\t. \0 \ t \ t \ 1\ \ h (t, 1) = h (t, 0)" + using ends by (simp add: pathfinish_def pathstart_def) + have adhoc_le: "c * 4 \ 1 + c * (d * 4)" if "\ d * 4 \ 3" "0 \ c" "c \ 1" for c d::real + proof - + have "c * 3 \ c * (d * 4)" using that less_eq_real_def by auto + with \c \ 1\ show ?thesis by fastforce + qed + have *: "\p x. (path p \ path(reversepath p)) \ + (path_image p \ s \ path_image(reversepath p) \ s) \ + (pathfinish p = pathstart(linepath a a +++ reversepath p) \ + pathstart(reversepath p) = a) \ pathstart p = x + \ homotopic_paths s (p +++ linepath a a +++ reversepath p) (linepath x x)" + by (metis homotopic_paths_lid homotopic_paths_join + homotopic_paths_trans homotopic_paths_sym homotopic_paths_rinv) + have 1: "homotopic_paths s p (p +++ linepath (pathfinish p) (pathfinish p))" + using \path p\ homotopic_paths_rid homotopic_paths_sym pip by blast + moreover have "homotopic_paths s (p +++ linepath (pathfinish p) (pathfinish p)) + (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p))" + apply (rule homotopic_paths_sym) + using homotopic_paths_lid [of "p +++ linepath (pathfinish p) (pathfinish p)" s] + by (metis 1 homotopic_paths_imp_path homotopic_paths_imp_pathstart homotopic_paths_imp_subset) + moreover have "homotopic_paths s (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p)) + ((\u. h (u, 0)) +++ linepath a a +++ reversepath (\u. h (u, 0)))" + apply (simp add: homotopic_paths_def homotopic_with_def) + apply (rule_tac x="\y. (subpath 0 (fst y) (\u. h (u, 0)) +++ (\u. h (Pair (fst y) u)) +++ subpath (fst y) 0 (\u. h (u, 0))) (snd y)" in exI) + apply (simp add: subpath_reversepath) + apply (intro conjI homotopic_join_lemma) + using ploop + apply (simp_all add: path_defs joinpaths_def o_def subpath_def conth c1 c2) + apply (force simp: algebra_simps mult_le_one mult_left_le intro: hs [THEN subsetD] adhoc_le) + done + moreover have "homotopic_paths s ((\u. h (u, 0)) +++ linepath a a +++ reversepath (\u. h (u, 0))) + (linepath (pathstart p) (pathstart p))" + apply (rule *) + apply (simp add: pih0 pathstart_def pathfinish_def conth0) + apply (simp add: reversepath_def joinpaths_def) + done + ultimately show ?thesis + by (blast intro: homotopic_paths_trans) +qed + +proposition homotopic_loops_conjugate: + fixes s :: "'a::real_normed_vector set" + assumes "path p" "path q" and pip: "path_image p \ s" and piq: "path_image q \ s" + and papp: "pathfinish p = pathstart q" and qloop: "pathfinish q = pathstart q" + shows "homotopic_loops s (p +++ q +++ reversepath p) q" +proof - + have contp: "continuous_on {0..1} p" using \path p\ [unfolded path_def] by blast + have contq: "continuous_on {0..1} q" using \path q\ [unfolded path_def] by blast + have c1: "continuous_on ({0..1} \ {0..1}) (\x. p ((1 - fst x) * snd x + fst x))" + apply (rule continuous_on_compose [of _ _ p, unfolded o_def]) + apply (force simp: mult_le_one intro!: continuous_intros) + apply (rule continuous_on_subset [OF contp]) + apply (auto simp: algebra_simps add_increasing2 mult_right_le_one_le sum_le_prod1) + done + have c2: "continuous_on ({0..1} \ {0..1}) (\x. p ((fst x - 1) * snd x + 1))" + apply (rule continuous_on_compose [of _ _ p, unfolded o_def]) + apply (force simp: mult_le_one intro!: continuous_intros) + apply (rule continuous_on_subset [OF contp]) + apply (auto simp: algebra_simps add_increasing2 mult_left_le_one_le) + done + have ps1: "\a b. \b * 2 \ 1; 0 \ b; 0 \ a; a \ 1\ \ p ((1 - a) * (2 * b) + a) \ s" + using sum_le_prod1 + by (force simp: algebra_simps add_increasing2 mult_left_le intro: pip [unfolded path_image_def, THEN subsetD]) + have ps2: "\a b. \\ 4 * b \ 3; b \ 1; 0 \ a; a \ 1\ \ p ((a - 1) * (4 * b - 3) + 1) \ s" + apply (rule pip [unfolded path_image_def, THEN subsetD]) + apply (rule image_eqI, blast) + apply (simp add: algebra_simps) + by (metis add_mono_thms_linordered_semiring(1) affine_ineq linear mult.commute mult.left_neutral mult_right_mono not_le + add.commute zero_le_numeral) + have qs: "\a b. \4 * b \ 3; \ b * 2 \ 1\ \ q (4 * b - 2) \ s" + using path_image_def piq by fastforce + have "homotopic_loops s (p +++ q +++ reversepath p) + (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q))" + apply (simp add: homotopic_loops_def homotopic_with_def) + apply (rule_tac x="(\y. (subpath (fst y) 1 p +++ q +++ subpath 1 (fst y) p) (snd y))" in exI) + apply (simp add: subpath_refl subpath_reversepath) + apply (intro conjI homotopic_join_lemma) + using papp qloop + apply (simp_all add: path_defs joinpaths_def o_def subpath_def c1 c2) + apply (force simp: contq intro: continuous_on_compose [of _ _ q, unfolded o_def] continuous_on_id continuous_on_snd) + apply (auto simp: ps1 ps2 qs) + done + moreover have "homotopic_loops s (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) q" + proof - + have "homotopic_paths s (linepath (pathfinish q) (pathfinish q) +++ q) q" + using \path q\ homotopic_paths_lid qloop piq by auto + hence 1: "\f. homotopic_paths s f q \ \ homotopic_paths s f (linepath (pathfinish q) (pathfinish q) +++ q)" + using homotopic_paths_trans by blast + hence "homotopic_paths s (linepath (pathfinish q) (pathfinish q) +++ q +++ linepath (pathfinish q) (pathfinish q)) q" + proof - + have "homotopic_paths s (q +++ linepath (pathfinish q) (pathfinish q)) q" + by (simp add: \path q\ homotopic_paths_rid piq) + thus ?thesis + by (metis (no_types) 1 \path q\ homotopic_paths_join homotopic_paths_rinv homotopic_paths_sym + homotopic_paths_trans qloop pathfinish_linepath piq) + qed + thus ?thesis + by (metis (no_types) qloop homotopic_loops_sym homotopic_paths_imp_homotopic_loops homotopic_paths_imp_pathfinish homotopic_paths_sym) + qed + ultimately show ?thesis + by (blast intro: homotopic_loops_trans) +qed + +lemma homotopic_paths_loop_parts: + assumes loops: "homotopic_loops S (p +++ reversepath q) (linepath a a)" and "path q" + shows "homotopic_paths S p q" +proof - + have paths: "homotopic_paths S (p +++ reversepath q) (linepath (pathstart p) (pathstart p))" + using homotopic_loops_imp_homotopic_paths_null [OF loops] by simp + then have "path p" + using \path q\ homotopic_loops_imp_path loops path_join path_join_path_ends path_reversepath by blast + show ?thesis + proof (cases "pathfinish p = pathfinish q") + case True + have pipq: "path_image p \ S" "path_image q \ S" + by (metis Un_subset_iff paths \path p\ \path q\ homotopic_loops_imp_subset homotopic_paths_imp_path loops + path_image_join path_image_reversepath path_imp_reversepath path_join_eq)+ + have "homotopic_paths S p (p +++ (linepath (pathfinish p) (pathfinish p)))" + using \path p\ \path_image p \ S\ homotopic_paths_rid homotopic_paths_sym by blast + moreover have "homotopic_paths S (p +++ (linepath (pathfinish p) (pathfinish p))) (p +++ (reversepath q +++ q))" + by (simp add: True \path p\ \path q\ pipq homotopic_paths_join homotopic_paths_linv homotopic_paths_sym) + moreover have "homotopic_paths S (p +++ (reversepath q +++ q)) ((p +++ reversepath q) +++ q)" + by (simp add: True \path p\ \path q\ homotopic_paths_assoc pipq) + moreover have "homotopic_paths S ((p +++ reversepath q) +++ q) (linepath (pathstart p) (pathstart p) +++ q)" + by (simp add: \path q\ homotopic_paths_join paths pipq) + moreover then have "homotopic_paths S (linepath (pathstart p) (pathstart p) +++ q) q" + by (metis \path q\ homotopic_paths_imp_path homotopic_paths_lid linepath_trivial path_join_path_ends pathfinish_def pipq(2)) + ultimately show ?thesis + using homotopic_paths_trans by metis + next + case False + then show ?thesis + using \path q\ homotopic_loops_imp_path loops path_join_path_ends by fastforce + qed +qed + + +subsection%unimportant\Homotopy of "nearby" function, paths and loops\ + +lemma homotopic_with_linear: + fixes f g :: "_ \ 'b::real_normed_vector" + assumes contf: "continuous_on s f" + and contg:"continuous_on s g" + and sub: "\x. x \ s \ closed_segment (f x) (g x) \ t" + shows "homotopic_with (\z. True) s t f g" + apply (simp add: homotopic_with_def) + apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R f(snd y) + (fst y) *\<^sub>R g(snd y))" in exI) + apply (intro conjI) + apply (rule subset_refl continuous_intros continuous_on_subset [OF contf] continuous_on_compose2 [where g=f] + continuous_on_subset [OF contg] continuous_on_compose2 [where g=g]| simp)+ + using sub closed_segment_def apply fastforce+ + done + +lemma homotopic_paths_linear: + fixes g h :: "real \ 'a::real_normed_vector" + assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" + "\t. t \ {0..1} \ closed_segment (g t) (h t) \ s" + shows "homotopic_paths s g h" + using assms + unfolding path_def + apply (simp add: closed_segment_def pathstart_def pathfinish_def homotopic_paths_def homotopic_with_def) + apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R (g \ snd) y + (fst y) *\<^sub>R (h \ snd) y)" in exI) + apply (intro conjI subsetI continuous_intros; force) + done + +lemma homotopic_loops_linear: + fixes g h :: "real \ 'a::real_normed_vector" + assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" + "\t x. t \ {0..1} \ closed_segment (g t) (h t) \ s" + shows "homotopic_loops s g h" + using assms + unfolding path_def + apply (simp add: pathstart_def pathfinish_def homotopic_loops_def homotopic_with_def) + apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R g(snd y) + (fst y) *\<^sub>R h(snd y))" in exI) + apply (auto intro!: continuous_intros intro: continuous_on_compose2 [where g=g] continuous_on_compose2 [where g=h]) + apply (force simp: closed_segment_def) + done + +lemma homotopic_paths_nearby_explicit: + assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" + and no: "\t x. \t \ {0..1}; x \ s\ \ norm(h t - g t) < norm(g t - x)" + shows "homotopic_paths s g h" + apply (rule homotopic_paths_linear [OF assms(1-4)]) + by (metis no segment_bound(1) subsetI norm_minus_commute not_le) + +lemma homotopic_loops_nearby_explicit: + assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" + and no: "\t x. \t \ {0..1}; x \ s\ \ norm(h t - g t) < norm(g t - x)" + shows "homotopic_loops s g h" + apply (rule homotopic_loops_linear [OF assms(1-4)]) + by (metis no segment_bound(1) subsetI norm_minus_commute not_le) + +lemma homotopic_nearby_paths: + fixes g h :: "real \ 'a::euclidean_space" + assumes "path g" "open s" "path_image g \ s" + shows "\e. 0 < e \ + (\h. path h \ + pathstart h = pathstart g \ pathfinish h = pathfinish g \ + (\t \ {0..1}. norm(h t - g t) < e) \ homotopic_paths s g h)" +proof - + obtain e where "e > 0" and e: "\x y. x \ path_image g \ y \ - s \ e \ dist x y" + using separate_compact_closed [of "path_image g" "-s"] assms by force + show ?thesis + apply (intro exI conjI) + using e [unfolded dist_norm] + apply (auto simp: intro!: homotopic_paths_nearby_explicit assms \e > 0\) + by (metis atLeastAtMost_iff imageI le_less_trans not_le path_image_def) +qed + +lemma homotopic_nearby_loops: + fixes g h :: "real \ 'a::euclidean_space" + assumes "path g" "open s" "path_image g \ s" "pathfinish g = pathstart g" + shows "\e. 0 < e \ + (\h. path h \ pathfinish h = pathstart h \ + (\t \ {0..1}. norm(h t - g t) < e) \ homotopic_loops s g h)" +proof - + obtain e where "e > 0" and e: "\x y. x \ path_image g \ y \ - s \ e \ dist x y" + using separate_compact_closed [of "path_image g" "-s"] assms by force + show ?thesis + apply (intro exI conjI) + using e [unfolded dist_norm] + apply (auto simp: intro!: homotopic_loops_nearby_explicit assms \e > 0\) + by (metis atLeastAtMost_iff imageI le_less_trans not_le path_image_def) +qed + + +subsection\ Homotopy and subpaths\ + +lemma homotopic_join_subpaths1: + assumes "path g" and pag: "path_image g \ s" + and u: "u \ {0..1}" and v: "v \ {0..1}" and w: "w \ {0..1}" "u \ v" "v \ w" + shows "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" +proof - + have 1: "t * 2 \ 1 \ u + t * (v * 2) \ v + t * (u * 2)" for t + using affine_ineq \u \ v\ by fastforce + have 2: "t * 2 > 1 \ u + (2*t - 1) * v \ v + (2*t - 1) * w" for t + by (metis add_mono_thms_linordered_semiring(1) diff_gt_0_iff_gt less_eq_real_def mult.commute mult_right_mono \u \ v\ \v \ w\) + have t2: "\t::real. t*2 = 1 \ t = 1/2" by auto + show ?thesis + apply (rule homotopic_paths_subset [OF _ pag]) + using assms + apply (cases "w = u") + using homotopic_paths_rinv [of "subpath u v g" "path_image g"] + apply (force simp: closed_segment_eq_real_ivl image_mono path_image_def subpath_refl) + apply (rule homotopic_paths_sym) + apply (rule homotopic_paths_reparametrize + [where f = "\t. if t \ 1 / 2 + then inverse((w - u)) *\<^sub>R (2 * (v - u)) *\<^sub>R t + else inverse((w - u)) *\<^sub>R ((v - u) + (w - v) *\<^sub>R (2 *\<^sub>R t - 1))"]) + using \path g\ path_subpath u w apply blast + using \path g\ path_image_subpath_subset u w(1) apply blast + apply simp_all + apply (subst split_01) + apply (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+ + apply (simp_all add: field_simps not_le) + apply (force dest!: t2) + apply (force simp: algebra_simps mult_left_mono affine_ineq dest!: 1 2) + apply (simp add: joinpaths_def subpath_def) + apply (force simp: algebra_simps) + done +qed + +lemma homotopic_join_subpaths2: + assumes "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" + shows "homotopic_paths s (subpath w v g +++ subpath v u g) (subpath w u g)" +by (metis assms homotopic_paths_reversepath_D pathfinish_subpath pathstart_subpath reversepath_joinpaths reversepath_subpath) + +lemma homotopic_join_subpaths3: + assumes hom: "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" + and "path g" and pag: "path_image g \ s" + and u: "u \ {0..1}" and v: "v \ {0..1}" and w: "w \ {0..1}" + shows "homotopic_paths s (subpath v w g +++ subpath w u g) (subpath v u g)" +proof - + have "homotopic_paths s (subpath u w g +++ subpath w v g) ((subpath u v g +++ subpath v w g) +++ subpath w v g)" + apply (rule homotopic_paths_join) + using hom homotopic_paths_sym_eq apply blast + apply (metis \path g\ homotopic_paths_eq pag path_image_subpath_subset path_subpath subset_trans v w, simp) + done + also have "homotopic_paths s ((subpath u v g +++ subpath v w g) +++ subpath w v g) (subpath u v g +++ subpath v w g +++ subpath w v g)" + apply (rule homotopic_paths_sym [OF homotopic_paths_assoc]) + using assms by (simp_all add: path_image_subpath_subset [THEN order_trans]) + also have "homotopic_paths s (subpath u v g +++ subpath v w g +++ subpath w v g) + (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))" + apply (rule homotopic_paths_join) + apply (metis \path g\ homotopic_paths_eq order.trans pag path_image_subpath_subset path_subpath u v) + apply (metis (no_types, lifting) \path g\ homotopic_paths_linv order_trans pag path_image_subpath_subset path_subpath pathfinish_subpath reversepath_subpath v w) + apply simp + done + also have "homotopic_paths s (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g))) (subpath u v g)" + apply (rule homotopic_paths_rid) + using \path g\ path_subpath u v apply blast + apply (meson \path g\ order.trans pag path_image_subpath_subset u v) + done + finally have "homotopic_paths s (subpath u w g +++ subpath w v g) (subpath u v g)" . + then show ?thesis + using homotopic_join_subpaths2 by blast +qed + +proposition homotopic_join_subpaths: + "\path g; path_image g \ s; u \ {0..1}; v \ {0..1}; w \ {0..1}\ + \ homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" + apply (rule le_cases3 [of u v w]) +using homotopic_join_subpaths1 homotopic_join_subpaths2 homotopic_join_subpaths3 by metis+ + +text\Relating homotopy of trivial loops to path-connectedness.\ + +lemma path_component_imp_homotopic_points: + "path_component S a b \ homotopic_loops S (linepath a a) (linepath b b)" +apply (simp add: path_component_def homotopic_loops_def homotopic_with_def + pathstart_def pathfinish_def path_image_def path_def, clarify) +apply (rule_tac x="g \ fst" in exI) +apply (intro conjI continuous_intros continuous_on_compose)+ +apply (auto elim!: continuous_on_subset) +done + +lemma homotopic_loops_imp_path_component_value: + "\homotopic_loops S p q; 0 \ t; t \ 1\ + \ path_component S (p t) (q t)" +apply (simp add: path_component_def homotopic_loops_def homotopic_with_def + pathstart_def pathfinish_def path_image_def path_def, clarify) +apply (rule_tac x="h \ (\u. (u, t))" in exI) +apply (intro conjI continuous_intros continuous_on_compose)+ +apply (auto elim!: continuous_on_subset) +done + +lemma homotopic_points_eq_path_component: + "homotopic_loops S (linepath a a) (linepath b b) \ + path_component S a b" +by (auto simp: path_component_imp_homotopic_points + dest: homotopic_loops_imp_path_component_value [where t=1]) + +lemma path_connected_eq_homotopic_points: + "path_connected S \ + (\a b. a \ S \ b \ S \ homotopic_loops S (linepath a a) (linepath b b))" +by (auto simp: path_connected_def path_component_def homotopic_points_eq_path_component) + + +subsection\Simply connected sets\ + +text%important\defined as "all loops are homotopic (as loops)\ + +definition%important simply_connected where + "simply_connected S \ + \p q. path p \ pathfinish p = pathstart p \ path_image p \ S \ + path q \ pathfinish q = pathstart q \ path_image q \ S + \ homotopic_loops S p q" + +lemma simply_connected_empty [iff]: "simply_connected {}" + by (simp add: simply_connected_def) + +lemma simply_connected_imp_path_connected: + fixes S :: "_::real_normed_vector set" + shows "simply_connected S \ path_connected S" +by (simp add: simply_connected_def path_connected_eq_homotopic_points) + +lemma simply_connected_imp_connected: + fixes S :: "_::real_normed_vector set" + shows "simply_connected S \ connected S" +by (simp add: path_connected_imp_connected simply_connected_imp_path_connected) + +lemma simply_connected_eq_contractible_loop_any: + fixes S :: "_::real_normed_vector set" + shows "simply_connected S \ + (\p a. path p \ path_image p \ S \ + pathfinish p = pathstart p \ a \ S + \ homotopic_loops S p (linepath a a))" +apply (simp add: simply_connected_def) +apply (rule iffI, force, clarify) +apply (rule_tac q = "linepath (pathstart p) (pathstart p)" in homotopic_loops_trans) +apply (fastforce simp add:) +using homotopic_loops_sym apply blast +done + +lemma simply_connected_eq_contractible_loop_some: + fixes S :: "_::real_normed_vector set" + shows "simply_connected S \ + path_connected S \ + (\p. path p \ path_image p \ S \ pathfinish p = pathstart p + \ (\a. a \ S \ homotopic_loops S p (linepath a a)))" +apply (rule iffI) + apply (fastforce simp: simply_connected_imp_path_connected simply_connected_eq_contractible_loop_any) +apply (clarsimp simp add: simply_connected_eq_contractible_loop_any) +apply (drule_tac x=p in spec) +using homotopic_loops_trans path_connected_eq_homotopic_points + apply blast +done + +lemma simply_connected_eq_contractible_loop_all: + fixes S :: "_::real_normed_vector set" + shows "simply_connected S \ + S = {} \ + (\a \ S. \p. path p \ path_image p \ S \ pathfinish p = pathstart p + \ homotopic_loops S p (linepath a a))" + (is "?lhs = ?rhs") +proof (cases "S = {}") + case True then show ?thesis by force +next + case False + then obtain a where "a \ S" by blast + show ?thesis + proof + assume "simply_connected S" + then show ?rhs + using \a \ S\ \simply_connected S\ simply_connected_eq_contractible_loop_any + by blast + next + assume ?rhs + then show "simply_connected S" + apply (simp add: simply_connected_eq_contractible_loop_any False) + by (meson homotopic_loops_refl homotopic_loops_sym homotopic_loops_trans + path_component_imp_homotopic_points path_component_refl) + qed +qed + +lemma simply_connected_eq_contractible_path: + fixes S :: "_::real_normed_vector set" + shows "simply_connected S \ + path_connected S \ + (\p. path p \ path_image p \ S \ pathfinish p = pathstart p + \ homotopic_paths S p (linepath (pathstart p) (pathstart p)))" +apply (rule iffI) + apply (simp add: simply_connected_imp_path_connected) + apply (metis simply_connected_eq_contractible_loop_some homotopic_loops_imp_homotopic_paths_null) +by (meson homotopic_paths_imp_homotopic_loops pathfinish_linepath pathstart_in_path_image + simply_connected_eq_contractible_loop_some subset_iff) + +lemma simply_connected_eq_homotopic_paths: + fixes S :: "_::real_normed_vector set" + shows "simply_connected S \ + path_connected S \ + (\p q. path p \ path_image p \ S \ + path q \ path_image q \ S \ + pathstart q = pathstart p \ pathfinish q = pathfinish p + \ homotopic_paths S p q)" + (is "?lhs = ?rhs") +proof + assume ?lhs + then have pc: "path_connected S" + and *: "\p. \path p; path_image p \ S; + pathfinish p = pathstart p\ + \ homotopic_paths S p (linepath (pathstart p) (pathstart p))" + by (auto simp: simply_connected_eq_contractible_path) + have "homotopic_paths S p q" + if "path p" "path_image p \ S" "path q" + "path_image q \ S" "pathstart q = pathstart p" + "pathfinish q = pathfinish p" for p q + proof - + have "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))" + by (simp add: homotopic_paths_rid homotopic_paths_sym that) + also have "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p)) + (p +++ reversepath q +++ q)" + using that + by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl homotopic_paths_sym_eq pathstart_linepath) + also have "homotopic_paths S (p +++ reversepath q +++ q) + ((p +++ reversepath q) +++ q)" + by (simp add: that homotopic_paths_assoc) + also have "homotopic_paths S ((p +++ reversepath q) +++ q) + (linepath (pathstart q) (pathstart q) +++ q)" + using * [of "p +++ reversepath q"] that + by (simp add: homotopic_paths_join path_image_join) + also have "homotopic_paths S (linepath (pathstart q) (pathstart q) +++ q) q" + using that homotopic_paths_lid by blast + finally show ?thesis . + qed + then show ?rhs + by (blast intro: pc *) +next + assume ?rhs + then show ?lhs + by (force simp: simply_connected_eq_contractible_path) +qed + +proposition simply_connected_Times: + fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" + assumes S: "simply_connected S" and T: "simply_connected T" + shows "simply_connected(S \ T)" +proof - + have "homotopic_loops (S \ T) p (linepath (a, b) (a, b))" + if "path p" "path_image p \ S \ T" "p 1 = p 0" "a \ S" "b \ T" + for p a b + proof - + have "path (fst \ p)" + apply (rule Path_Connected.path_continuous_image [OF \path p\]) + apply (rule continuous_intros)+ + done + moreover have "path_image (fst \ p) \ S" + using that apply (simp add: path_image_def) by force + ultimately have p1: "homotopic_loops S (fst \ p) (linepath a a)" + using S that + apply (simp add: simply_connected_eq_contractible_loop_any) + apply (drule_tac x="fst \ p" in spec) + apply (drule_tac x=a in spec) + apply (auto simp: pathstart_def pathfinish_def) + done + have "path (snd \ p)" + apply (rule Path_Connected.path_continuous_image [OF \path p\]) + apply (rule continuous_intros)+ + done + moreover have "path_image (snd \ p) \ T" + using that apply (simp add: path_image_def) by force + ultimately have p2: "homotopic_loops T (snd \ p) (linepath b b)" + using T that + apply (simp add: simply_connected_eq_contractible_loop_any) + apply (drule_tac x="snd \ p" in spec) + apply (drule_tac x=b in spec) + apply (auto simp: pathstart_def pathfinish_def) + done + show ?thesis + using p1 p2 + apply (simp add: homotopic_loops, clarify) + apply (rename_tac h k) + apply (rule_tac x="\z. Pair (h z) (k z)" in exI) + apply (intro conjI continuous_intros | assumption)+ + apply (auto simp: pathstart_def pathfinish_def) + done + qed + with assms show ?thesis + by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) +qed + + +subsection\Contractible sets\ + +definition%important contractible where + "contractible S \ \a. homotopic_with (\x. True) S S id (\x. a)" + +proposition contractible_imp_simply_connected: + fixes S :: "_::real_normed_vector set" + assumes "contractible S" shows "simply_connected S" +proof (cases "S = {}") + case True then show ?thesis by force +next + case False + obtain a where a: "homotopic_with (\x. True) S S id (\x. a)" + using assms by (force simp: contractible_def) + then have "a \ S" + by (metis False homotopic_constant_maps homotopic_with_symD homotopic_with_trans path_component_mem(2)) + show ?thesis + apply (simp add: simply_connected_eq_contractible_loop_all False) + apply (rule bexI [OF _ \a \ S\]) + using a apply (simp add: homotopic_loops_def homotopic_with_def path_def path_image_def pathfinish_def pathstart_def, clarify) + apply (rule_tac x="(h \ (\y. (fst y, (p \ snd) y)))" in exI) + apply (intro conjI continuous_on_compose continuous_intros) + apply (erule continuous_on_subset | force)+ + done +qed + +corollary contractible_imp_connected: + fixes S :: "_::real_normed_vector set" + shows "contractible S \ connected S" +by (simp add: contractible_imp_simply_connected simply_connected_imp_connected) + +lemma contractible_imp_path_connected: + fixes S :: "_::real_normed_vector set" + shows "contractible S \ path_connected S" +by (simp add: contractible_imp_simply_connected simply_connected_imp_path_connected) + +lemma nullhomotopic_through_contractible: + fixes S :: "_::topological_space set" + assumes f: "continuous_on S f" "f ` S \ T" + and g: "continuous_on T g" "g ` T \ U" + and T: "contractible T" + obtains c where "homotopic_with (\h. True) S U (g \ f) (\x. c)" +proof - + obtain b where b: "homotopic_with (\x. True) T T id (\x. b)" + using assms by (force simp: contractible_def) + have "homotopic_with (\f. True) T U (g \ id) (g \ (\x. b))" + by (rule homotopic_compose_continuous_left [OF b g]) + then have "homotopic_with (\f. True) S U (g \ id \ f) (g \ (\x. b) \ f)" + by (rule homotopic_compose_continuous_right [OF _ f]) + then show ?thesis + by (simp add: comp_def that) +qed + +lemma nullhomotopic_into_contractible: + assumes f: "continuous_on S f" "f ` S \ T" + and T: "contractible T" + obtains c where "homotopic_with (\h. True) S T f (\x. c)" +apply (rule nullhomotopic_through_contractible [OF f, of id T]) +using assms +apply (auto simp: continuous_on_id) +done + +lemma nullhomotopic_from_contractible: + assumes f: "continuous_on S f" "f ` S \ T" + and S: "contractible S" + obtains c where "homotopic_with (\h. True) S T f (\x. c)" +apply (rule nullhomotopic_through_contractible [OF continuous_on_id _ f S, of S]) +using assms +apply (auto simp: comp_def) +done + +lemma homotopic_through_contractible: + fixes S :: "_::real_normed_vector set" + assumes "continuous_on S f1" "f1 ` S \ T" + "continuous_on T g1" "g1 ` T \ U" + "continuous_on S f2" "f2 ` S \ T" + "continuous_on T g2" "g2 ` T \ U" + "contractible T" "path_connected U" + shows "homotopic_with (\h. True) S U (g1 \ f1) (g2 \ f2)" +proof - + obtain c1 where c1: "homotopic_with (\h. True) S U (g1 \ f1) (\x. c1)" + apply (rule nullhomotopic_through_contractible [of S f1 T g1 U]) + using assms apply auto + done + obtain c2 where c2: "homotopic_with (\h. True) S U (g2 \ f2) (\x. c2)" + apply (rule nullhomotopic_through_contractible [of S f2 T g2 U]) + using assms apply auto + done + have *: "S = {} \ (\t. path_connected t \ t \ U \ c2 \ t \ c1 \ t)" + proof (cases "S = {}") + case True then show ?thesis by force + next + case False + with c1 c2 have "c1 \ U" "c2 \ U" + using homotopic_with_imp_subset2 all_not_in_conv image_subset_iff by blast+ + with \path_connected U\ show ?thesis by blast + qed + show ?thesis + apply (rule homotopic_with_trans [OF c1]) + apply (rule homotopic_with_symD) + apply (rule homotopic_with_trans [OF c2]) + apply (simp add: path_component homotopic_constant_maps *) + done +qed + +lemma homotopic_into_contractible: + fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" + assumes f: "continuous_on S f" "f ` S \ T" + and g: "continuous_on S g" "g ` S \ T" + and T: "contractible T" + shows "homotopic_with (\h. True) S T f g" +using homotopic_through_contractible [of S f T id T g id] +by (simp add: assms contractible_imp_path_connected continuous_on_id) + +lemma homotopic_from_contractible: + fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" + assumes f: "continuous_on S f" "f ` S \ T" + and g: "continuous_on S g" "g ` S \ T" + and "contractible S" "path_connected T" + shows "homotopic_with (\h. True) S T f g" +using homotopic_through_contractible [of S id S f T id g] +by (simp add: assms contractible_imp_path_connected continuous_on_id) + +lemma starlike_imp_contractible_gen: + fixes S :: "'a::real_normed_vector set" + assumes S: "starlike S" + and P: "\a T. \a \ S; 0 \ T; T \ 1\ \ P(\x. (1 - T) *\<^sub>R x + T *\<^sub>R a)" + obtains a where "homotopic_with P S S (\x. x) (\x. a)" +proof - + obtain a where "a \ S" and a: "\x. x \ S \ closed_segment a x \ S" + using S by (auto simp: starlike_def) + have "(\y. (1 - fst y) *\<^sub>R snd y + fst y *\<^sub>R a) ` ({0..1} \ S) \ S" + apply clarify + apply (erule a [unfolded closed_segment_def, THEN subsetD], simp) + apply (metis add_diff_cancel_right' diff_ge_0_iff_ge le_add_diff_inverse pth_c(1)) + done + then show ?thesis + apply (rule_tac a=a in that) + using \a \ S\ + apply (simp add: homotopic_with_def) + apply (rule_tac x="\y. (1 - (fst y)) *\<^sub>R snd y + (fst y) *\<^sub>R a" in exI) + apply (intro conjI ballI continuous_on_compose continuous_intros) + apply (simp_all add: P) + done +qed + +lemma starlike_imp_contractible: + fixes S :: "'a::real_normed_vector set" + shows "starlike S \ contractible S" +using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def) + +lemma contractible_UNIV [simp]: "contractible (UNIV :: 'a::real_normed_vector set)" + by (simp add: starlike_imp_contractible) + +lemma starlike_imp_simply_connected: + fixes S :: "'a::real_normed_vector set" + shows "starlike S \ simply_connected S" +by (simp add: contractible_imp_simply_connected starlike_imp_contractible) + +lemma convex_imp_simply_connected: + fixes S :: "'a::real_normed_vector set" + shows "convex S \ simply_connected S" +using convex_imp_starlike starlike_imp_simply_connected by blast + +lemma starlike_imp_path_connected: + fixes S :: "'a::real_normed_vector set" + shows "starlike S \ path_connected S" +by (simp add: simply_connected_imp_path_connected starlike_imp_simply_connected) + +lemma starlike_imp_connected: + fixes S :: "'a::real_normed_vector set" + shows "starlike S \ connected S" +by (simp add: path_connected_imp_connected starlike_imp_path_connected) + +lemma is_interval_simply_connected_1: + fixes S :: "real set" + shows "is_interval S \ simply_connected S" +using convex_imp_simply_connected is_interval_convex_1 is_interval_path_connected_1 simply_connected_imp_path_connected by auto + +lemma contractible_empty [simp]: "contractible {}" + by (simp add: contractible_def homotopic_with) + +lemma contractible_convex_tweak_boundary_points: + fixes S :: "'a::euclidean_space set" + assumes "convex S" and TS: "rel_interior S \ T" "T \ closure S" + shows "contractible T" +proof (cases "S = {}") + case True + with assms show ?thesis + by (simp add: subsetCE) +next + case False + show ?thesis + apply (rule starlike_imp_contractible) + apply (rule starlike_convex_tweak_boundary_points [OF \convex S\ False TS]) + done +qed + +lemma convex_imp_contractible: + fixes S :: "'a::real_normed_vector set" + shows "convex S \ contractible S" + using contractible_empty convex_imp_starlike starlike_imp_contractible by blast + +lemma contractible_sing [simp]: + fixes a :: "'a::real_normed_vector" + shows "contractible {a}" +by (rule convex_imp_contractible [OF convex_singleton]) + +lemma is_interval_contractible_1: + fixes S :: "real set" + shows "is_interval S \ contractible S" +using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1 + is_interval_simply_connected_1 by auto + +lemma contractible_Times: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes S: "contractible S" and T: "contractible T" + shows "contractible (S \ T)" +proof - + obtain a h where conth: "continuous_on ({0..1} \ S) h" + and hsub: "h ` ({0..1} \ S) \ S" + and [simp]: "\x. x \ S \ h (0, x) = x" + and [simp]: "\x. x \ S \ h (1::real, x) = a" + using S by (auto simp: contractible_def homotopic_with) + obtain b k where contk: "continuous_on ({0..1} \ T) k" + and ksub: "k ` ({0..1} \ T) \ T" + and [simp]: "\x. x \ T \ k (0, x) = x" + and [simp]: "\x. x \ T \ k (1::real, x) = b" + using T by (auto simp: contractible_def homotopic_with) + show ?thesis + apply (simp add: contractible_def homotopic_with) + apply (rule exI [where x=a]) + apply (rule exI [where x=b]) + apply (rule exI [where x = "\z. (h (fst z, fst(snd z)), k (fst z, snd(snd z)))"]) + apply (intro conjI ballI continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF contk]) + using hsub ksub + apply auto + done +qed + +lemma homotopy_dominated_contractibility: + fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" + assumes S: "contractible S" + and f: "continuous_on S f" "image f S \ T" + and g: "continuous_on T g" "image g T \ S" + and hom: "homotopic_with (\x. True) T T (f \ g) id" + shows "contractible T" +proof - + obtain b where "homotopic_with (\h. True) S T f (\x. b)" + using nullhomotopic_from_contractible [OF f S] . + then have homg: "homotopic_with (\x. True) T T ((\x. b) \ g) (f \ g)" + by (rule homotopic_with_compose_continuous_right [OF homotopic_with_symD g]) + show ?thesis + apply (simp add: contractible_def) + apply (rule exI [where x = b]) + apply (rule homotopic_with_symD) + apply (rule homotopic_with_trans [OF _ hom]) + using homg apply (simp add: o_def) + done +qed + + +subsection\Local versions of topological properties in general\ + +definition%important locally :: "('a::topological_space set \ bool) \ 'a set \ bool" +where + "locally P S \ + \w x. openin (subtopology euclidean S) w \ x \ w + \ (\u v. openin (subtopology euclidean S) u \ P v \ + x \ u \ u \ v \ v \ w)" + +lemma locallyI: + assumes "\w x. \openin (subtopology euclidean S) w; x \ w\ + \ \u v. openin (subtopology euclidean S) u \ P v \ + x \ u \ u \ v \ v \ w" + shows "locally P S" +using assms by (force simp: locally_def) + +lemma locallyE: + assumes "locally P S" "openin (subtopology euclidean S) w" "x \ w" + obtains u v where "openin (subtopology euclidean S) u" + "P v" "x \ u" "u \ v" "v \ w" + using assms unfolding locally_def by meson + +lemma locally_mono: + assumes "locally P S" "\t. P t \ Q t" + shows "locally Q S" +by (metis assms locally_def) + +lemma locally_open_subset: + assumes "locally P S" "openin (subtopology euclidean S) t" + shows "locally P t" +using assms +apply (simp add: locally_def) +apply (erule all_forward)+ +apply (rule impI) +apply (erule impCE) + using openin_trans apply blast +apply (erule ex_forward) +by (metis (no_types, hide_lams) Int_absorb1 Int_lower1 Int_subset_iff openin_open openin_subtopology_Int_subset) + +lemma locally_diff_closed: + "\locally P S; closedin (subtopology euclidean S) t\ \ locally P (S - t)" + using locally_open_subset closedin_def by fastforce + +lemma locally_empty [iff]: "locally P {}" + by (simp add: locally_def openin_subtopology) + +lemma locally_singleton [iff]: + fixes a :: "'a::metric_space" + shows "locally P {a} \ P {a}" +apply (simp add: locally_def openin_euclidean_subtopology_iff subset_singleton_iff conj_disj_distribR cong: conj_cong) +using zero_less_one by blast + +lemma locally_iff: + "locally P S \ + (\T x. open T \ x \ S \ T \ (\U. open U \ (\v. P v \ x \ S \ U \ S \ U \ v \ v \ S \ T)))" +apply (simp add: le_inf_iff locally_def openin_open, safe) +apply (metis IntE IntI le_inf_iff) +apply (metis IntI Int_subset_iff) +done + +lemma locally_Int: + assumes S: "locally P S" and t: "locally P t" + and P: "\S t. P S \ P t \ P(S \ t)" + shows "locally P (S \ t)" +using S t unfolding locally_iff +apply clarify +apply (drule_tac x=T in spec)+ +apply (drule_tac x=x in spec)+ +apply clarsimp +apply (rename_tac U1 U2 V1 V2) +apply (rule_tac x="U1 \ U2" in exI) +apply (simp add: open_Int) +apply (rule_tac x="V1 \ V2" in exI) +apply (auto intro: P) +done + +lemma locally_Times: + fixes S :: "('a::metric_space) set" and T :: "('b::metric_space) set" + assumes PS: "locally P S" and QT: "locally Q T" and R: "\S T. P S \ Q T \ R(S \ T)" + shows "locally R (S \ T)" + unfolding locally_def +proof (clarify) + fix W x y + assume W: "openin (subtopology euclidean (S \ T)) W" and xy: "(x, y) \ W" + then obtain U V where "openin (subtopology euclidean S) U" "x \ U" + "openin (subtopology euclidean T) V" "y \ V" "U \ V \ W" + using Times_in_interior_subtopology by metis + then obtain U1 U2 V1 V2 + where opeS: "openin (subtopology euclidean S) U1 \ P U2 \ x \ U1 \ U1 \ U2 \ U2 \ U" + and opeT: "openin (subtopology euclidean T) V1 \ Q V2 \ y \ V1 \ V1 \ V2 \ V2 \ V" + by (meson PS QT locallyE) + with \U \ V \ W\ show "\u v. openin (subtopology euclidean (S \ T)) u \ R v \ (x,y) \ u \ u \ v \ v \ W" + apply (rule_tac x="U1 \ V1" in exI) + apply (rule_tac x="U2 \ V2" in exI) + apply (auto simp: openin_Times R) + done +qed + + +proposition homeomorphism_locally_imp: + fixes S :: "'a::metric_space set" and t :: "'b::t2_space set" + assumes S: "locally P S" and hom: "homeomorphism S t f g" + and Q: "\S S'. \P S; homeomorphism S S' f g\ \ Q S'" + shows "locally Q t" +proof (clarsimp simp: locally_def) + fix W y + assume "y \ W" and "openin (subtopology euclidean t) W" + then obtain T where T: "open T" "W = t \ T" + by (force simp: openin_open) + then have "W \ t" by auto + have f: "\x. x \ S \ g(f x) = x" "f ` S = t" "continuous_on S f" + and g: "\y. y \ t \ f(g y) = y" "g ` t = S" "continuous_on t g" + using hom by (auto simp: homeomorphism_def) + have gw: "g ` W = S \ f -` W" + using \W \ t\ + apply auto + using \g ` t = S\ \W \ t\ apply blast + using g \W \ t\ apply auto[1] + by (simp add: f rev_image_eqI) + have \: "openin (subtopology euclidean S) (g ` W)" + proof - + have "continuous_on S f" + using f(3) by blast + then show "openin (subtopology euclidean S) (g ` W)" + by (simp add: gw Collect_conj_eq \openin (subtopology euclidean t) W\ continuous_on_open f(2)) + qed + then obtain u v + where osu: "openin (subtopology euclidean S) u" and uv: "P v" "g y \ u" "u \ v" "v \ g ` W" + using S [unfolded locally_def, rule_format, of "g ` W" "g y"] \y \ W\ by force + have "v \ S" using uv by (simp add: gw) + have fv: "f ` v = t \ {x. g x \ v}" + using \f ` S = t\ f \v \ S\ by auto + have "f ` v \ W" + using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto + have contvf: "continuous_on v f" + using \v \ S\ continuous_on_subset f(3) by blast + have contvg: "continuous_on (f ` v) g" + using \f ` v \ W\ \W \ t\ continuous_on_subset [OF g(3)] by blast + have homv: "homeomorphism v (f ` v) f g" + using \v \ S\ \W \ t\ f + apply (simp add: homeomorphism_def contvf contvg, auto) + by (metis f(1) rev_image_eqI rev_subsetD) + have 1: "openin (subtopology euclidean t) (t \ g -` u)" + apply (rule continuous_on_open [THEN iffD1, rule_format]) + apply (rule \continuous_on t g\) + using \g ` t = S\ apply (simp add: osu) + done + have 2: "\V. Q V \ y \ (t \ g -` u) \ (t \ g -` u) \ V \ V \ W" + apply (rule_tac x="f ` v" in exI) + apply (intro conjI Q [OF \P v\ homv]) + using \W \ t\ \y \ W\ \f ` v \ W\ uv apply (auto simp: fv) + done + show "\U. openin (subtopology euclidean t) U \ (\v. Q v \ y \ U \ U \ v \ v \ W)" + by (meson 1 2) +qed + +lemma homeomorphism_locally: + fixes f:: "'a::metric_space \ 'b::metric_space" + assumes hom: "homeomorphism S t f g" + and eq: "\S t. homeomorphism S t f g \ (P S \ Q t)" + shows "locally P S \ locally Q t" +apply (rule iffI) +apply (erule homeomorphism_locally_imp [OF _ hom]) +apply (simp add: eq) +apply (erule homeomorphism_locally_imp) +using eq homeomorphism_sym homeomorphism_symD [OF hom] apply blast+ +done + +lemma homeomorphic_locally: + fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" + assumes hom: "S homeomorphic T" + and iff: "\X Y. X homeomorphic Y \ (P X \ Q Y)" + shows "locally P S \ locally Q T" +proof - + obtain f g where hom: "homeomorphism S T f g" + using assms by (force simp: homeomorphic_def) + then show ?thesis + using homeomorphic_def local.iff + by (blast intro!: homeomorphism_locally) +qed + +lemma homeomorphic_local_compactness: + fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" + shows "S homeomorphic T \ locally compact S \ locally compact T" +by (simp add: homeomorphic_compactness homeomorphic_locally) + +lemma locally_translation: + fixes P :: "'a :: real_normed_vector set \ bool" + shows + "(\S. P (image (\x. a + x) S) \ P S) + \ locally P (image (\x. a + x) S) \ locally P S" +apply (rule homeomorphism_locally [OF homeomorphism_translation]) +apply (simp add: homeomorphism_def) +by metis + +lemma locally_injective_linear_image: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes f: "linear f" "inj f" and iff: "\S. P (f ` S) \ Q S" + shows "locally P (f ` S) \ locally Q S" +apply (rule linear_homeomorphism_image [OF f]) +apply (rule_tac f=g and g = f in homeomorphism_locally, assumption) +by (metis iff homeomorphism_def) + +lemma locally_open_map_image: + fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes P: "locally P S" + and f: "continuous_on S f" + and oo: "\t. openin (subtopology euclidean S) t + \ openin (subtopology euclidean (f ` S)) (f ` t)" + and Q: "\t. \t \ S; P t\ \ Q(f ` t)" + shows "locally Q (f ` S)" +proof (clarsimp simp add: locally_def) + fix W y + assume oiw: "openin (subtopology euclidean (f ` S)) W" and "y \ W" + then have "W \ f ` S" by (simp add: openin_euclidean_subtopology_iff) + have oivf: "openin (subtopology euclidean S) (S \ f -` W)" + by (rule continuous_on_open [THEN iffD1, rule_format, OF f oiw]) + then obtain x where "x \ S" "f x = y" + using \W \ f ` S\ \y \ W\ by blast + then obtain U V + where "openin (subtopology euclidean S) U" "P V" "x \ U" "U \ V" "V \ S \ f -` W" + using P [unfolded locally_def, rule_format, of "(S \ f -` W)" x] oivf \y \ W\ + by auto + then show "\X. openin (subtopology euclidean (f ` S)) X \ (\Y. Q Y \ y \ X \ X \ Y \ Y \ W)" + apply (rule_tac x="f ` U" in exI) + apply (rule conjI, blast intro!: oo) + apply (rule_tac x="f ` V" in exI) + apply (force simp: \f x = y\ rev_image_eqI intro: Q) + done +qed + + +subsection\An induction principle for connected sets\ + +proposition connected_induction: + assumes "connected S" + and opD: "\T a. \openin (subtopology euclidean S) T; a \ T\ \ \z. z \ T \ P z" + and opI: "\a. a \ S + \ \T. openin (subtopology euclidean S) T \ a \ T \ + (\x \ T. \y \ T. P x \ P y \ Q x \ Q y)" + and etc: "a \ S" "b \ S" "P a" "P b" "Q a" + shows "Q b" +proof - + have 1: "openin (subtopology euclidean S) + {b. \T. openin (subtopology euclidean S) T \ + b \ T \ (\x\T. P x \ Q x)}" + apply (subst openin_subopen, clarify) + apply (rule_tac x=T in exI, auto) + done + have 2: "openin (subtopology euclidean S) + {b. \T. openin (subtopology euclidean S) T \ + b \ T \ (\x\T. P x \ \ Q x)}" + apply (subst openin_subopen, clarify) + apply (rule_tac x=T in exI, auto) + done + show ?thesis + using \connected S\ + apply (simp only: connected_openin HOL.not_ex HOL.de_Morgan_conj) + apply (elim disjE allE) + apply (blast intro: 1) + apply (blast intro: 2, simp_all) + apply clarify apply (metis opI) + using opD apply (blast intro: etc elim: dest:) + using opI etc apply meson+ + done +qed + +lemma connected_equivalence_relation_gen: + assumes "connected S" + and etc: "a \ S" "b \ S" "P a" "P b" + and trans: "\x y z. \R x y; R y z\ \ R x z" + and opD: "\T a. \openin (subtopology euclidean S) T; a \ T\ \ \z. z \ T \ P z" + and opI: "\a. a \ S + \ \T. openin (subtopology euclidean S) T \ a \ T \ + (\x \ T. \y \ T. P x \ P y \ R x y)" + shows "R a b" +proof - + have "\a b c. \a \ S; P a; b \ S; c \ S; P b; P c; R a b\ \ R a c" + apply (rule connected_induction [OF \connected S\ opD], simp_all) + by (meson trans opI) + then show ?thesis by (metis etc opI) +qed + +lemma connected_induction_simple: + assumes "connected S" + and etc: "a \ S" "b \ S" "P a" + and opI: "\a. a \ S + \ \T. openin (subtopology euclidean S) T \ a \ T \ + (\x \ T. \y \ T. P x \ P y)" + shows "P b" +apply (rule connected_induction [OF \connected S\ _, where P = "\x. True"], blast) +apply (frule opI) +using etc apply simp_all +done + +lemma connected_equivalence_relation: + assumes "connected S" + and etc: "a \ S" "b \ S" + and sym: "\x y. \R x y; x \ S; y \ S\ \ R y x" + and trans: "\x y z. \R x y; R y z; x \ S; y \ S; z \ S\ \ R x z" + and opI: "\a. a \ S \ \T. openin (subtopology euclidean S) T \ a \ T \ (\x \ T. R a x)" + shows "R a b" +proof - + have "\a b c. \a \ S; b \ S; c \ S; R a b\ \ R a c" + apply (rule connected_induction_simple [OF \connected S\], simp_all) + by (meson local.sym local.trans opI openin_imp_subset subsetCE) + then show ?thesis by (metis etc opI) +qed + +lemma locally_constant_imp_constant: + assumes "connected S" + and opI: "\a. a \ S + \ \T. openin (subtopology euclidean S) T \ a \ T \ (\x \ T. f x = f a)" + shows "f constant_on S" +proof - + have "\x y. x \ S \ y \ S \ f x = f y" + apply (rule connected_equivalence_relation [OF \connected S\], simp_all) + by (metis opI) + then show ?thesis + by (metis constant_on_def) +qed + +lemma locally_constant: + "connected S \ locally (\U. f constant_on U) S \ f constant_on S" +apply (simp add: locally_def) +apply (rule iffI) + apply (rule locally_constant_imp_constant, assumption) + apply (metis (mono_tags, hide_lams) constant_on_def constant_on_subset openin_subtopology_self) +by (meson constant_on_subset openin_imp_subset order_refl) + + +subsection\Basic properties of local compactness\ + +proposition locally_compact: + fixes s :: "'a :: metric_space set" + shows + "locally compact s \ + (\x \ s. \u v. x \ u \ u \ v \ v \ s \ + openin (subtopology euclidean s) u \ compact v)" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + apply clarify + apply (erule_tac w = "s \ ball x 1" in locallyE) + by auto +next + assume r [rule_format]: ?rhs + have *: "\u v. + openin (subtopology euclidean s) u \ + compact v \ x \ u \ u \ v \ v \ s \ T" + if "open T" "x \ s" "x \ T" for x T + proof - + obtain u v where uv: "x \ u" "u \ v" "v \ s" "compact v" "openin (subtopology euclidean s) u" + using r [OF \x \ s\] by auto + obtain e where "e>0" and e: "cball x e \ T" + using open_contains_cball \open T\ \x \ T\ by blast + show ?thesis + apply (rule_tac x="(s \ ball x e) \ u" in exI) + apply (rule_tac x="cball x e \ v" in exI) + using that \e > 0\ e uv + apply auto + done + qed + show ?lhs + apply (rule locallyI) + apply (subst (asm) openin_open) + apply (blast intro: *) + done +qed + +lemma locally_compactE: + fixes s :: "'a :: metric_space set" + assumes "locally compact s" + obtains u v where "\x. x \ s \ x \ u x \ u x \ v x \ v x \ s \ + openin (subtopology euclidean s) (u x) \ compact (v x)" +using assms +unfolding locally_compact by metis + +lemma locally_compact_alt: + fixes s :: "'a :: heine_borel set" + shows "locally compact s \ + (\x \ s. \u. x \ u \ + openin (subtopology euclidean s) u \ compact(closure u) \ closure u \ s)" +apply (simp add: locally_compact) +apply (intro ball_cong ex_cong refl iffI) +apply (metis bounded_subset closure_eq closure_mono compact_eq_bounded_closed dual_order.trans) +by (meson closure_subset compact_closure) + +lemma locally_compact_Int_cball: + fixes s :: "'a :: heine_borel set" + shows "locally compact s \ (\x \ s. \e. 0 < e \ closed(cball x e \ s))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + apply (simp add: locally_compact openin_contains_cball) + apply (clarify | assumption | drule bspec)+ + by (metis (no_types, lifting) compact_cball compact_imp_closed compact_Int inf.absorb_iff2 inf.orderE inf_sup_aci(2)) +next + assume ?rhs + then show ?lhs + apply (simp add: locally_compact openin_contains_cball) + apply (clarify | assumption | drule bspec)+ + apply (rule_tac x="ball x e \ s" in exI, simp) + apply (rule_tac x="cball x e \ s" in exI) + using compact_eq_bounded_closed + apply auto + apply (metis open_ball le_infI1 mem_ball open_contains_cball_eq) + done +qed + +lemma locally_compact_compact: + fixes s :: "'a :: heine_borel set" + shows "locally compact s \ + (\k. k \ s \ compact k + \ (\u v. k \ u \ u \ v \ v \ s \ + openin (subtopology euclidean s) u \ compact v))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then obtain u v where + uv: "\x. x \ s \ x \ u x \ u x \ v x \ v x \ s \ + openin (subtopology euclidean s) (u x) \ compact (v x)" + by (metis locally_compactE) + have *: "\u v. k \ u \ u \ v \ v \ s \ openin (subtopology euclidean s) u \ compact v" + if "k \ s" "compact k" for k + proof - + have "\C. (\c\C. openin (subtopology euclidean k) c) \ k \ \C \ + \D\C. finite D \ k \ \D" + using that by (simp add: compact_eq_openin_cover) + moreover have "\c \ (\x. k \ u x) ` k. openin (subtopology euclidean k) c" + using that by clarify (metis subsetD inf.absorb_iff2 openin_subset openin_subtopology_Int_subset topspace_euclidean_subtopology uv) + moreover have "k \ \((\x. k \ u x) ` k)" + using that by clarsimp (meson subsetCE uv) + ultimately obtain D where "D \ (\x. k \ u x) ` k" "finite D" "k \ \D" + by metis + then obtain T where T: "T \ k" "finite T" "k \ \((\x. k \ u x) ` T)" + by (metis finite_subset_image) + have Tuv: "\(u ` T) \ \(v ` T)" + using T that by (force simp: dest!: uv) + show ?thesis + apply (rule_tac x="\(u ` T)" in exI) + apply (rule_tac x="\(v ` T)" in exI) + apply (simp add: Tuv) + using T that + apply (auto simp: dest!: uv) + done + qed + show ?rhs + by (blast intro: *) +next + assume ?rhs + then show ?lhs + apply (clarsimp simp add: locally_compact) + apply (drule_tac x="{x}" in spec, simp) + done +qed + +lemma open_imp_locally_compact: + fixes s :: "'a :: heine_borel set" + assumes "open s" + shows "locally compact s" +proof - + have *: "\u v. x \ u \ u \ v \ v \ s \ openin (subtopology euclidean s) u \ compact v" + if "x \ s" for x + proof - + obtain e where "e>0" and e: "cball x e \ s" + using open_contains_cball assms \x \ s\ by blast + have ope: "openin (subtopology euclidean s) (ball x e)" + by (meson e open_ball ball_subset_cball dual_order.trans open_subset) + show ?thesis + apply (rule_tac x="ball x e" in exI) + apply (rule_tac x="cball x e" in exI) + using \e > 0\ e apply (auto simp: ope) + done + qed + show ?thesis + unfolding locally_compact + by (blast intro: *) +qed + +lemma closed_imp_locally_compact: + fixes s :: "'a :: heine_borel set" + assumes "closed s" + shows "locally compact s" +proof - + have *: "\u v. x \ u \ u \ v \ v \ s \ + openin (subtopology euclidean s) u \ compact v" + if "x \ s" for x + proof - + show ?thesis + apply (rule_tac x = "s \ ball x 1" in exI) + apply (rule_tac x = "s \ cball x 1" in exI) + using \x \ s\ assms apply auto + done + qed + show ?thesis + unfolding locally_compact + by (blast intro: *) +qed + +lemma locally_compact_UNIV: "locally compact (UNIV :: 'a :: heine_borel set)" + by (simp add: closed_imp_locally_compact) + +lemma locally_compact_Int: + fixes s :: "'a :: t2_space set" + shows "\locally compact s; locally compact t\ \ locally compact (s \ t)" +by (simp add: compact_Int locally_Int) + +lemma locally_compact_closedin: + fixes s :: "'a :: heine_borel set" + shows "\closedin (subtopology euclidean s) t; locally compact s\ + \ locally compact t" +unfolding closedin_closed +using closed_imp_locally_compact locally_compact_Int by blast + +lemma locally_compact_delete: + fixes s :: "'a :: t1_space set" + shows "locally compact s \ locally compact (s - {a})" + by (auto simp: openin_delete locally_open_subset) + +lemma locally_closed: + fixes s :: "'a :: heine_borel set" + shows "locally closed s \ locally compact s" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + apply (simp only: locally_def) + apply (erule all_forward imp_forward asm_rl exE)+ + apply (rule_tac x = "u \ ball x 1" in exI) + apply (rule_tac x = "v \ cball x 1" in exI) + apply (force intro: openin_trans) + done +next + assume ?rhs then show ?lhs + using compact_eq_bounded_closed locally_mono by blast +qed + +lemma locally_compact_openin_Un: + fixes S :: "'a::euclidean_space set" + assumes LCS: "locally compact S" and LCT:"locally compact T" + and opS: "openin (subtopology euclidean (S \ T)) S" + and opT: "openin (subtopology euclidean (S \ T)) T" + shows "locally compact (S \ T)" +proof - + have "\e>0. closed (cball x e \ (S \ T))" if "x \ S" for x + proof - + obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" + using LCS \x \ S\ unfolding locally_compact_Int_cball by blast + moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \ (S \ T) \ S" + by (meson \x \ S\ opS openin_contains_cball) + then have "cball x e2 \ (S \ T) = cball x e2 \ S" + by force + ultimately show ?thesis + apply (rule_tac x="min e1 e2" in exI) + apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int) + by (metis closed_Int closed_cball inf_left_commute) + qed + moreover have "\e>0. closed (cball x e \ (S \ T))" if "x \ T" for x + proof - + obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ T)" + using LCT \x \ T\ unfolding locally_compact_Int_cball by blast + moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \ (S \ T) \ T" + by (meson \x \ T\ opT openin_contains_cball) + then have "cball x e2 \ (S \ T) = cball x e2 \ T" + by force + ultimately show ?thesis + apply (rule_tac x="min e1 e2" in exI) + apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int) + by (metis closed_Int closed_cball inf_left_commute) + qed + ultimately show ?thesis + by (force simp: locally_compact_Int_cball) +qed + +lemma locally_compact_closedin_Un: + fixes S :: "'a::euclidean_space set" + assumes LCS: "locally compact S" and LCT:"locally compact T" + and clS: "closedin (subtopology euclidean (S \ T)) S" + and clT: "closedin (subtopology euclidean (S \ T)) T" + shows "locally compact (S \ T)" +proof - + have "\e>0. closed (cball x e \ (S \ T))" if "x \ S" "x \ T" for x + proof - + obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" + using LCS \x \ S\ unfolding locally_compact_Int_cball by blast + moreover + obtain e2 where "e2 > 0" and e2: "closed (cball x e2 \ T)" + using LCT \x \ T\ unfolding locally_compact_Int_cball by blast + ultimately show ?thesis + apply (rule_tac x="min e1 e2" in exI) + apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int Int_Un_distrib) + by (metis closed_Int closed_Un closed_cball inf_left_commute) + qed + moreover + have "\e>0. closed (cball x e \ (S \ T))" if x: "x \ S" "x \ T" for x + proof - + obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" + using LCS \x \ S\ unfolding locally_compact_Int_cball by blast + moreover + obtain e2 where "e2>0" and "cball x e2 \ (S \ T) \ S - T" + using clT x by (fastforce simp: openin_contains_cball closedin_def) + then have "closed (cball x e2 \ T)" + proof - + have "{} = T - (T - cball x e2)" + using Diff_subset Int_Diff \cball x e2 \ (S \ T) \ S - T\ by auto + then show ?thesis + by (simp add: Diff_Diff_Int inf_commute) + qed + ultimately show ?thesis + apply (rule_tac x="min e1 e2" in exI) + apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int Int_Un_distrib) + by (metis closed_Int closed_Un closed_cball inf_left_commute) + qed + moreover + have "\e>0. closed (cball x e \ (S \ T))" if x: "x \ S" "x \ T" for x + proof - + obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ T)" + using LCT \x \ T\ unfolding locally_compact_Int_cball by blast + moreover + obtain e2 where "e2>0" and "cball x e2 \ (S \ T) \ S \ T - S" + using clS x by (fastforce simp: openin_contains_cball closedin_def) + then have "closed (cball x e2 \ S)" + by (metis Diff_disjoint Int_empty_right closed_empty inf.left_commute inf.orderE inf_sup_absorb) + ultimately show ?thesis + apply (rule_tac x="min e1 e2" in exI) + apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int Int_Un_distrib) + by (metis closed_Int closed_Un closed_cball inf_left_commute) + qed + ultimately show ?thesis + by (auto simp: locally_compact_Int_cball) +qed + +lemma locally_compact_Times: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + shows "\locally compact S; locally compact T\ \ locally compact (S \ T)" + by (auto simp: compact_Times locally_Times) + +lemma locally_compact_compact_subopen: + fixes S :: "'a :: heine_borel set" + shows + "locally compact S \ + (\K T. K \ S \ compact K \ open T \ K \ T + \ (\U V. K \ U \ U \ V \ U \ T \ V \ S \ + openin (subtopology euclidean S) U \ compact V))" + (is "?lhs = ?rhs") +proof + assume L: ?lhs + show ?rhs + proof clarify + fix K :: "'a set" and T :: "'a set" + assume "K \ S" and "compact K" and "open T" and "K \ T" + obtain U V where "K \ U" "U \ V" "V \ S" "compact V" + and ope: "openin (subtopology euclidean S) U" + using L unfolding locally_compact_compact by (meson \K \ S\ \compact K\) + show "\U V. K \ U \ U \ V \ U \ T \ V \ S \ + openin (subtopology euclidean S) U \ compact V" + proof (intro exI conjI) + show "K \ U \ T" + by (simp add: \K \ T\ \K \ U\) + show "U \ T \ closure(U \ T)" + by (rule closure_subset) + show "closure (U \ T) \ S" + by (metis \U \ V\ \V \ S\ \compact V\ closure_closed closure_mono compact_imp_closed inf.cobounded1 subset_trans) + show "openin (subtopology euclidean S) (U \ T)" + by (simp add: \open T\ ope openin_Int_open) + show "compact (closure (U \ T))" + by (meson Int_lower1 \U \ V\ \compact V\ bounded_subset compact_closure compact_eq_bounded_closed) + qed auto + qed +next + assume ?rhs then show ?lhs + unfolding locally_compact_compact + by (metis open_openin openin_topspace subtopology_superset top.extremum topspace_euclidean_subtopology) +qed + + +subsection\Sura-Bura's results about compact components of sets\ + +proposition Sura_Bura_compact: + fixes S :: "'a::euclidean_space set" + assumes "compact S" and C: "C \ components S" + shows "C = \{T. C \ T \ openin (subtopology euclidean S) T \ + closedin (subtopology euclidean S) T}" + (is "C = \?\") +proof + obtain x where x: "C = connected_component_set S x" and "x \ S" + using C by (auto simp: components_def) + have "C \ S" + by (simp add: C in_components_subset) + have "\?\ \ connected_component_set S x" + proof (rule connected_component_maximal) + have "x \ C" + by (simp add: \x \ S\ x) + then show "x \ \?\" + by blast + have clo: "closed (\?\)" + by (simp add: \compact S\ closed_Inter closedin_compact_eq compact_imp_closed) + have False + if K1: "closedin (subtopology euclidean (\?\)) K1" and + K2: "closedin (subtopology euclidean (\?\)) K2" and + K12_Int: "K1 \ K2 = {}" and K12_Un: "K1 \ K2 = \?\" and "K1 \ {}" "K2 \ {}" + for K1 K2 + proof - + have "closed K1" "closed K2" + using closedin_closed_trans clo K1 K2 by blast+ + then obtain V1 V2 where "open V1" "open V2" "K1 \ V1" "K2 \ V2" and V12: "V1 \ V2 = {}" + using separation_normal \K1 \ K2 = {}\ by metis + have SV12_ne: "(S - (V1 \ V2)) \ (\?\) \ {}" + proof (rule compact_imp_fip) + show "compact (S - (V1 \ V2))" + by (simp add: \open V1\ \open V2\ \compact S\ compact_diff open_Un) + show clo\: "closed T" if "T \ ?\" for T + using that \compact S\ + by (force intro: closedin_closed_trans simp add: compact_imp_closed) + show "(S - (V1 \ V2)) \ \\ \ {}" if "finite \" and \: "\ \ ?\" for \ + proof + assume djo: "(S - (V1 \ V2)) \ \\ = {}" + obtain D where opeD: "openin (subtopology euclidean S) D" + and cloD: "closedin (subtopology euclidean S) D" + and "C \ D" and DV12: "D \ V1 \ V2" + proof (cases "\ = {}") + case True + with \C \ S\ djo that show ?thesis + by force + next + case False show ?thesis + proof + show ope: "openin (subtopology euclidean S) (\\)" + using openin_Inter \finite \\ False \ by blast + then show "closedin (subtopology euclidean S) (\\)" + by (meson clo\ \ closed_Inter closed_subset openin_imp_subset subset_eq) + show "C \ \\" + using \ by auto + show "\\ \ V1 \ V2" + using ope djo openin_imp_subset by fastforce + qed + qed + have "connected C" + by (simp add: x) + have "closed D" + using \compact S\ cloD closedin_closed_trans compact_imp_closed by blast + have cloV1: "closedin (subtopology euclidean D) (D \ closure V1)" + and cloV2: "closedin (subtopology euclidean D) (D \ closure V2)" + by (simp_all add: closedin_closed_Int) + moreover have "D \ closure V1 = D \ V1" "D \ closure V2 = D \ V2" + apply safe + using \D \ V1 \ V2\ \open V1\ \open V2\ V12 + apply (simp_all add: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+) + done + ultimately have cloDV1: "closedin (subtopology euclidean D) (D \ V1)" + and cloDV2: "closedin (subtopology euclidean D) (D \ V2)" + by metis+ + then obtain U1 U2 where "closed U1" "closed U2" + and D1: "D \ V1 = D \ U1" and D2: "D \ V2 = D \ U2" + by (auto simp: closedin_closed) + have "D \ U1 \ C \ {}" + proof + assume "D \ U1 \ C = {}" + then have *: "C \ D \ V2" + using D1 DV12 \C \ D\ by auto + have "\?\ \ D \ V2" + apply (rule Inter_lower) + using * apply simp + by (meson cloDV2 \open V2\ cloD closedin_trans le_inf_iff opeD openin_Int_open) + then show False + using K1 V12 \K1 \ {}\ \K1 \ V1\ closedin_imp_subset by blast + qed + moreover have "D \ U2 \ C \ {}" + proof + assume "D \ U2 \ C = {}" + then have *: "C \ D \ V1" + using D2 DV12 \C \ D\ by auto + have "\?\ \ D \ V1" + apply (rule Inter_lower) + using * apply simp + by (meson cloDV1 \open V1\ cloD closedin_trans le_inf_iff opeD openin_Int_open) + then show False + using K2 V12 \K2 \ {}\ \K2 \ V2\ closedin_imp_subset by blast + qed + ultimately show False + using \connected C\ unfolding connected_closed + apply (simp only: not_ex) + apply (drule_tac x="D \ U1" in spec) + apply (drule_tac x="D \ U2" in spec) + using \C \ D\ D1 D2 V12 DV12 \closed U1\ \closed U2\ \closed D\ + by blast + qed + qed + show False + by (metis (full_types) DiffE UnE Un_upper2 SV12_ne \K1 \ V1\ \K2 \ V2\ disjoint_iff_not_equal subsetCE sup_ge1 K12_Un) + qed + then show "connected (\?\)" + by (auto simp: connected_closedin_eq) + show "\?\ \ S" + by (fastforce simp: C in_components_subset) + qed + with x show "\?\ \ C" by simp +qed auto + + +corollary Sura_Bura_clopen_subset: + fixes S :: "'a::euclidean_space set" + assumes S: "locally compact S" and C: "C \ components S" and "compact C" + and U: "open U" "C \ U" + obtains K where "openin (subtopology euclidean S) K" "compact K" "C \ K" "K \ U" +proof (rule ccontr) + assume "\ thesis" + with that have neg: "\K. openin (subtopology euclidean S) K \ compact K \ C \ K \ K \ U" + by metis + obtain V K where "C \ V" "V \ U" "V \ K" "K \ S" "compact K" + and opeSV: "openin (subtopology euclidean S) V" + using S U \compact C\ + apply (simp add: locally_compact_compact_subopen) + by (meson C in_components_subset) + let ?\ = "{T. C \ T \ openin (subtopology euclidean K) T \ compact T \ T \ K}" + have CK: "C \ components K" + by (meson C \C \ V\ \K \ S\ \V \ K\ components_intermediate_subset subset_trans) + with \compact K\ + have "C = \{T. C \ T \ openin (subtopology euclidean K) T \ closedin (subtopology euclidean K) T}" + by (simp add: Sura_Bura_compact) + then have Ceq: "C = \?\" + by (simp add: closedin_compact_eq \compact K\) + obtain W where "open W" and W: "V = S \ W" + using opeSV by (auto simp: openin_open) + have "-(U \ W) \ \?\ \ {}" + proof (rule closed_imp_fip_compact) + show "- (U \ W) \ \\ \ {}" + if "finite \" and \: "\ \ ?\" for \ + proof (cases "\ = {}") + case True + have False if "U = UNIV" "W = UNIV" + proof - + have "V = S" + by (simp add: W \W = UNIV\) + with neg show False + using \C \ V\ \K \ S\ \V \ K\ \V \ U\ \compact K\ by auto + qed + with True show ?thesis + by auto + next + case False + show ?thesis + proof + assume "- (U \ W) \ \\ = {}" + then have FUW: "\\ \ U \ W" + by blast + have "C \ \\" + using \ by auto + moreover have "compact (\\)" + by (metis (no_types, lifting) compact_Inter False mem_Collect_eq subsetCE \) + moreover have "\\ \ K" + using False that(2) by fastforce + moreover have opeKF: "openin (subtopology euclidean K) (\\)" + using False \ \finite \\ by blast + then have opeVF: "openin (subtopology euclidean V) (\\)" + using W \K \ S\ \V \ K\ opeKF \\\ \ K\ FUW openin_subset_trans by fastforce + then have "openin (subtopology euclidean S) (\\)" + by (metis opeSV openin_trans) + moreover have "\\ \ U" + by (meson \V \ U\ opeVF dual_order.trans openin_imp_subset) + ultimately show False + using neg by blast + qed + qed + qed (use \open W\ \open U\ in auto) + with W Ceq \C \ V\ \C \ U\ show False + by auto +qed + + +corollary Sura_Bura_clopen_subset_alt: + fixes S :: "'a::euclidean_space set" + assumes S: "locally compact S" and C: "C \ components S" and "compact C" + and opeSU: "openin (subtopology euclidean S) U" and "C \ U" + obtains K where "openin (subtopology euclidean S) K" "compact K" "C \ K" "K \ U" +proof - + obtain V where "open V" "U = S \ V" + using opeSU by (auto simp: openin_open) + with \C \ U\ have "C \ V" + by auto + then show ?thesis + using Sura_Bura_clopen_subset [OF S C \compact C\ \open V\] + by (metis \U = S \ V\ inf.bounded_iff openin_imp_subset that) +qed + +corollary Sura_Bura: + fixes S :: "'a::euclidean_space set" + assumes "locally compact S" "C \ components S" "compact C" + shows "C = \ {K. C \ K \ compact K \ openin (subtopology euclidean S) K}" + (is "C = ?rhs") +proof + show "?rhs \ C" + proof (clarsimp, rule ccontr) + fix x + assume *: "\X. C \ X \ compact X \ openin (subtopology euclidean S) X \ x \ X" + and "x \ C" + obtain U V where "open U" "open V" "{x} \ U" "C \ V" "U \ V = {}" + using separation_normal [of "{x}" C] + by (metis Int_empty_left \x \ C\ \compact C\ closed_empty closed_insert compact_imp_closed insert_disjoint(1)) + have "x \ V" + using \U \ V = {}\ \{x} \ U\ by blast + then show False + by (meson "*" Sura_Bura_clopen_subset \C \ V\ \open V\ assms(1) assms(2) assms(3) subsetCE) + qed +qed blast + + +subsection\Special cases of local connectedness and path connectedness\ + +lemma locally_connected_1: + assumes + "\v x. \openin (subtopology euclidean S) v; x \ v\ + \ \u. openin (subtopology euclidean S) u \ + connected u \ x \ u \ u \ v" + shows "locally connected S" +apply (clarsimp simp add: locally_def) +apply (drule assms; blast) +done + +lemma locally_connected_2: + assumes "locally connected S" + "openin (subtopology euclidean S) t" + "x \ t" + shows "openin (subtopology euclidean S) (connected_component_set t x)" +proof - + { fix y :: 'a + let ?SS = "subtopology euclidean S" + assume 1: "openin ?SS t" + "\w x. openin ?SS w \ x \ w \ (\u. openin ?SS u \ (\v. connected v \ x \ u \ u \ v \ v \ w))" + and "connected_component t x y" + then have "y \ t" and y: "y \ connected_component_set t x" + using connected_component_subset by blast+ + obtain F where + "\x y. (\w. openin ?SS w \ (\u. connected u \ x \ w \ w \ u \ u \ y)) = (openin ?SS (F x y) \ (\u. connected u \ x \ F x y \ F x y \ u \ u \ y))" + by moura + then obtain G where + "\a A. (\U. openin ?SS U \ (\V. connected V \ a \ U \ U \ V \ V \ A)) = (openin ?SS (F a A) \ connected (G a A) \ a \ F a A \ F a A \ G a A \ G a A \ A)" + by moura + then have *: "openin ?SS (F y t) \ connected (G y t) \ y \ F y t \ F y t \ G y t \ G y t \ t" + using 1 \y \ t\ by presburger + have "G y t \ connected_component_set t y" + by (metis (no_types) * connected_component_eq_self connected_component_mono contra_subsetD) + then have "\A. openin ?SS A \ y \ A \ A \ connected_component_set t x" + by (metis (no_types) * connected_component_eq dual_order.trans y) + } + then show ?thesis + using assms openin_subopen by (force simp: locally_def) +qed + +lemma locally_connected_3: + assumes "\t x. \openin (subtopology euclidean S) t; x \ t\ + \ openin (subtopology euclidean S) + (connected_component_set t x)" + "openin (subtopology euclidean S) v" "x \ v" + shows "\u. openin (subtopology euclidean S) u \ connected u \ x \ u \ u \ v" +using assms connected_component_subset by fastforce + +lemma locally_connected: + "locally connected S \ + (\v x. openin (subtopology euclidean S) v \ x \ v + \ (\u. openin (subtopology euclidean S) u \ connected u \ x \ u \ u \ v))" +by (metis locally_connected_1 locally_connected_2 locally_connected_3) + +lemma locally_connected_open_connected_component: + "locally connected S \ + (\t x. openin (subtopology euclidean S) t \ x \ t + \ openin (subtopology euclidean S) (connected_component_set t x))" +by (metis locally_connected_1 locally_connected_2 locally_connected_3) + +lemma locally_path_connected_1: + assumes + "\v x. \openin (subtopology euclidean S) v; x \ v\ + \ \u. openin (subtopology euclidean S) u \ path_connected u \ x \ u \ u \ v" + shows "locally path_connected S" +apply (clarsimp simp add: locally_def) +apply (drule assms; blast) +done + +lemma locally_path_connected_2: + assumes "locally path_connected S" + "openin (subtopology euclidean S) t" + "x \ t" + shows "openin (subtopology euclidean S) (path_component_set t x)" +proof - + { fix y :: 'a + let ?SS = "subtopology euclidean S" + assume 1: "openin ?SS t" + "\w x. openin ?SS w \ x \ w \ (\u. openin ?SS u \ (\v. path_connected v \ x \ u \ u \ v \ v \ w))" + and "path_component t x y" + then have "y \ t" and y: "y \ path_component_set t x" + using path_component_mem(2) by blast+ + obtain F where + "\x y. (\w. openin ?SS w \ (\u. path_connected u \ x \ w \ w \ u \ u \ y)) = (openin ?SS (F x y) \ (\u. path_connected u \ x \ F x y \ F x y \ u \ u \ y))" + by moura + then obtain G where + "\a A. (\U. openin ?SS U \ (\V. path_connected V \ a \ U \ U \ V \ V \ A)) = (openin ?SS (F a A) \ path_connected (G a A) \ a \ F a A \ F a A \ G a A \ G a A \ A)" + by moura + then have *: "openin ?SS (F y t) \ path_connected (G y t) \ y \ F y t \ F y t \ G y t \ G y t \ t" + using 1 \y \ t\ by presburger + have "G y t \ path_component_set t y" + using * path_component_maximal set_rev_mp by blast + then have "\A. openin ?SS A \ y \ A \ A \ path_component_set t x" + by (metis "*" \G y t \ path_component_set t y\ dual_order.trans path_component_eq y) + } + then show ?thesis + using assms openin_subopen by (force simp: locally_def) +qed + +lemma locally_path_connected_3: + assumes "\t x. \openin (subtopology euclidean S) t; x \ t\ + \ openin (subtopology euclidean S) (path_component_set t x)" + "openin (subtopology euclidean S) v" "x \ v" + shows "\u. openin (subtopology euclidean S) u \ path_connected u \ x \ u \ u \ v" +proof - + have "path_component v x x" + by (meson assms(3) path_component_refl) + then show ?thesis + by (metis assms(1) assms(2) assms(3) mem_Collect_eq path_component_subset path_connected_path_component) +qed + +proposition locally_path_connected: + "locally path_connected S \ + (\v x. openin (subtopology euclidean S) v \ x \ v + \ (\u. openin (subtopology euclidean S) u \ path_connected u \ x \ u \ u \ v))" + by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) + +proposition locally_path_connected_open_path_component: + "locally path_connected S \ + (\t x. openin (subtopology euclidean S) t \ x \ t + \ openin (subtopology euclidean S) (path_component_set t x))" + by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) + +lemma locally_connected_open_component: + "locally connected S \ + (\t c. openin (subtopology euclidean S) t \ c \ components t + \ openin (subtopology euclidean S) c)" +by (metis components_iff locally_connected_open_connected_component) + +proposition locally_connected_im_kleinen: + "locally connected S \ + (\v x. openin (subtopology euclidean S) v \ x \ v + \ (\u. openin (subtopology euclidean S) u \ + x \ u \ u \ v \ + (\y. y \ u \ (\c. connected c \ c \ v \ x \ c \ y \ c))))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + by (fastforce simp add: locally_connected) +next + assume ?rhs + have *: "\T. openin (subtopology euclidean S) T \ x \ T \ T \ c" + if "openin (subtopology euclidean S) t" and c: "c \ components t" and "x \ c" for t c x + proof - + from that \?rhs\ [rule_format, of t x] + obtain u where u: + "openin (subtopology euclidean S) u \ x \ u \ u \ t \ + (\y. y \ u \ (\c. connected c \ c \ t \ x \ c \ y \ c))" + using in_components_subset by auto + obtain F :: "'a set \ 'a set \ 'a" where + "\x y. (\z. z \ x \ y = connected_component_set x z) = (F x y \ x \ y = connected_component_set x (F x y))" + by moura + then have F: "F t c \ t \ c = connected_component_set t (F t c)" + by (meson components_iff c) + obtain G :: "'a set \ 'a set \ 'a" where + G: "\x y. (\z. z \ y \ z \ x) = (G x y \ y \ G x y \ x)" + by moura + have "G c u \ u \ G c u \ c" + using F by (metis (full_types) u connected_componentI connected_component_eq mem_Collect_eq that(3)) + then show ?thesis + using G u by auto + qed + show ?lhs + apply (clarsimp simp add: locally_connected_open_component) + apply (subst openin_subopen) + apply (blast intro: *) + done +qed + +proposition locally_path_connected_im_kleinen: + "locally path_connected S \ + (\v x. openin (subtopology euclidean S) v \ x \ v + \ (\u. openin (subtopology euclidean S) u \ + x \ u \ u \ v \ + (\y. y \ u \ (\p. path p \ path_image p \ v \ + pathstart p = x \ pathfinish p = y))))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + apply (simp add: locally_path_connected path_connected_def) + apply (erule all_forward ex_forward imp_forward conjE | simp)+ + by (meson dual_order.trans) +next + assume ?rhs + have *: "\T. openin (subtopology euclidean S) T \ + x \ T \ T \ path_component_set u z" + if "openin (subtopology euclidean S) u" and "z \ u" and c: "path_component u z x" for u z x + proof - + have "x \ u" + by (meson c path_component_mem(2)) + with that \?rhs\ [rule_format, of u x] + obtain U where U: + "openin (subtopology euclidean S) U \ x \ U \ U \ u \ + (\y. y \ U \ (\p. path p \ path_image p \ u \ pathstart p = x \ pathfinish p = y))" + by blast + show ?thesis + apply (rule_tac x=U in exI) + apply (auto simp: U) + apply (metis U c path_component_trans path_component_def) + done + qed + show ?lhs + apply (clarsimp simp add: locally_path_connected_open_path_component) + apply (subst openin_subopen) + apply (blast intro: *) + done +qed + +lemma locally_path_connected_imp_locally_connected: + "locally path_connected S \ locally connected S" +using locally_mono path_connected_imp_connected by blast + +lemma locally_connected_components: + "\locally connected S; c \ components S\ \ locally connected c" +by (meson locally_connected_open_component locally_open_subset openin_subtopology_self) + +lemma locally_path_connected_components: + "\locally path_connected S; c \ components S\ \ locally path_connected c" +by (meson locally_connected_open_component locally_open_subset locally_path_connected_imp_locally_connected openin_subtopology_self) + +lemma locally_path_connected_connected_component: + "locally path_connected S \ locally path_connected (connected_component_set S x)" +by (metis components_iff connected_component_eq_empty locally_empty locally_path_connected_components) + +lemma open_imp_locally_path_connected: + fixes S :: "'a :: real_normed_vector set" + shows "open S \ locally path_connected S" +apply (rule locally_mono [of convex]) +apply (simp_all add: locally_def openin_open_eq convex_imp_path_connected) +apply (meson open_ball centre_in_ball convex_ball openE order_trans) +done + +lemma open_imp_locally_connected: + fixes S :: "'a :: real_normed_vector set" + shows "open S \ locally connected S" +by (simp add: locally_path_connected_imp_locally_connected open_imp_locally_path_connected) + +lemma locally_path_connected_UNIV: "locally path_connected (UNIV::'a :: real_normed_vector set)" + by (simp add: open_imp_locally_path_connected) + +lemma locally_connected_UNIV: "locally connected (UNIV::'a :: real_normed_vector set)" + by (simp add: open_imp_locally_connected) + +lemma openin_connected_component_locally_connected: + "locally connected S + \ openin (subtopology euclidean S) (connected_component_set S x)" +apply (simp add: locally_connected_open_connected_component) +by (metis connected_component_eq_empty connected_component_subset open_empty open_subset openin_subtopology_self) + +lemma openin_components_locally_connected: + "\locally connected S; c \ components S\ \ openin (subtopology euclidean S) c" + using locally_connected_open_component openin_subtopology_self by blast + +lemma openin_path_component_locally_path_connected: + "locally path_connected S + \ openin (subtopology euclidean S) (path_component_set S x)" +by (metis (no_types) empty_iff locally_path_connected_2 openin_subopen openin_subtopology_self path_component_eq_empty) + +lemma closedin_path_component_locally_path_connected: + "locally path_connected S + \ closedin (subtopology euclidean S) (path_component_set S x)" +apply (simp add: closedin_def path_component_subset complement_path_component_Union) +apply (rule openin_Union) +using openin_path_component_locally_path_connected by auto + +lemma convex_imp_locally_path_connected: + fixes S :: "'a:: real_normed_vector set" + shows "convex S \ locally path_connected S" +apply (clarsimp simp add: locally_path_connected) +apply (subst (asm) openin_open) +apply clarify +apply (erule (1) openE) +apply (rule_tac x = "S \ ball x e" in exI) +apply (force simp: convex_Int convex_imp_path_connected) +done + +lemma convex_imp_locally_connected: + fixes S :: "'a:: real_normed_vector set" + shows "convex S \ locally connected S" + by (simp add: locally_path_connected_imp_locally_connected convex_imp_locally_path_connected) + + +subsection\Relations between components and path components\ + +lemma path_component_eq_connected_component: + assumes "locally path_connected S" + shows "(path_component S x = connected_component S x)" +proof (cases "x \ S") + case True + have "openin (subtopology euclidean (connected_component_set S x)) (path_component_set S x)" + apply (rule openin_subset_trans [of S]) + apply (intro conjI openin_path_component_locally_path_connected [OF assms]) + using path_component_subset_connected_component apply (auto simp: connected_component_subset) + done + moreover have "closedin (subtopology euclidean (connected_component_set S x)) (path_component_set S x)" + apply (rule closedin_subset_trans [of S]) + apply (intro conjI closedin_path_component_locally_path_connected [OF assms]) + using path_component_subset_connected_component apply (auto simp: connected_component_subset) + done + ultimately have *: "path_component_set S x = connected_component_set S x" + by (metis connected_connected_component connected_clopen True path_component_eq_empty) + then show ?thesis + by blast +next + case False then show ?thesis + by (metis Collect_empty_eq_bot connected_component_eq_empty path_component_eq_empty) +qed + +lemma path_component_eq_connected_component_set: + "locally path_connected S \ (path_component_set S x = connected_component_set S x)" +by (simp add: path_component_eq_connected_component) + +lemma locally_path_connected_path_component: + "locally path_connected S \ locally path_connected (path_component_set S x)" +using locally_path_connected_connected_component path_component_eq_connected_component by fastforce + +lemma open_path_connected_component: + fixes S :: "'a :: real_normed_vector set" + shows "open S \ path_component S x = connected_component S x" +by (simp add: path_component_eq_connected_component open_imp_locally_path_connected) + +lemma open_path_connected_component_set: + fixes S :: "'a :: real_normed_vector set" + shows "open S \ path_component_set S x = connected_component_set S x" +by (simp add: open_path_connected_component) + +proposition locally_connected_quotient_image: + assumes lcS: "locally connected S" + and oo: "\T. T \ f ` S + \ openin (subtopology euclidean S) (S \ f -` T) \ + openin (subtopology euclidean (f ` S)) T" + shows "locally connected (f ` S)" +proof (clarsimp simp: locally_connected_open_component) + fix U C + assume opefSU: "openin (subtopology euclidean (f ` S)) U" and "C \ components U" + then have "C \ U" "U \ f ` S" + by (meson in_components_subset openin_imp_subset)+ + then have "openin (subtopology euclidean (f ` S)) C \ + openin (subtopology euclidean S) (S \ f -` C)" + by (auto simp: oo) + moreover have "openin (subtopology euclidean S) (S \ f -` C)" + proof (subst openin_subopen, clarify) + fix x + assume "x \ S" "f x \ C" + show "\T. openin (subtopology euclidean S) T \ x \ T \ T \ (S \ f -` C)" + proof (intro conjI exI) + show "openin (subtopology euclidean S) (connected_component_set (S \ f -` U) x)" + proof (rule ccontr) + assume **: "\ openin (subtopology euclidean S) (connected_component_set (S \ f -` U) x)" + then have "x \ (S \ f -` U)" + using \U \ f ` S\ opefSU lcS locally_connected_2 oo by blast + with ** show False + by (metis (no_types) connected_component_eq_empty empty_iff openin_subopen) + qed + next + show "x \ connected_component_set (S \ f -` U) x" + using \C \ U\ \f x \ C\ \x \ S\ by auto + next + have contf: "continuous_on S f" + by (simp add: continuous_on_open oo openin_imp_subset) + then have "continuous_on (connected_component_set (S \ f -` U) x) f" + apply (rule continuous_on_subset) + using connected_component_subset apply blast + done + then have "connected (f ` connected_component_set (S \ f -` U) x)" + by (rule connected_continuous_image [OF _ connected_connected_component]) + moreover have "f ` connected_component_set (S \ f -` U) x \ U" + using connected_component_in by blast + moreover have "C \ f ` connected_component_set (S \ f -` U) x \ {}" + using \C \ U\ \f x \ C\ \x \ S\ by fastforce + ultimately have fC: "f ` (connected_component_set (S \ f -` U) x) \ C" + by (rule components_maximal [OF \C \ components U\]) + have cUC: "connected_component_set (S \ f -` U) x \ (S \ f -` C)" + using connected_component_subset fC by blast + have "connected_component_set (S \ f -` U) x \ connected_component_set (S \ f -` C) x" + proof - + { assume "x \ connected_component_set (S \ f -` U) x" + then have ?thesis + using cUC connected_component_idemp connected_component_mono by blast } + then show ?thesis + using connected_component_eq_empty by auto + qed + also have "\ \ (S \ f -` C)" + by (rule connected_component_subset) + finally show "connected_component_set (S \ f -` U) x \ (S \ f -` C)" . + qed + qed + ultimately show "openin (subtopology euclidean (f ` S)) C" + by metis +qed + +text\The proof resembles that above but is not identical!\ +proposition locally_path_connected_quotient_image: + assumes lcS: "locally path_connected S" + and oo: "\T. T \ f ` S + \ openin (subtopology euclidean S) (S \ f -` T) \ openin (subtopology euclidean (f ` S)) T" + shows "locally path_connected (f ` S)" +proof (clarsimp simp: locally_path_connected_open_path_component) + fix U y + assume opefSU: "openin (subtopology euclidean (f ` S)) U" and "y \ U" + then have "path_component_set U y \ U" "U \ f ` S" + by (meson path_component_subset openin_imp_subset)+ + then have "openin (subtopology euclidean (f ` S)) (path_component_set U y) \ + openin (subtopology euclidean S) (S \ f -` path_component_set U y)" + proof - + have "path_component_set U y \ f ` S" + using \U \ f ` S\ \path_component_set U y \ U\ by blast + then show ?thesis + using oo by blast + qed + moreover have "openin (subtopology euclidean S) (S \ f -` path_component_set U y)" + proof (subst openin_subopen, clarify) + fix x + assume "x \ S" and Uyfx: "path_component U y (f x)" + then have "f x \ U" + using path_component_mem by blast + show "\T. openin (subtopology euclidean S) T \ x \ T \ T \ (S \ f -` path_component_set U y)" + proof (intro conjI exI) + show "openin (subtopology euclidean S) (path_component_set (S \ f -` U) x)" + proof (rule ccontr) + assume **: "\ openin (subtopology euclidean S) (path_component_set (S \ f -` U) x)" + then have "x \ (S \ f -` U)" + by (metis (no_types, lifting) \U \ f ` S\ opefSU lcS oo locally_path_connected_open_path_component) + then show False + using ** \path_component_set U y \ U\ \x \ S\ \path_component U y (f x)\ by blast + qed + next + show "x \ path_component_set (S \ f -` U) x" + by (simp add: \f x \ U\ \x \ S\ path_component_refl) + next + have contf: "continuous_on S f" + by (simp add: continuous_on_open oo openin_imp_subset) + then have "continuous_on (path_component_set (S \ f -` U) x) f" + apply (rule continuous_on_subset) + using path_component_subset apply blast + done + then have "path_connected (f ` path_component_set (S \ f -` U) x)" + by (simp add: path_connected_continuous_image) + moreover have "f ` path_component_set (S \ f -` U) x \ U" + using path_component_mem by fastforce + moreover have "f x \ f ` path_component_set (S \ f -` U) x" + by (force simp: \x \ S\ \f x \ U\ path_component_refl_eq) + ultimately have "f ` (path_component_set (S \ f -` U) x) \ path_component_set U (f x)" + by (meson path_component_maximal) + also have "\ \ path_component_set U y" + by (simp add: Uyfx path_component_maximal path_component_subset path_component_sym) + finally have fC: "f ` (path_component_set (S \ f -` U) x) \ path_component_set U y" . + have cUC: "path_component_set (S \ f -` U) x \ (S \ f -` path_component_set U y)" + using path_component_subset fC by blast + have "path_component_set (S \ f -` U) x \ path_component_set (S \ f -` path_component_set U y) x" + proof - + have "\a. path_component_set (path_component_set (S \ f -` U) x) a \ path_component_set (S \ f -` path_component_set U y) a" + using cUC path_component_mono by blast + then show ?thesis + using path_component_path_component by blast + qed + also have "\ \ (S \ f -` path_component_set U y)" + by (rule path_component_subset) + finally show "path_component_set (S \ f -` U) x \ (S \ f -` path_component_set U y)" . + qed + qed + ultimately show "openin (subtopology euclidean (f ` S)) (path_component_set U y)" + by metis +qed + +subsection%unimportant\Components, continuity, openin, closedin\ + +lemma continuous_on_components_gen: + fixes f :: "'a::topological_space \ 'b::topological_space" + assumes "\c. c \ components S \ + openin (subtopology euclidean S) c \ continuous_on c f" + shows "continuous_on S f" +proof (clarsimp simp: continuous_openin_preimage_eq) + fix t :: "'b set" + assume "open t" + have *: "S \ f -` t = (\c \ components S. c \ f -` t)" + by auto + show "openin (subtopology euclidean S) (S \ f -` t)" + unfolding * using \open t\ assms continuous_openin_preimage_gen openin_trans openin_Union by blast +qed + +lemma continuous_on_components: + fixes f :: "'a::topological_space \ 'b::topological_space" + assumes "locally connected S " + "\c. c \ components S \ continuous_on c f" + shows "continuous_on S f" +apply (rule continuous_on_components_gen) +apply (auto simp: assms intro: openin_components_locally_connected) +done + +lemma continuous_on_components_eq: + "locally connected S + \ (continuous_on S f \ (\c \ components S. continuous_on c f))" +by (meson continuous_on_components continuous_on_subset in_components_subset) + +lemma continuous_on_components_open: + fixes S :: "'a::real_normed_vector set" + assumes "open S " + "\c. c \ components S \ continuous_on c f" + shows "continuous_on S f" +using continuous_on_components open_imp_locally_connected assms by blast + +lemma continuous_on_components_open_eq: + fixes S :: "'a::real_normed_vector set" + shows "open S \ (continuous_on S f \ (\c \ components S. continuous_on c f))" +using continuous_on_subset in_components_subset +by (blast intro: continuous_on_components_open) + +lemma closedin_union_complement_components: + assumes u: "locally connected u" + and S: "closedin (subtopology euclidean u) S" + and cuS: "c \ components(u - S)" + shows "closedin (subtopology euclidean u) (S \ \c)" +proof - + have di: "(\S t. S \ c \ t \ c' \ disjnt S t) \ disjnt (\ c) (\ c')" for c' + by (simp add: disjnt_def) blast + have "S \ u" + using S closedin_imp_subset by blast + moreover have "u - S = \c \ \(components (u - S) - c)" + by (metis Diff_partition Union_components Union_Un_distrib assms(3)) + moreover have "disjnt (\c) (\(components (u - S) - c))" + apply (rule di) + by (metis DiffD1 DiffD2 assms(3) components_nonoverlap disjnt_def subsetCE) + ultimately have eq: "S \ \c = u - (\(components(u - S) - c))" + by (auto simp: disjnt_def) + have *: "openin (subtopology euclidean u) (\(components (u - S) - c))" + apply (rule openin_Union) + apply (rule openin_trans [of "u - S"]) + apply (simp add: u S locally_diff_closed openin_components_locally_connected) + apply (simp add: openin_diff S) + done + have "openin (subtopology euclidean u) (u - (u - \(components (u - S) - c)))" + apply (rule openin_diff, simp) + apply (metis closedin_diff closedin_topspace topspace_euclidean_subtopology *) + done + then show ?thesis + by (force simp: eq closedin_def) +qed + +lemma closed_union_complement_components: + fixes S :: "'a::real_normed_vector set" + assumes S: "closed S" and c: "c \ components(- S)" + shows "closed(S \ \ c)" +proof - + have "closedin (subtopology euclidean UNIV) (S \ \c)" + apply (rule closedin_union_complement_components [OF locally_connected_UNIV]) + using S c apply (simp_all add: Compl_eq_Diff_UNIV) + done + then show ?thesis by simp +qed + +lemma closedin_Un_complement_component: + fixes S :: "'a::real_normed_vector set" + assumes u: "locally connected u" + and S: "closedin (subtopology euclidean u) S" + and c: " c \ components(u - S)" + shows "closedin (subtopology euclidean u) (S \ c)" +proof - + have "closedin (subtopology euclidean u) (S \ \{c})" + using c by (blast intro: closedin_union_complement_components [OF u S]) + then show ?thesis + by simp +qed + +lemma closed_Un_complement_component: + fixes S :: "'a::real_normed_vector set" + assumes S: "closed S" and c: " c \ components(-S)" + shows "closed (S \ c)" + by (metis Compl_eq_Diff_UNIV S c closed_closedin closedin_Un_complement_component + locally_connected_UNIV subtopology_UNIV) + + +subsection\Existence of isometry between subspaces of same dimension\ + +lemma isometry_subset_subspace: + fixes S :: "'a::euclidean_space set" + and T :: "'b::euclidean_space set" + assumes S: "subspace S" + and T: "subspace T" + and d: "dim S \ dim T" + obtains f where "linear f" "f ` S \ T" "\x. x \ S \ norm(f x) = norm x" +proof - + obtain B where "B \ S" and Borth: "pairwise orthogonal B" + and B1: "\x. x \ B \ norm x = 1" + and "independent B" "finite B" "card B = dim S" "span B = S" + by (metis orthonormal_basis_subspace [OF S] independent_finite) + obtain C where "C \ T" and Corth: "pairwise orthogonal C" + and C1:"\x. x \ C \ norm x = 1" + and "independent C" "finite C" "card C = dim T" "span C = T" + by (metis orthonormal_basis_subspace [OF T] independent_finite) + obtain fb where "fb ` B \ C" "inj_on fb B" + by (metis \card B = dim S\ \card C = dim T\ \finite B\ \finite C\ card_le_inj d) + then have pairwise_orth_fb: "pairwise (\v j. orthogonal (fb v) (fb j)) B" + using Corth + apply (auto simp: pairwise_def orthogonal_clauses) + by (meson subsetD image_eqI inj_on_def) + obtain f where "linear f" and ffb: "\x. x \ B \ f x = fb x" + using linear_independent_extend \independent B\ by fastforce + have "span (f ` B) \ span C" + by (metis \fb ` B \ C\ ffb image_cong span_mono) + then have "f ` S \ T" + unfolding \span B = S\ \span C = T\ span_linear_image[OF \linear f\] . + have [simp]: "\x. x \ B \ norm (fb x) = norm x" + using B1 C1 \fb ` B \ C\ by auto + have "norm (f x) = norm x" if "x \ S" for x + proof - + interpret linear f by fact + obtain a where x: "x = (\v \ B. a v *\<^sub>R v)" + using \finite B\ \span B = S\ \x \ S\ span_finite by fastforce + have "norm (f x)^2 = norm (\v\B. a v *\<^sub>R fb v)^2" by (simp add: sum scale ffb x) + also have "\ = (\v\B. norm ((a v *\<^sub>R fb v))^2)" + apply (rule norm_sum_Pythagorean [OF \finite B\]) + apply (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) + done + also have "\ = norm x ^2" + by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \finite B\]) + finally show ?thesis + by (simp add: norm_eq_sqrt_inner) + qed + then show ?thesis + by (rule that [OF \linear f\ \f ` S \ T\]) +qed + +proposition isometries_subspaces: + fixes S :: "'a::euclidean_space set" + and T :: "'b::euclidean_space set" + assumes S: "subspace S" + and T: "subspace T" + and d: "dim S = dim T" + obtains f g where "linear f" "linear g" "f ` S = T" "g ` T = S" + "\x. x \ S \ norm(f x) = norm x" + "\x. x \ T \ norm(g x) = norm x" + "\x. x \ S \ g(f x) = x" + "\x. x \ T \ f(g x) = x" +proof - + obtain B where "B \ S" and Borth: "pairwise orthogonal B" + and B1: "\x. x \ B \ norm x = 1" + and "independent B" "finite B" "card B = dim S" "span B = S" + by (metis orthonormal_basis_subspace [OF S] independent_finite) + obtain C where "C \ T" and Corth: "pairwise orthogonal C" + and C1:"\x. x \ C \ norm x = 1" + and "independent C" "finite C" "card C = dim T" "span C = T" + by (metis orthonormal_basis_subspace [OF T] independent_finite) + obtain fb where "bij_betw fb B C" + by (metis \finite B\ \finite C\ bij_betw_iff_card \card B = dim S\ \card C = dim T\ d) + then have pairwise_orth_fb: "pairwise (\v j. orthogonal (fb v) (fb j)) B" + using Corth + apply (auto simp: pairwise_def orthogonal_clauses bij_betw_def) + by (meson subsetD image_eqI inj_on_def) + obtain f where "linear f" and ffb: "\x. x \ B \ f x = fb x" + using linear_independent_extend \independent B\ by fastforce + interpret f: linear f by fact + define gb where "gb \ inv_into B fb" + then have pairwise_orth_gb: "pairwise (\v j. orthogonal (gb v) (gb j)) C" + using Borth + apply (auto simp: pairwise_def orthogonal_clauses bij_betw_def) + by (metis \bij_betw fb B C\ bij_betw_imp_surj_on bij_betw_inv_into_right inv_into_into) + obtain g where "linear g" and ggb: "\x. x \ C \ g x = gb x" + using linear_independent_extend \independent C\ by fastforce + interpret g: linear g by fact + have "span (f ` B) \ span C" + by (metis \bij_betw fb B C\ bij_betw_imp_surj_on eq_iff ffb image_cong) + then have "f ` S \ T" + unfolding \span B = S\ \span C = T\ + span_linear_image[OF \linear f\] . + have [simp]: "\x. x \ B \ norm (fb x) = norm x" + using B1 C1 \bij_betw fb B C\ bij_betw_imp_surj_on by fastforce + have f [simp]: "norm (f x) = norm x" "g (f x) = x" if "x \ S" for x + proof - + obtain a where x: "x = (\v \ B. a v *\<^sub>R v)" + using \finite B\ \span B = S\ \x \ S\ span_finite by fastforce + have "f x = (\v \ B. f (a v *\<^sub>R v))" + using linear_sum [OF \linear f\] x by auto + also have "\ = (\v \ B. a v *\<^sub>R f v)" + by (simp add: f.sum f.scale) + also have "\ = (\v \ B. a v *\<^sub>R fb v)" + by (simp add: ffb cong: sum.cong) + finally have *: "f x = (\v\B. a v *\<^sub>R fb v)" . + then have "(norm (f x))\<^sup>2 = (norm (\v\B. a v *\<^sub>R fb v))\<^sup>2" by simp + also have "\ = (\v\B. norm ((a v *\<^sub>R fb v))^2)" + apply (rule norm_sum_Pythagorean [OF \finite B\]) + apply (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) + done + also have "\ = (norm x)\<^sup>2" + by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \finite B\]) + finally show "norm (f x) = norm x" + by (simp add: norm_eq_sqrt_inner) + have "g (f x) = g (\v\B. a v *\<^sub>R fb v)" by (simp add: *) + also have "\ = (\v\B. g (a v *\<^sub>R fb v))" + by (simp add: g.sum g.scale) + also have "\ = (\v\B. a v *\<^sub>R g (fb v))" + by (simp add: g.scale) + also have "\ = (\v\B. a v *\<^sub>R v)" + apply (rule sum.cong [OF refl]) + using \bij_betw fb B C\ gb_def bij_betwE bij_betw_inv_into_left gb_def ggb by fastforce + also have "\ = x" + using x by blast + finally show "g (f x) = x" . + qed + have [simp]: "\x. x \ C \ norm (gb x) = norm x" + by (metis B1 C1 \bij_betw fb B C\ bij_betw_imp_surj_on gb_def inv_into_into) + have g [simp]: "f (g x) = x" if "x \ T" for x + proof - + obtain a where x: "x = (\v \ C. a v *\<^sub>R v)" + using \finite C\ \span C = T\ \x \ T\ span_finite by fastforce + have "g x = (\v \ C. g (a v *\<^sub>R v))" + by (simp add: x g.sum) + also have "\ = (\v \ C. a v *\<^sub>R g v)" + by (simp add: g.scale) + also have "\ = (\v \ C. a v *\<^sub>R gb v)" + by (simp add: ggb cong: sum.cong) + finally have "f (g x) = f (\v\C. a v *\<^sub>R gb v)" by simp + also have "\ = (\v\C. f (a v *\<^sub>R gb v))" + by (simp add: f.scale f.sum) + also have "\ = (\v\C. a v *\<^sub>R f (gb v))" + by (simp add: f.scale f.sum) + also have "\ = (\v\C. a v *\<^sub>R v)" + using \bij_betw fb B C\ + by (simp add: bij_betw_def gb_def bij_betw_inv_into_right ffb inv_into_into) + also have "\ = x" + using x by blast + finally show "f (g x) = x" . + qed + have gim: "g ` T = S" + by (metis (full_types) S T \f ` S \ T\ d dim_eq_span dim_image_le f(2) g.linear_axioms + image_iff linear_subspace_image span_eq_iff subset_iff) + have fim: "f ` S = T" + using \g ` T = S\ image_iff by fastforce + have [simp]: "norm (g x) = norm x" if "x \ T" for x + using fim that by auto + show ?thesis + apply (rule that [OF \linear f\ \linear g\]) + apply (simp_all add: fim gim) + done +qed + +corollary isometry_subspaces: + fixes S :: "'a::euclidean_space set" + and T :: "'b::euclidean_space set" + assumes S: "subspace S" + and T: "subspace T" + and d: "dim S = dim T" + obtains f where "linear f" "f ` S = T" "\x. x \ S \ norm(f x) = norm x" +using isometries_subspaces [OF assms] +by metis + +corollary isomorphisms_UNIV_UNIV: + assumes "DIM('M) = DIM('N)" + obtains f::"'M::euclidean_space \'N::euclidean_space" and g + where "linear f" "linear g" + "\x. norm(f x) = norm x" "\y. norm(g y) = norm y" + "\x. g (f x) = x" "\y. f(g y) = y" + using assms by (auto intro: isometries_subspaces [of "UNIV::'M set" "UNIV::'N set"]) + +lemma homeomorphic_subspaces: + fixes S :: "'a::euclidean_space set" + and T :: "'b::euclidean_space set" + assumes S: "subspace S" + and T: "subspace T" + and d: "dim S = dim T" + shows "S homeomorphic T" +proof - + obtain f g where "linear f" "linear g" "f ` S = T" "g ` T = S" + "\x. x \ S \ g(f x) = x" "\x. x \ T \ f(g x) = x" + by (blast intro: isometries_subspaces [OF assms]) + then show ?thesis + apply (simp add: homeomorphic_def homeomorphism_def) + apply (rule_tac x=f in exI) + apply (rule_tac x=g in exI) + apply (auto simp: linear_continuous_on linear_conv_bounded_linear) + done +qed + +lemma homeomorphic_affine_sets: + assumes "affine S" "affine T" "aff_dim S = aff_dim T" + shows "S homeomorphic T" +proof (cases "S = {} \ T = {}") + case True with assms aff_dim_empty homeomorphic_empty show ?thesis + by metis +next + case False + then obtain a b where ab: "a \ S" "b \ T" by auto + then have ss: "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)" + using affine_diffs_subspace assms by blast+ + have dd: "dim ((+) (- a) ` S) = dim ((+) (- b) ` T)" + using assms ab by (simp add: aff_dim_eq_dim [OF hull_inc] image_def) + have "S homeomorphic ((+) (- a) ` S)" + by (simp add: homeomorphic_translation) + also have "\ homeomorphic ((+) (- b) ` T)" + by (rule homeomorphic_subspaces [OF ss dd]) + also have "\ homeomorphic T" + using homeomorphic_sym homeomorphic_translation by auto + finally show ?thesis . +qed + + +subsection\Retracts, in a general sense, preserve (co)homotopic triviality)\ + +locale%important Retracts = + fixes s h t k + assumes conth: "continuous_on s h" + and imh: "h ` s = t" + and contk: "continuous_on t k" + and imk: "k ` t \ s" + and idhk: "\y. y \ t \ h(k y) = y" + +begin + +lemma homotopically_trivial_retraction_gen: + assumes P: "\f. \continuous_on u f; f ` u \ t; Q f\ \ P(k \ f)" + and Q: "\f. \continuous_on u f; f ` u \ s; P f\ \ Q(h \ f)" + and Qeq: "\h k. (\x. x \ u \ h x = k x) \ Q h = Q k" + and hom: "\f g. \continuous_on u f; f ` u \ s; P f; + continuous_on u g; g ` u \ s; P g\ + \ homotopic_with P u s f g" + and contf: "continuous_on u f" and imf: "f ` u \ t" and Qf: "Q f" + and contg: "continuous_on u g" and img: "g ` u \ t" and Qg: "Q g" + shows "homotopic_with Q u t f g" +proof - + have feq: "\x. x \ u \ (h \ (k \ f)) x = f x" using idhk imf by auto + have geq: "\x. x \ u \ (h \ (k \ g)) x = g x" using idhk img by auto + have "continuous_on u (k \ f)" + using contf continuous_on_compose continuous_on_subset contk imf by blast + moreover have "(k \ f) ` u \ s" + using imf imk by fastforce + moreover have "P (k \ f)" + by (simp add: P Qf contf imf) + moreover have "continuous_on u (k \ g)" + using contg continuous_on_compose continuous_on_subset contk img by blast + moreover have "(k \ g) ` u \ s" + using img imk by fastforce + moreover have "P (k \ g)" + by (simp add: P Qg contg img) + ultimately have "homotopic_with P u s (k \ f) (k \ g)" + by (rule hom) + then have "homotopic_with Q u t (h \ (k \ f)) (h \ (k \ g))" + apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) + using Q by (auto simp: conth imh) + then show ?thesis + apply (rule homotopic_with_eq) + apply (metis feq) + apply (metis geq) + apply (metis Qeq) + done +qed + +lemma homotopically_trivial_retraction_null_gen: + assumes P: "\f. \continuous_on u f; f ` u \ t; Q f\ \ P(k \ f)" + and Q: "\f. \continuous_on u f; f ` u \ s; P f\ \ Q(h \ f)" + and Qeq: "\h k. (\x. x \ u \ h x = k x) \ Q h = Q k" + and hom: "\f. \continuous_on u f; f ` u \ s; P f\ + \ \c. homotopic_with P u s f (\x. c)" + and contf: "continuous_on u f" and imf:"f ` u \ t" and Qf: "Q f" + obtains c where "homotopic_with Q u t f (\x. c)" +proof - + have feq: "\x. x \ u \ (h \ (k \ f)) x = f x" using idhk imf by auto + have "continuous_on u (k \ f)" + using contf continuous_on_compose continuous_on_subset contk imf by blast + moreover have "(k \ f) ` u \ s" + using imf imk by fastforce + moreover have "P (k \ f)" + by (simp add: P Qf contf imf) + ultimately obtain c where "homotopic_with P u s (k \ f) (\x. c)" + by (metis hom) + then have "homotopic_with Q u t (h \ (k \ f)) (h \ (\x. c))" + apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) + using Q by (auto simp: conth imh) + then show ?thesis + apply (rule_tac c = "h c" in that) + apply (erule homotopic_with_eq) + apply (metis feq, simp) + apply (metis Qeq) + done +qed + +lemma cohomotopically_trivial_retraction_gen: + assumes P: "\f. \continuous_on t f; f ` t \ u; Q f\ \ P(f \ h)" + and Q: "\f. \continuous_on s f; f ` s \ u; P f\ \ Q(f \ k)" + and Qeq: "\h k. (\x. x \ t \ h x = k x) \ Q h = Q k" + and hom: "\f g. \continuous_on s f; f ` s \ u; P f; + continuous_on s g; g ` s \ u; P g\ + \ homotopic_with P s u f g" + and contf: "continuous_on t f" and imf: "f ` t \ u" and Qf: "Q f" + and contg: "continuous_on t g" and img: "g ` t \ u" and Qg: "Q g" + shows "homotopic_with Q t u f g" +proof - + have feq: "\x. x \ t \ (f \ h \ k) x = f x" using idhk imf by auto + have geq: "\x. x \ t \ (g \ h \ k) x = g x" using idhk img by auto + have "continuous_on s (f \ h)" + using contf conth continuous_on_compose imh by blast + moreover have "(f \ h) ` s \ u" + using imf imh by fastforce + moreover have "P (f \ h)" + by (simp add: P Qf contf imf) + moreover have "continuous_on s (g \ h)" + using contg continuous_on_compose continuous_on_subset conth imh by blast + moreover have "(g \ h) ` s \ u" + using img imh by fastforce + moreover have "P (g \ h)" + by (simp add: P Qg contg img) + ultimately have "homotopic_with P s u (f \ h) (g \ h)" + by (rule hom) + then have "homotopic_with Q t u (f \ h \ k) (g \ h \ k)" + apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) + using Q by (auto simp: contk imk) + then show ?thesis + apply (rule homotopic_with_eq) + apply (metis feq) + apply (metis geq) + apply (metis Qeq) + done +qed + +lemma cohomotopically_trivial_retraction_null_gen: + assumes P: "\f. \continuous_on t f; f ` t \ u; Q f\ \ P(f \ h)" + and Q: "\f. \continuous_on s f; f ` s \ u; P f\ \ Q(f \ k)" + and Qeq: "\h k. (\x. x \ t \ h x = k x) \ Q h = Q k" + and hom: "\f g. \continuous_on s f; f ` s \ u; P f\ + \ \c. homotopic_with P s u f (\x. c)" + and contf: "continuous_on t f" and imf: "f ` t \ u" and Qf: "Q f" + obtains c where "homotopic_with Q t u f (\x. c)" +proof - + have feq: "\x. x \ t \ (f \ h \ k) x = f x" using idhk imf by auto + have "continuous_on s (f \ h)" + using contf conth continuous_on_compose imh by blast + moreover have "(f \ h) ` s \ u" + using imf imh by fastforce + moreover have "P (f \ h)" + by (simp add: P Qf contf imf) + ultimately obtain c where "homotopic_with P s u (f \ h) (\x. c)" + by (metis hom) + then have "homotopic_with Q t u (f \ h \ k) ((\x. c) \ k)" + apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) + using Q by (auto simp: contk imk) + then show ?thesis + apply (rule_tac c = c in that) + apply (erule homotopic_with_eq) + apply (metis feq, simp) + apply (metis Qeq) + done +qed + +end + +lemma simply_connected_retraction_gen: + shows "\simply_connected S; continuous_on S h; h ` S = T; + continuous_on T k; k ` T \ S; \y. y \ T \ h(k y) = y\ + \ simply_connected T" +apply (simp add: simply_connected_def path_def path_image_def homotopic_loops_def, clarify) +apply (rule Retracts.homotopically_trivial_retraction_gen + [of S h _ k _ "\p. pathfinish p = pathstart p" "\p. pathfinish p = pathstart p"]) +apply (simp_all add: Retracts_def pathfinish_def pathstart_def) +done + +lemma homeomorphic_simply_connected: + "\S homeomorphic T; simply_connected S\ \ simply_connected T" + by (auto simp: homeomorphic_def homeomorphism_def intro: simply_connected_retraction_gen) + +lemma homeomorphic_simply_connected_eq: + "S homeomorphic T \ (simply_connected S \ simply_connected T)" + by (metis homeomorphic_simply_connected homeomorphic_sym) + + +subsection\Homotopy equivalence\ + +definition%important homotopy_eqv :: "'a::topological_space set \ 'b::topological_space set \ bool" + (infix "homotopy'_eqv" 50) + where "S homotopy_eqv T \ + \f g. continuous_on S f \ f ` S \ T \ + continuous_on T g \ g ` T \ S \ + homotopic_with (\x. True) S S (g \ f) id \ + homotopic_with (\x. True) T T (f \ g) id" + +lemma homeomorphic_imp_homotopy_eqv: "S homeomorphic T \ S homotopy_eqv T" + unfolding homeomorphic_def homotopy_eqv_def homeomorphism_def + by (fastforce intro!: homotopic_with_equal continuous_on_compose) + +lemma homotopy_eqv_refl: "S homotopy_eqv S" + by (rule homeomorphic_imp_homotopy_eqv homeomorphic_refl)+ + +lemma homotopy_eqv_sym: "S homotopy_eqv T \ T homotopy_eqv S" + by (auto simp: homotopy_eqv_def) + +lemma homotopy_eqv_trans [trans]: + fixes S :: "'a::real_normed_vector set" and U :: "'c::real_normed_vector set" + assumes ST: "S homotopy_eqv T" and TU: "T homotopy_eqv U" + shows "S homotopy_eqv U" +proof - + obtain f1 g1 where f1: "continuous_on S f1" "f1 ` S \ T" + and g1: "continuous_on T g1" "g1 ` T \ S" + and hom1: "homotopic_with (\x. True) S S (g1 \ f1) id" + "homotopic_with (\x. True) T T (f1 \ g1) id" + using ST by (auto simp: homotopy_eqv_def) + obtain f2 g2 where f2: "continuous_on T f2" "f2 ` T \ U" + and g2: "continuous_on U g2" "g2 ` U \ T" + and hom2: "homotopic_with (\x. True) T T (g2 \ f2) id" + "homotopic_with (\x. True) U U (f2 \ g2) id" + using TU by (auto simp: homotopy_eqv_def) + have "homotopic_with (\f. True) S T (g2 \ f2 \ f1) (id \ f1)" + by (rule homotopic_with_compose_continuous_right hom2 f1)+ + then have "homotopic_with (\f. True) S T (g2 \ (f2 \ f1)) (id \ f1)" + by (simp add: o_assoc) + then have "homotopic_with (\x. True) S S + (g1 \ (g2 \ (f2 \ f1))) (g1 \ (id \ f1))" + by (simp add: g1 homotopic_with_compose_continuous_left) + moreover have "homotopic_with (\x. True) S S (g1 \ id \ f1) id" + using hom1 by simp + ultimately have SS: "homotopic_with (\x. True) S S (g1 \ g2 \ (f2 \ f1)) id" + apply (simp add: o_assoc) + apply (blast intro: homotopic_with_trans) + done + have "homotopic_with (\f. True) U T (f1 \ g1 \ g2) (id \ g2)" + by (rule homotopic_with_compose_continuous_right hom1 g2)+ + then have "homotopic_with (\f. True) U T (f1 \ (g1 \ g2)) (id \ g2)" + by (simp add: o_assoc) + then have "homotopic_with (\x. True) U U + (f2 \ (f1 \ (g1 \ g2))) (f2 \ (id \ g2))" + by (simp add: f2 homotopic_with_compose_continuous_left) + moreover have "homotopic_with (\x. True) U U (f2 \ id \ g2) id" + using hom2 by simp + ultimately have UU: "homotopic_with (\x. True) U U (f2 \ f1 \ (g1 \ g2)) id" + apply (simp add: o_assoc) + apply (blast intro: homotopic_with_trans) + done + show ?thesis + unfolding homotopy_eqv_def + apply (rule_tac x = "f2 \ f1" in exI) + apply (rule_tac x = "g1 \ g2" in exI) + apply (intro conjI continuous_on_compose SS UU) + using f1 f2 g1 g2 apply (force simp: elim!: continuous_on_subset)+ + done +qed + +lemma homotopy_eqv_inj_linear_image: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "linear f" "inj f" + shows "(f ` S) homotopy_eqv S" +apply (rule homeomorphic_imp_homotopy_eqv) +using assms homeomorphic_sym linear_homeomorphic_image by auto + +lemma homotopy_eqv_translation: + fixes S :: "'a::real_normed_vector set" + shows "(+) a ` S homotopy_eqv S" + apply (rule homeomorphic_imp_homotopy_eqv) + using homeomorphic_translation homeomorphic_sym by blast + +lemma homotopy_eqv_homotopic_triviality_imp: + fixes S :: "'a::real_normed_vector set" + and T :: "'b::real_normed_vector set" + and U :: "'c::real_normed_vector set" + assumes "S homotopy_eqv T" + and f: "continuous_on U f" "f ` U \ T" + and g: "continuous_on U g" "g ` U \ T" + and homUS: "\f g. \continuous_on U f; f ` U \ S; + continuous_on U g; g ` U \ S\ + \ homotopic_with (\x. True) U S f g" + shows "homotopic_with (\x. True) U T f g" +proof - + obtain h k where h: "continuous_on S h" "h ` S \ T" + and k: "continuous_on T k" "k ` T \ S" + and hom: "homotopic_with (\x. True) S S (k \ h) id" + "homotopic_with (\x. True) T T (h \ k) id" + using assms by (auto simp: homotopy_eqv_def) + have "homotopic_with (\f. True) U S (k \ f) (k \ g)" + apply (rule homUS) + using f g k + apply (safe intro!: continuous_on_compose h k f elim!: continuous_on_subset) + apply (force simp: o_def)+ + done + then have "homotopic_with (\x. True) U T (h \ (k \ f)) (h \ (k \ g))" + apply (rule homotopic_with_compose_continuous_left) + apply (simp_all add: h) + done + moreover have "homotopic_with (\x. True) U T (h \ k \ f) (id \ f)" + apply (rule homotopic_with_compose_continuous_right [where X=T and Y=T]) + apply (auto simp: hom f) + done + moreover have "homotopic_with (\x. True) U T (h \ k \ g) (id \ g)" + apply (rule homotopic_with_compose_continuous_right [where X=T and Y=T]) + apply (auto simp: hom g) + done + ultimately show "homotopic_with (\x. True) U T f g" + apply (simp add: o_assoc) + using homotopic_with_trans homotopic_with_sym by blast +qed + +lemma homotopy_eqv_homotopic_triviality: + fixes S :: "'a::real_normed_vector set" + and T :: "'b::real_normed_vector set" + and U :: "'c::real_normed_vector set" + assumes "S homotopy_eqv T" + shows "(\f g. continuous_on U f \ f ` U \ S \ + continuous_on U g \ g ` U \ S + \ homotopic_with (\x. True) U S f g) \ + (\f g. continuous_on U f \ f ` U \ T \ + continuous_on U g \ g ` U \ T + \ homotopic_with (\x. True) U T f g)" +apply (rule iffI) +apply (metis assms homotopy_eqv_homotopic_triviality_imp) +by (metis (no_types) assms homotopy_eqv_homotopic_triviality_imp homotopy_eqv_sym) + +lemma homotopy_eqv_cohomotopic_triviality_null_imp: + fixes S :: "'a::real_normed_vector set" + and T :: "'b::real_normed_vector set" + and U :: "'c::real_normed_vector set" + assumes "S homotopy_eqv T" + and f: "continuous_on T f" "f ` T \ U" + and homSU: "\f. \continuous_on S f; f ` S \ U\ + \ \c. homotopic_with (\x. True) S U f (\x. c)" + obtains c where "homotopic_with (\x. True) T U f (\x. c)" +proof - + obtain h k where h: "continuous_on S h" "h ` S \ T" + and k: "continuous_on T k" "k ` T \ S" + and hom: "homotopic_with (\x. True) S S (k \ h) id" + "homotopic_with (\x. True) T T (h \ k) id" + using assms by (auto simp: homotopy_eqv_def) + obtain c where "homotopic_with (\x. True) S U (f \ h) (\x. c)" + apply (rule exE [OF homSU [of "f \ h"]]) + apply (intro continuous_on_compose h) + using h f apply (force elim!: continuous_on_subset)+ + done + then have "homotopic_with (\x. True) T U ((f \ h) \ k) ((\x. c) \ k)" + apply (rule homotopic_with_compose_continuous_right [where X=S]) + using k by auto + moreover have "homotopic_with (\x. True) T U (f \ id) (f \ (h \ k))" + apply (rule homotopic_with_compose_continuous_left [where Y=T]) + apply (simp add: hom homotopic_with_symD) + using f apply auto + done + ultimately show ?thesis + apply (rule_tac c=c in that) + apply (simp add: o_def) + using homotopic_with_trans by blast +qed + +lemma homotopy_eqv_cohomotopic_triviality_null: + fixes S :: "'a::real_normed_vector set" + and T :: "'b::real_normed_vector set" + and U :: "'c::real_normed_vector set" + assumes "S homotopy_eqv T" + shows "(\f. continuous_on S f \ f ` S \ U + \ (\c. homotopic_with (\x. True) S U f (\x. c))) \ + (\f. continuous_on T f \ f ` T \ U + \ (\c. homotopic_with (\x. True) T U f (\x. c)))" +apply (rule iffI) +apply (metis assms homotopy_eqv_cohomotopic_triviality_null_imp) +by (metis assms homotopy_eqv_cohomotopic_triviality_null_imp homotopy_eqv_sym) + +lemma homotopy_eqv_homotopic_triviality_null_imp: + fixes S :: "'a::real_normed_vector set" + and T :: "'b::real_normed_vector set" + and U :: "'c::real_normed_vector set" + assumes "S homotopy_eqv T" + and f: "continuous_on U f" "f ` U \ T" + and homSU: "\f. \continuous_on U f; f ` U \ S\ + \ \c. homotopic_with (\x. True) U S f (\x. c)" + shows "\c. homotopic_with (\x. True) U T f (\x. c)" +proof - + obtain h k where h: "continuous_on S h" "h ` S \ T" + and k: "continuous_on T k" "k ` T \ S" + and hom: "homotopic_with (\x. True) S S (k \ h) id" + "homotopic_with (\x. True) T T (h \ k) id" + using assms by (auto simp: homotopy_eqv_def) + obtain c::'a where "homotopic_with (\x. True) U S (k \ f) (\x. c)" + apply (rule exE [OF homSU [of "k \ f"]]) + apply (intro continuous_on_compose h) + using k f apply (force elim!: continuous_on_subset)+ + done + then have "homotopic_with (\x. True) U T (h \ (k \ f)) (h \ (\x. c))" + apply (rule homotopic_with_compose_continuous_left [where Y=S]) + using h by auto + moreover have "homotopic_with (\x. True) U T (id \ f) ((h \ k) \ f)" + apply (rule homotopic_with_compose_continuous_right [where X=T]) + apply (simp add: hom homotopic_with_symD) + using f apply auto + done + ultimately show ?thesis + using homotopic_with_trans by (fastforce simp add: o_def) +qed + +lemma homotopy_eqv_homotopic_triviality_null: + fixes S :: "'a::real_normed_vector set" + and T :: "'b::real_normed_vector set" + and U :: "'c::real_normed_vector set" + assumes "S homotopy_eqv T" + shows "(\f. continuous_on U f \ f ` U \ S + \ (\c. homotopic_with (\x. True) U S f (\x. c))) \ + (\f. continuous_on U f \ f ` U \ T + \ (\c. homotopic_with (\x. True) U T f (\x. c)))" +apply (rule iffI) +apply (metis assms homotopy_eqv_homotopic_triviality_null_imp) +by (metis assms homotopy_eqv_homotopic_triviality_null_imp homotopy_eqv_sym) + +lemma homotopy_eqv_contractible_sets: + fixes S :: "'a::real_normed_vector set" + and T :: "'b::real_normed_vector set" + assumes "contractible S" "contractible T" "S = {} \ T = {}" + shows "S homotopy_eqv T" +proof (cases "S = {}") + case True with assms show ?thesis + by (simp add: homeomorphic_imp_homotopy_eqv) +next + case False + with assms obtain a b where "a \ S" "b \ T" + by auto + then show ?thesis + unfolding homotopy_eqv_def + apply (rule_tac x="\x. b" in exI) + apply (rule_tac x="\x. a" in exI) + apply (intro assms conjI continuous_on_id' homotopic_into_contractible) + apply (auto simp: o_def continuous_on_const) + done +qed + +lemma homotopy_eqv_empty1 [simp]: + fixes S :: "'a::real_normed_vector set" + shows "S homotopy_eqv ({}::'b::real_normed_vector set) \ S = {}" +apply (rule iffI) +using homotopy_eqv_def apply fastforce +by (simp add: homotopy_eqv_contractible_sets) + +lemma homotopy_eqv_empty2 [simp]: + fixes S :: "'a::real_normed_vector set" + shows "({}::'b::real_normed_vector set) homotopy_eqv S \ S = {}" +by (metis homotopy_eqv_empty1 homotopy_eqv_sym) + +lemma homotopy_eqv_contractibility: + fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" + shows "S homotopy_eqv T \ (contractible S \ contractible T)" +unfolding homotopy_eqv_def +by (blast intro: homotopy_dominated_contractibility) + +lemma homotopy_eqv_sing: + fixes S :: "'a::real_normed_vector set" and a :: "'b::real_normed_vector" + shows "S homotopy_eqv {a} \ S \ {} \ contractible S" +proof (cases "S = {}") + case True then show ?thesis + by simp +next + case False then show ?thesis + by (metis contractible_sing empty_not_insert homotopy_eqv_contractibility homotopy_eqv_contractible_sets) +qed + +lemma homeomorphic_contractible_eq: + fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" + shows "S homeomorphic T \ (contractible S \ contractible T)" +by (simp add: homeomorphic_imp_homotopy_eqv homotopy_eqv_contractibility) + +lemma homeomorphic_contractible: + fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" + shows "\contractible S; S homeomorphic T\ \ contractible T" + by (metis homeomorphic_contractible_eq) + + +subsection%unimportant\Misc other results\ + +lemma bounded_connected_Compl_real: + fixes S :: "real set" + assumes "bounded S" and conn: "connected(- S)" + shows "S = {}" +proof - + obtain a b where "S \ box a b" + by (meson assms bounded_subset_box_symmetric) + then have "a \ S" "b \ S" + by auto + then have "\x. a \ x \ x \ b \ x \ - S" + by (meson Compl_iff conn connected_iff_interval) + then show ?thesis + using \S \ box a b\ by auto +qed + +lemma bounded_connected_Compl_1: + fixes S :: "'a::{euclidean_space} set" + assumes "bounded S" and conn: "connected(- S)" and 1: "DIM('a) = 1" + shows "S = {}" +proof - + have "DIM('a) = DIM(real)" + by (simp add: "1") + then obtain f::"'a \ real" and g + where "linear f" "\x. norm(f x) = norm x" "\x. g(f x) = x" "\y. f(g y) = y" + by (rule isomorphisms_UNIV_UNIV) blast + with \bounded S\ have "bounded (f ` S)" + using bounded_linear_image linear_linear by blast + have "connected (f ` (-S))" + using connected_linear_image assms \linear f\ by blast + moreover have "f ` (-S) = - (f ` S)" + apply (rule bij_image_Compl_eq) + apply (auto simp: bij_def) + apply (metis \\x. g (f x) = x\ injI) + by (metis UNIV_I \\y. f (g y) = y\ image_iff) + finally have "connected (- (f ` S))" + by simp + then have "f ` S = {}" + using \bounded (f ` S)\ bounded_connected_Compl_real by blast + then show ?thesis + by blast +qed + + +subsection%unimportant\Some Uncountable Sets\ + +lemma uncountable_closed_segment: + fixes a :: "'a::real_normed_vector" + assumes "a \ b" shows "uncountable (closed_segment a b)" +unfolding path_image_linepath [symmetric] path_image_def + using inj_on_linepath [OF assms] uncountable_closed_interval [of 0 1] + countable_image_inj_on by auto + +lemma uncountable_open_segment: + fixes a :: "'a::real_normed_vector" + assumes "a \ b" shows "uncountable (open_segment a b)" + by (simp add: assms open_segment_def uncountable_closed_segment uncountable_minus_countable) + +lemma uncountable_convex: + fixes a :: "'a::real_normed_vector" + assumes "convex S" "a \ S" "b \ S" "a \ b" + shows "uncountable S" +proof - + have "uncountable (closed_segment a b)" + by (simp add: uncountable_closed_segment assms) + then show ?thesis + by (meson assms convex_contains_segment countable_subset) +qed + +lemma uncountable_ball: + fixes a :: "'a::euclidean_space" + assumes "r > 0" + shows "uncountable (ball a r)" +proof - + have "uncountable (open_segment a (a + r *\<^sub>R (SOME i. i \ Basis)))" + by (metis Basis_zero SOME_Basis add_cancel_right_right assms less_le scale_eq_0_iff uncountable_open_segment) + moreover have "open_segment a (a + r *\<^sub>R (SOME i. i \ Basis)) \ ball a r" + using assms by (auto simp: in_segment algebra_simps dist_norm SOME_Basis) + ultimately show ?thesis + by (metis countable_subset) +qed + +lemma ball_minus_countable_nonempty: + assumes "countable (A :: 'a :: euclidean_space set)" "r > 0" + shows "ball z r - A \ {}" +proof + assume *: "ball z r - A = {}" + have "uncountable (ball z r - A)" + by (intro uncountable_minus_countable assms uncountable_ball) + thus False by (subst (asm) *) auto +qed + +lemma uncountable_cball: + fixes a :: "'a::euclidean_space" + assumes "r > 0" + shows "uncountable (cball a r)" + using assms countable_subset uncountable_ball by auto + +lemma pairwise_disjnt_countable: + fixes \ :: "nat set set" + assumes "pairwise disjnt \" + shows "countable \" +proof - + have "inj_on (\X. SOME n. n \ X) (\ - {{}})" + apply (clarsimp simp add: inj_on_def) + by (metis assms disjnt_insert2 insert_absorb pairwise_def subsetI subset_empty tfl_some) + then show ?thesis + by (metis countable_Diff_eq countable_def) +qed + +lemma pairwise_disjnt_countable_Union: + assumes "countable (\\)" and pwd: "pairwise disjnt \" + shows "countable \" +proof - + obtain f :: "_ \ nat" where f: "inj_on f (\\)" + using assms by blast + then have "pairwise disjnt (\ X \ \. {f ` X})" + using assms by (force simp: pairwise_def disjnt_inj_on_iff [OF f]) + then have "countable (\ X \ \. {f ` X})" + using pairwise_disjnt_countable by blast + then show ?thesis + by (meson pwd countable_image_inj_on disjoint_image f inj_on_image pairwise_disjnt_countable) +qed + +lemma connected_uncountable: + fixes S :: "'a::metric_space set" + assumes "connected S" "a \ S" "b \ S" "a \ b" shows "uncountable S" +proof - + have "continuous_on S (dist a)" + by (intro continuous_intros) + then have "connected (dist a ` S)" + by (metis connected_continuous_image \connected S\) + then have "closed_segment 0 (dist a b) \ (dist a ` S)" + by (simp add: assms closed_segment_subset is_interval_connected_1 is_interval_convex) + then have "uncountable (dist a ` S)" + by (metis \a \ b\ countable_subset dist_eq_0_iff uncountable_closed_segment) + then show ?thesis + by blast +qed + +lemma path_connected_uncountable: + fixes S :: "'a::metric_space set" + assumes "path_connected S" "a \ S" "b \ S" "a \ b" shows "uncountable S" + using path_connected_imp_connected assms connected_uncountable by metis + +lemma connected_finite_iff_sing: + fixes S :: "'a::metric_space set" + assumes "connected S" + shows "finite S \ S = {} \ (\a. S = {a})" (is "_ = ?rhs") +proof - + have "uncountable S" if "\ ?rhs" + using connected_uncountable assms that by blast + then show ?thesis + using uncountable_infinite by auto +qed + +lemma connected_card_eq_iff_nontrivial: + fixes S :: "'a::metric_space set" + shows "connected S \ uncountable S \ \(\a. S \ {a})" + apply (auto simp: countable_finite finite_subset) + by (metis connected_uncountable is_singletonI' is_singleton_the_elem subset_singleton_iff) + +lemma simple_path_image_uncountable: + fixes g :: "real \ 'a::metric_space" + assumes "simple_path g" + shows "uncountable (path_image g)" +proof - + have "g 0 \ path_image g" "g (1/2) \ path_image g" + by (simp_all add: path_defs) + moreover have "g 0 \ g (1/2)" + using assms by (fastforce simp add: simple_path_def) + ultimately show ?thesis + apply (simp add: assms connected_card_eq_iff_nontrivial connected_simple_path_image) + by blast +qed + +lemma arc_image_uncountable: + fixes g :: "real \ 'a::metric_space" + assumes "arc g" + shows "uncountable (path_image g)" + by (simp add: arc_imp_simple_path assms simple_path_image_uncountable) + + +subsection%unimportant\ Some simple positive connection theorems\ + +proposition path_connected_convex_diff_countable: + fixes U :: "'a::euclidean_space set" + assumes "convex U" "\ collinear U" "countable S" + shows "path_connected(U - S)" +proof (clarsimp simp add: path_connected_def) + fix a b + assume "a \ U" "a \ S" "b \ U" "b \ S" + let ?m = "midpoint a b" + show "\g. path g \ path_image g \ U - S \ pathstart g = a \ pathfinish g = b" + proof (cases "a = b") + case True + then show ?thesis + by (metis DiffI \a \ U\ \a \ S\ path_component_def path_component_refl) + next + case False + then have "a \ ?m" "b \ ?m" + using midpoint_eq_endpoint by fastforce+ + have "?m \ U" + using \a \ U\ \b \ U\ \convex U\ convex_contains_segment by force + obtain c where "c \ U" and nc_abc: "\ collinear {a,b,c}" + by (metis False \a \ U\ \b \ U\ \\ collinear U\ collinear_triples insert_absorb) + have ncoll_mca: "\ collinear {?m,c,a}" + by (metis (full_types) \a \ ?m\ collinear_3_trans collinear_midpoint insert_commute nc_abc) + have ncoll_mcb: "\ collinear {?m,c,b}" + by (metis (full_types) \b \ ?m\ collinear_3_trans collinear_midpoint insert_commute nc_abc) + have "c \ ?m" + by (metis collinear_midpoint insert_commute nc_abc) + then have "closed_segment ?m c \ U" + by (simp add: \c \ U\ \?m \ U\ \convex U\ closed_segment_subset) + then obtain z where z: "z \ closed_segment ?m c" + and disjS: "(closed_segment a z \ closed_segment z b) \ S = {}" + proof - + have False if "closed_segment ?m c \ {z. (closed_segment a z \ closed_segment z b) \ S \ {}}" + proof - + have closb: "closed_segment ?m c \ + {z \ closed_segment ?m c. closed_segment a z \ S \ {}} \ {z \ closed_segment ?m c. closed_segment z b \ S \ {}}" + using that by blast + have *: "countable {z \ closed_segment ?m c. closed_segment z u \ S \ {}}" + if "u \ U" "u \ S" and ncoll: "\ collinear {?m, c, u}" for u + proof - + have **: False if x1: "x1 \ closed_segment ?m c" and x2: "x2 \ closed_segment ?m c" + and "x1 \ x2" "x1 \ u" + and w: "w \ closed_segment x1 u" "w \ closed_segment x2 u" + and "w \ S" for x1 x2 w + proof - + have "x1 \ affine hull {?m,c}" "x2 \ affine hull {?m,c}" + using segment_as_ball x1 x2 by auto + then have coll_x1: "collinear {x1, ?m, c}" and coll_x2: "collinear {?m, c, x2}" + by (simp_all add: affine_hull_3_imp_collinear) (metis affine_hull_3_imp_collinear insert_commute) + have "\ collinear {x1, u, x2}" + proof + assume "collinear {x1, u, x2}" + then have "collinear {?m, c, u}" + by (metis (full_types) \c \ ?m\ coll_x1 coll_x2 collinear_3_trans insert_commute ncoll \x1 \ x2\) + with ncoll show False .. + qed + then have "closed_segment x1 u \ closed_segment u x2 = {u}" + by (blast intro!: Int_closed_segment) + then have "w = u" + using closed_segment_commute w by auto + show ?thesis + using \u \ S\ \w = u\ that(7) by auto + qed + then have disj: "disjoint ((\z\closed_segment ?m c. {closed_segment z u \ S}))" + by (fastforce simp: pairwise_def disjnt_def) + have cou: "countable ((\z \ closed_segment ?m c. {closed_segment z u \ S}) - {{}})" + apply (rule pairwise_disjnt_countable_Union [OF _ pairwise_subset [OF disj]]) + apply (rule countable_subset [OF _ \countable S\], auto) + done + define f where "f \ \X. (THE z. z \ closed_segment ?m c \ X = closed_segment z u \ S)" + show ?thesis + proof (rule countable_subset [OF _ countable_image [OF cou, where f=f]], clarify) + fix x + assume x: "x \ closed_segment ?m c" "closed_segment x u \ S \ {}" + show "x \ f ` ((\z\closed_segment ?m c. {closed_segment z u \ S}) - {{}})" + proof (rule_tac x="closed_segment x u \ S" in image_eqI) + show "x = f (closed_segment x u \ S)" + unfolding f_def + apply (rule the_equality [symmetric]) + using x apply (auto simp: dest: **) + done + qed (use x in auto) + qed + qed + have "uncountable (closed_segment ?m c)" + by (metis \c \ ?m\ uncountable_closed_segment) + then show False + using closb * [OF \a \ U\ \a \ S\ ncoll_mca] * [OF \b \ U\ \b \ S\ ncoll_mcb] + apply (simp add: closed_segment_commute) + by (simp add: countable_subset) + qed + then show ?thesis + by (force intro: that) + qed + show ?thesis + proof (intro exI conjI) + have "path_image (linepath a z +++ linepath z b) \ U" + by (metis \a \ U\ \b \ U\ \closed_segment ?m c \ U\ z \convex U\ closed_segment_subset contra_subsetD path_image_linepath subset_path_image_join) + with disjS show "path_image (linepath a z +++ linepath z b) \ U - S" + by (force simp: path_image_join) + qed auto + qed +qed + + +corollary connected_convex_diff_countable: + fixes U :: "'a::euclidean_space set" + assumes "convex U" "\ collinear U" "countable S" + shows "connected(U - S)" + by (simp add: assms path_connected_convex_diff_countable path_connected_imp_connected) + +lemma path_connected_punctured_convex: + assumes "convex S" and aff: "aff_dim S \ 1" + shows "path_connected(S - {a})" +proof - + consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S \ 2" + using assms aff_dim_geq [of S] by linarith + then show ?thesis + proof cases + assume "aff_dim S = -1" + then show ?thesis + by (metis aff_dim_empty empty_Diff path_connected_empty) + next + assume "aff_dim S = 0" + then show ?thesis + by (metis aff_dim_eq_0 Diff_cancel Diff_empty Diff_insert0 convex_empty convex_imp_path_connected path_connected_singleton singletonD) + next + assume ge2: "aff_dim S \ 2" + then have "\ collinear S" + proof (clarsimp simp add: collinear_affine_hull) + fix u v + assume "S \ affine hull {u, v}" + then have "aff_dim S \ aff_dim {u, v}" + by (metis (no_types) aff_dim_affine_hull aff_dim_subset) + with ge2 show False + by (metis (no_types) aff_dim_2 antisym aff not_numeral_le_zero one_le_numeral order_trans) + qed + then show ?thesis + apply (rule path_connected_convex_diff_countable [OF \convex S\]) + by simp + qed +qed + +lemma connected_punctured_convex: + shows "\convex S; aff_dim S \ 1\ \ connected(S - {a})" + using path_connected_imp_connected path_connected_punctured_convex by blast + +lemma path_connected_complement_countable: + fixes S :: "'a::euclidean_space set" + assumes "2 \ DIM('a)" "countable S" + shows "path_connected(- S)" +proof - + have "path_connected(UNIV - S)" + apply (rule path_connected_convex_diff_countable) + using assms by (auto simp: collinear_aff_dim [of "UNIV :: 'a set"]) + then show ?thesis + by (simp add: Compl_eq_Diff_UNIV) +qed + +proposition path_connected_openin_diff_countable: + fixes S :: "'a::euclidean_space set" + assumes "connected S" and ope: "openin (subtopology euclidean (affine hull S)) S" + and "\ collinear S" "countable T" + shows "path_connected(S - T)" +proof (clarsimp simp add: path_connected_component) + fix x y + assume xy: "x \ S" "x \ T" "y \ S" "y \ T" + show "path_component (S - T) x y" + proof (rule connected_equivalence_relation_gen [OF \connected S\, where P = "\x. x \ T"]) + show "\z. z \ U \ z \ T" if opeU: "openin (subtopology euclidean S) U" and "x \ U" for U x + proof - + have "openin (subtopology euclidean (affine hull S)) U" + using opeU ope openin_trans by blast + with \x \ U\ obtain r where Usub: "U \ affine hull S" and "r > 0" + and subU: "ball x r \ affine hull S \ U" + by (auto simp: openin_contains_ball) + with \x \ U\ have x: "x \ ball x r \ affine hull S" + by auto + have "\ S \ {x}" + using \\ collinear S\ collinear_subset by blast + then obtain x' where "x' \ x" "x' \ S" + by blast + obtain y where y: "y \ x" "y \ ball x r \ affine hull S" + proof + show "x + (r / 2 / norm(x' - x)) *\<^sub>R (x' - x) \ x" + using \x' \ x\ \r > 0\ by auto + show "x + (r / 2 / norm (x' - x)) *\<^sub>R (x' - x) \ ball x r \ affine hull S" + using \x' \ x\ \r > 0\ \x' \ S\ x + by (simp add: dist_norm mem_affine_3_minus hull_inc) + qed + have "convex (ball x r \ affine hull S)" + by (simp add: affine_imp_convex convex_Int) + with x y subU have "uncountable U" + by (meson countable_subset uncountable_convex) + then have "\ U \ T" + using \countable T\ countable_subset by blast + then show ?thesis by blast + qed + show "\U. openin (subtopology euclidean S) U \ x \ U \ + (\x\U. \y\U. x \ T \ y \ T \ path_component (S - T) x y)" + if "x \ S" for x + proof - + obtain r where Ssub: "S \ affine hull S" and "r > 0" + and subS: "ball x r \ affine hull S \ S" + using ope \x \ S\ by (auto simp: openin_contains_ball) + then have conv: "convex (ball x r \ affine hull S)" + by (simp add: affine_imp_convex convex_Int) + have "\ aff_dim (affine hull S) \ 1" + using \\ collinear S\ collinear_aff_dim by auto + then have "\ collinear (ball x r \ affine hull S)" + apply (simp add: collinear_aff_dim) + by (metis (no_types, hide_lams) aff_dim_convex_Int_open IntI open_ball \0 < r\ aff_dim_affine_hull affine_affine_hull affine_imp_convex centre_in_ball empty_iff hull_subset inf_commute subsetCE that) + then have *: "path_connected ((ball x r \ affine hull S) - T)" + by (rule path_connected_convex_diff_countable [OF conv _ \countable T\]) + have ST: "ball x r \ affine hull S - T \ S - T" + using subS by auto + show ?thesis + proof (intro exI conjI) + show "x \ ball x r \ affine hull S" + using \x \ S\ \r > 0\ by (simp add: hull_inc) + have "openin (subtopology euclidean (affine hull S)) (ball x r \ affine hull S)" + by (subst inf.commute) (simp add: openin_Int_open) + then show "openin (subtopology euclidean S) (ball x r \ affine hull S)" + by (rule openin_subset_trans [OF _ subS Ssub]) + qed (use * path_component_trans in \auto simp: path_connected_component path_component_of_subset [OF ST]\) + qed + qed (use xy path_component_trans in auto) +qed + +corollary connected_openin_diff_countable: + fixes S :: "'a::euclidean_space set" + assumes "connected S" and ope: "openin (subtopology euclidean (affine hull S)) S" + and "\ collinear S" "countable T" + shows "connected(S - T)" + by (metis path_connected_imp_connected path_connected_openin_diff_countable [OF assms]) + +corollary path_connected_open_diff_countable: + fixes S :: "'a::euclidean_space set" + assumes "2 \ DIM('a)" "open S" "connected S" "countable T" + shows "path_connected(S - T)" +proof (cases "S = {}") + case True + then show ?thesis + by (simp add: path_connected_empty) +next + case False + show ?thesis + proof (rule path_connected_openin_diff_countable) + show "openin (subtopology euclidean (affine hull S)) S" + by (simp add: assms hull_subset open_subset) + show "\ collinear S" + using assms False by (simp add: collinear_aff_dim aff_dim_open) + qed (simp_all add: assms) +qed + +corollary connected_open_diff_countable: + fixes S :: "'a::euclidean_space set" + assumes "2 \ DIM('a)" "open S" "connected S" "countable T" + shows "connected(S - T)" +by (simp add: assms path_connected_imp_connected path_connected_open_diff_countable) + + + +subsection%unimportant \Self-homeomorphisms shuffling points about\ + +subsubsection%unimportant\The theorem \homeomorphism_moving_points_exists\\ + +lemma homeomorphism_moving_point_1: + fixes a :: "'a::euclidean_space" + assumes "affine T" "a \ T" and u: "u \ ball a r \ T" + obtains f g where "homeomorphism (cball a r \ T) (cball a r \ T) f g" + "f a = u" "\x. x \ sphere a r \ f x = x" +proof - + have nou: "norm (u - a) < r" and "u \ T" + using u by (auto simp: dist_norm norm_minus_commute) + then have "0 < r" + by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) + define f where "f \ \x. (1 - norm(x - a) / r) *\<^sub>R (u - a) + x" + have *: "False" if eq: "x + (norm y / r) *\<^sub>R u = y + (norm x / r) *\<^sub>R u" + and nou: "norm u < r" and yx: "norm y < norm x" for x y and u::'a + proof - + have "x = y + (norm x / r - (norm y / r)) *\<^sub>R u" + using eq by (simp add: algebra_simps) + then have "norm x = norm (y + ((norm x - norm y) / r) *\<^sub>R u)" + by (metis diff_divide_distrib) + also have "\ \ norm y + norm(((norm x - norm y) / r) *\<^sub>R u)" + using norm_triangle_ineq by blast + also have "\ = norm y + (norm x - norm y) * (norm u / r)" + using yx \r > 0\ + by (simp add: divide_simps) + also have "\ < norm y + (norm x - norm y) * 1" + apply (subst add_less_cancel_left) + apply (rule mult_strict_left_mono) + using nou \0 < r\ yx + apply (simp_all add: field_simps) + done + also have "\ = norm x" + by simp + finally show False by simp + qed + have "inj f" + unfolding f_def + proof (clarsimp simp: inj_on_def) + fix x y + assume "(1 - norm (x - a) / r) *\<^sub>R (u - a) + x = + (1 - norm (y - a) / r) *\<^sub>R (u - a) + y" + then have eq: "(x - a) + (norm (y - a) / r) *\<^sub>R (u - a) = (y - a) + (norm (x - a) / r) *\<^sub>R (u - a)" + by (auto simp: algebra_simps) + show "x=y" + proof (cases "norm (x - a) = norm (y - a)") + case True + then show ?thesis + using eq by auto + next + case False + then consider "norm (x - a) < norm (y - a)" | "norm (x - a) > norm (y - a)" + by linarith + then have "False" + proof cases + case 1 show False + using * [OF _ nou 1] eq by simp + next + case 2 with * [OF eq nou] show False + by auto + qed + then show "x=y" .. + qed + qed + then have inj_onf: "inj_on f (cball a r \ T)" + using inj_on_Int by fastforce + have contf: "continuous_on (cball a r \ T) f" + unfolding f_def using \0 < r\ by (intro continuous_intros) blast + have fim: "f ` (cball a r \ T) = cball a r \ T" + proof + have *: "norm (y + (1 - norm y / r) *\<^sub>R u) \ r" if "norm y \ r" "norm u < r" for y u::'a + proof - + have "norm (y + (1 - norm y / r) *\<^sub>R u) \ norm y + norm((1 - norm y / r) *\<^sub>R u)" + using norm_triangle_ineq by blast + also have "\ = norm y + abs(1 - norm y / r) * norm u" + by simp + also have "\ \ r" + proof - + have "(r - norm u) * (r - norm y) \ 0" + using that by auto + then have "r * norm u + r * norm y \ r * r + norm u * norm y" + by (simp add: algebra_simps) + then show ?thesis + using that \0 < r\ by (simp add: abs_if field_simps) + qed + finally show ?thesis . + qed + have "f ` (cball a r) \ cball a r" + apply (clarsimp simp add: dist_norm norm_minus_commute f_def) + using * by (metis diff_add_eq diff_diff_add diff_diff_eq2 norm_minus_commute nou) + moreover have "f ` T \ T" + unfolding f_def using \affine T\ \a \ T\ \u \ T\ + by (force simp: add.commute mem_affine_3_minus) + ultimately show "f ` (cball a r \ T) \ cball a r \ T" + by blast + next + show "cball a r \ T \ f ` (cball a r \ T)" + proof (clarsimp simp add: dist_norm norm_minus_commute) + fix x + assume x: "norm (x - a) \ r" and "x \ T" + have "\v \ {0..1}. ((1 - v) * r - norm ((x - a) - v *\<^sub>R (u - a))) \ 1 = 0" + by (rule ivt_decreasing_component_on_1) (auto simp: x continuous_intros) + then obtain v where "0\v" "v\1" and v: "(1 - v) * r = norm ((x - a) - v *\<^sub>R (u - a))" + by auto + show "x \ f ` (cball a r \ T)" + proof (rule image_eqI) + show "x = f (x - v *\<^sub>R (u - a))" + using \r > 0\ v by (simp add: f_def field_simps) + have "x - v *\<^sub>R (u - a) \ cball a r" + using \r > 0\ v \0 \ v\ + apply (simp add: field_simps dist_norm norm_minus_commute) + by (metis le_add_same_cancel2 order.order_iff_strict zero_le_mult_iff) + moreover have "x - v *\<^sub>R (u - a) \ T" + by (simp add: f_def \affine T\ \u \ T\ \x \ T\ assms mem_affine_3_minus2) + ultimately show "x - v *\<^sub>R (u - a) \ cball a r \ T" + by blast + qed + qed + qed + have "\g. homeomorphism (cball a r \ T) (cball a r \ T) f g" + apply (rule homeomorphism_compact [OF _ contf fim inj_onf]) + apply (simp add: affine_closed compact_Int_closed \affine T\) + done + then show ?thesis + apply (rule exE) + apply (erule_tac f=f in that) + using \r > 0\ + apply (simp_all add: f_def dist_norm norm_minus_commute) + done +qed + +corollary%unimportant homeomorphism_moving_point_2: + fixes a :: "'a::euclidean_space" + assumes "affine T" "a \ T" and u: "u \ ball a r \ T" and v: "v \ ball a r \ T" + obtains f g where "homeomorphism (cball a r \ T) (cball a r \ T) f g" + "f u = v" "\x. \x \ sphere a r; x \ T\ \ f x = x" +proof - + have "0 < r" + by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) + obtain f1 g1 where hom1: "homeomorphism (cball a r \ T) (cball a r \ T) f1 g1" + and "f1 a = u" and f1: "\x. x \ sphere a r \ f1 x = x" + using homeomorphism_moving_point_1 [OF \affine T\ \a \ T\ u] by blast + obtain f2 g2 where hom2: "homeomorphism (cball a r \ T) (cball a r \ T) f2 g2" + and "f2 a = v" and f2: "\x. x \ sphere a r \ f2 x = x" + using homeomorphism_moving_point_1 [OF \affine T\ \a \ T\ v] by blast + show ?thesis + proof + show "homeomorphism (cball a r \ T) (cball a r \ T) (f2 \ g1) (f1 \ g2)" + by (metis homeomorphism_compose homeomorphism_symD hom1 hom2) + have "g1 u = a" + using \0 < r\ \f1 a = u\ assms hom1 homeomorphism_apply1 by fastforce + then show "(f2 \ g1) u = v" + by (simp add: \f2 a = v\) + show "\x. \x \ sphere a r; x \ T\ \ (f2 \ g1) x = x" + using f1 f2 hom1 homeomorphism_apply1 by fastforce + qed +qed + + +corollary%unimportant homeomorphism_moving_point_3: + fixes a :: "'a::euclidean_space" + assumes "affine T" "a \ T" and ST: "ball a r \ T \ S" "S \ T" + and u: "u \ ball a r \ T" and v: "v \ ball a r \ T" + obtains f g where "homeomorphism S S f g" + "f u = v" "{x. \ (f x = x \ g x = x)} \ ball a r \ T" +proof - + obtain f g where hom: "homeomorphism (cball a r \ T) (cball a r \ T) f g" + and "f u = v" and fid: "\x. \x \ sphere a r; x \ T\ \ f x = x" + using homeomorphism_moving_point_2 [OF \affine T\ \a \ T\ u v] by blast + have gid: "\x. \x \ sphere a r; x \ T\ \ g x = x" + using fid hom homeomorphism_apply1 by fastforce + define ff where "ff \ \x. if x \ ball a r \ T then f x else x" + define gg where "gg \ \x. if x \ ball a r \ T then g x else x" + show ?thesis + proof + show "homeomorphism S S ff gg" + proof (rule homeomorphismI) + have "continuous_on ((cball a r \ T) \ (T - ball a r)) ff" + apply (simp add: ff_def) + apply (rule continuous_on_cases) + using homeomorphism_cont1 [OF hom] + apply (auto simp: affine_closed \affine T\ continuous_on_id fid) + done + then show "continuous_on S ff" + apply (rule continuous_on_subset) + using ST by auto + have "continuous_on ((cball a r \ T) \ (T - ball a r)) gg" + apply (simp add: gg_def) + apply (rule continuous_on_cases) + using homeomorphism_cont2 [OF hom] + apply (auto simp: affine_closed \affine T\ continuous_on_id gid) + done + then show "continuous_on S gg" + apply (rule continuous_on_subset) + using ST by auto + show "ff ` S \ S" + proof (clarsimp simp add: ff_def) + fix x + assume "x \ S" and x: "dist a x < r" and "x \ T" + then have "f x \ cball a r \ T" + using homeomorphism_image1 [OF hom] by force + then show "f x \ S" + using ST(1) \x \ T\ gid hom homeomorphism_def x by fastforce + qed + show "gg ` S \ S" + proof (clarsimp simp add: gg_def) + fix x + assume "x \ S" and x: "dist a x < r" and "x \ T" + then have "g x \ cball a r \ T" + using homeomorphism_image2 [OF hom] by force + then have "g x \ ball a r" + using homeomorphism_apply2 [OF hom] + by (metis Diff_Diff_Int Diff_iff \x \ T\ cball_def fid le_less mem_Collect_eq mem_ball mem_sphere x) + then show "g x \ S" + using ST(1) \g x \ cball a r \ T\ by force + qed + show "\x. x \ S \ gg (ff x) = x" + apply (simp add: ff_def gg_def) + using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] + apply auto + apply (metis Int_iff homeomorphism_apply1 [OF hom] fid image_eqI less_eq_real_def mem_cball mem_sphere) + done + show "\x. x \ S \ ff (gg x) = x" + apply (simp add: ff_def gg_def) + using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] + apply auto + apply (metis Int_iff fid image_eqI less_eq_real_def mem_cball mem_sphere) + done + qed + show "ff u = v" + using u by (auto simp: ff_def \f u = v\) + show "{x. \ (ff x = x \ gg x = x)} \ ball a r \ T" + by (auto simp: ff_def gg_def) + qed +qed + + +proposition%unimportant homeomorphism_moving_point: + fixes a :: "'a::euclidean_space" + assumes ope: "openin (subtopology euclidean (affine hull S)) S" + and "S \ T" + and TS: "T \ affine hull S" + and S: "connected S" "a \ S" "b \ S" + obtains f g where "homeomorphism T T f g" "f a = b" + "{x. \ (f x = x \ g x = x)} \ S" + "bounded {x. \ (f x = x \ g x = x)}" +proof - + have 1: "\h k. homeomorphism T T h k \ h (f d) = d \ + {x. \ (h x = x \ k x = x)} \ S \ bounded {x. \ (h x = x \ k x = x)}" + if "d \ S" "f d \ S" and homfg: "homeomorphism T T f g" + and S: "{x. \ (f x = x \ g x = x)} \ S" + and bo: "bounded {x. \ (f x = x \ g x = x)}" for d f g + proof (intro exI conjI) + show homgf: "homeomorphism T T g f" + by (metis homeomorphism_symD homfg) + then show "g (f d) = d" + by (meson \S \ T\ homeomorphism_def subsetD \d \ S\) + show "{x. \ (g x = x \ f x = x)} \ S" + using S by blast + show "bounded {x. \ (g x = x \ f x = x)}" + using bo by (simp add: conj_commute) + qed + have 2: "\f g. homeomorphism T T f g \ f x = f2 (f1 x) \ + {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" + if "x \ S" "f1 x \ S" "f2 (f1 x) \ S" + and hom: "homeomorphism T T f1 g1" "homeomorphism T T f2 g2" + and sub: "{x. \ (f1 x = x \ g1 x = x)} \ S" "{x. \ (f2 x = x \ g2 x = x)} \ S" + and bo: "bounded {x. \ (f1 x = x \ g1 x = x)}" "bounded {x. \ (f2 x = x \ g2 x = x)}" + for x f1 f2 g1 g2 + proof (intro exI conjI) + show homgf: "homeomorphism T T (f2 \ f1) (g1 \ g2)" + by (metis homeomorphism_compose hom) + then show "(f2 \ f1) x = f2 (f1 x)" + by force + show "{x. \ ((f2 \ f1) x = x \ (g1 \ g2) x = x)} \ S" + using sub by force + have "bounded ({x. \(f1 x = x \ g1 x = x)} \ {x. \(f2 x = x \ g2 x = x)})" + using bo by simp + then show "bounded {x. \ ((f2 \ f1) x = x \ (g1 \ g2) x = x)}" + by (rule bounded_subset) auto + qed + have 3: "\U. openin (subtopology euclidean S) U \ + d \ U \ + (\x\U. + \f g. homeomorphism T T f g \ f d = x \ + {x. \ (f x = x \ g x = x)} \ S \ + bounded {x. \ (f x = x \ g x = x)})" + if "d \ S" for d + proof - + obtain r where "r > 0" and r: "ball d r \ affine hull S \ S" + by (metis \d \ S\ ope openin_contains_ball) + have *: "\f g. homeomorphism T T f g \ f d = e \ + {x. \ (f x = x \ g x = x)} \ S \ + bounded {x. \ (f x = x \ g x = x)}" if "e \ S" "e \ ball d r" for e + apply (rule homeomorphism_moving_point_3 [of "affine hull S" d r T d e]) + using r \S \ T\ TS that + apply (auto simp: \d \ S\ \0 < r\ hull_inc) + using bounded_subset by blast + show ?thesis + apply (rule_tac x="S \ ball d r" in exI) + apply (intro conjI) + apply (simp add: openin_open_Int) + apply (simp add: \0 < r\ that) + apply (blast intro: *) + done + qed + have "\f g. homeomorphism T T f g \ f a = b \ + {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" + apply (rule connected_equivalence_relation [OF S], safe) + apply (blast intro: 1 2 3)+ + done + then show ?thesis + using that by auto +qed + + +lemma homeomorphism_moving_points_exists_gen: + assumes K: "finite K" "\i. i \ K \ x i \ S \ y i \ S" + "pairwise (\i j. (x i \ x j) \ (y i \ y j)) K" + and "2 \ aff_dim S" + and ope: "openin (subtopology euclidean (affine hull S)) S" + and "S \ T" "T \ affine hull S" "connected S" + shows "\f g. homeomorphism T T f g \ (\i \ K. f(x i) = y i) \ + {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" + using assms +proof (induction K) + case empty + then show ?case + by (force simp: homeomorphism_ident) +next + case (insert i K) + then have xney: "\j. \j \ K; j \ i\ \ x i \ x j \ y i \ y j" + and pw: "pairwise (\i j. x i \ x j \ y i \ y j) K" + and "x i \ S" "y i \ S" + and xyS: "\i. i \ K \ x i \ S \ y i \ S" + by (simp_all add: pairwise_insert) + obtain f g where homfg: "homeomorphism T T f g" and feq: "\i. i \ K \ f(x i) = y i" + and fg_sub: "{x. \ (f x = x \ g x = x)} \ S" + and bo_fg: "bounded {x. \ (f x = x \ g x = x)}" + using insert.IH [OF xyS pw] insert.prems by (blast intro: that) + then have "\f g. homeomorphism T T f g \ (\i \ K. f(x i) = y i) \ + {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" + using insert by blast + have aff_eq: "affine hull (S - y ` K) = affine hull S" + apply (rule affine_hull_Diff) + apply (auto simp: insert) + using \y i \ S\ insert.hyps(2) xney xyS by fastforce + have f_in_S: "f x \ S" if "x \ S" for x + using homfg fg_sub homeomorphism_apply1 \S \ T\ + proof - + have "(f (f x) \ f x \ g (f x) \ f x) \ f x \ S" + by (metis \S \ T\ homfg subsetD homeomorphism_apply1 that) + then show ?thesis + using fg_sub by force + qed + obtain h k where homhk: "homeomorphism T T h k" and heq: "h (f (x i)) = y i" + and hk_sub: "{x. \ (h x = x \ k x = x)} \ S - y ` K" + and bo_hk: "bounded {x. \ (h x = x \ k x = x)}" + proof (rule homeomorphism_moving_point [of "S - y`K" T "f(x i)" "y i"]) + show "openin (subtopology euclidean (affine hull (S - y ` K))) (S - y ` K)" + by (simp add: aff_eq openin_diff finite_imp_closedin image_subset_iff hull_inc insert xyS) + show "S - y ` K \ T" + using \S \ T\ by auto + show "T \ affine hull (S - y ` K)" + using insert by (simp add: aff_eq) + show "connected (S - y ` K)" + proof (rule connected_openin_diff_countable [OF \connected S\ ope]) + show "\ collinear S" + using collinear_aff_dim \2 \ aff_dim S\ by force + show "countable (y ` K)" + using countable_finite insert.hyps(1) by blast + qed + show "f (x i) \ S - y ` K" + apply (auto simp: f_in_S \x i \ S\) + by (metis feq homfg \x i \ S\ homeomorphism_def \S \ T\ \i \ K\ subsetCE xney xyS) + show "y i \ S - y ` K" + using insert.hyps xney by (auto simp: \y i \ S\) + qed blast + show ?case + proof (intro exI conjI) + show "homeomorphism T T (h \ f) (g \ k)" + using homfg homhk homeomorphism_compose by blast + show "\i \ insert i K. (h \ f) (x i) = y i" + using feq hk_sub by (auto simp: heq) + show "{x. \ ((h \ f) x = x \ (g \ k) x = x)} \ S" + using fg_sub hk_sub by force + have "bounded ({x. \(f x = x \ g x = x)} \ {x. \(h x = x \ k x = x)})" + using bo_fg bo_hk bounded_Un by blast + then show "bounded {x. \ ((h \ f) x = x \ (g \ k) x = x)}" + by (rule bounded_subset) auto + qed +qed + +proposition%unimportant homeomorphism_moving_points_exists: + fixes S :: "'a::euclidean_space set" + assumes 2: "2 \ DIM('a)" "open S" "connected S" "S \ T" "finite K" + and KS: "\i. i \ K \ x i \ S \ y i \ S" + and pw: "pairwise (\i j. (x i \ x j) \ (y i \ y j)) K" + and S: "S \ T" "T \ affine hull S" "connected S" + obtains f g where "homeomorphism T T f g" "\i. i \ K \ f(x i) = y i" + "{x. \ (f x = x \ g x = x)} \ S" "bounded {x. (\ (f x = x \ g x = x))}" +proof (cases "S = {}") + case True + then show ?thesis + using KS homeomorphism_ident that by fastforce +next + case False + then have affS: "affine hull S = UNIV" + by (simp add: affine_hull_open \open S\) + then have ope: "openin (subtopology euclidean (affine hull S)) S" + using \open S\ open_openin by auto + have "2 \ DIM('a)" by (rule 2) + also have "\ = aff_dim (UNIV :: 'a set)" + by simp + also have "\ \ aff_dim S" + by (metis aff_dim_UNIV aff_dim_affine_hull aff_dim_le_DIM affS) + finally have "2 \ aff_dim S" + by linarith + then show ?thesis + using homeomorphism_moving_points_exists_gen [OF \finite K\ KS pw _ ope S] that by fastforce +qed + +subsubsection%unimportant\The theorem \homeomorphism_grouping_points_exists\\ + +lemma homeomorphism_grouping_point_1: + fixes a::real and c::real + assumes "a < b" "c < d" + obtains f g where "homeomorphism (cbox a b) (cbox c d) f g" "f a = c" "f b = d" +proof - + define f where "f \ \x. ((d - c) / (b - a)) * x + (c - a * ((d - c) / (b - a)))" + have "\g. homeomorphism (cbox a b) (cbox c d) f g" + proof (rule homeomorphism_compact) + show "continuous_on (cbox a b) f" + apply (simp add: f_def) + apply (intro continuous_intros) + using assms by auto + have "f ` {a..b} = {c..d}" + unfolding f_def image_affinity_atLeastAtMost + using assms sum_sqs_eq by (auto simp: divide_simps algebra_simps) + then show "f ` cbox a b = cbox c d" + by auto + show "inj_on f (cbox a b)" + unfolding f_def inj_on_def using assms by auto + qed auto + then obtain g where "homeomorphism (cbox a b) (cbox c d) f g" .. + then show ?thesis + proof + show "f a = c" + by (simp add: f_def) + show "f b = d" + using assms sum_sqs_eq [of a b] by (auto simp: f_def divide_simps algebra_simps) + qed +qed + +lemma homeomorphism_grouping_point_2: + fixes a::real and w::real + assumes hom_ab: "homeomorphism (cbox a b) (cbox u v) f1 g1" + and hom_bc: "homeomorphism (cbox b c) (cbox v w) f2 g2" + and "b \ cbox a c" "v \ cbox u w" + and eq: "f1 a = u" "f1 b = v" "f2 b = v" "f2 c = w" + obtains f g where "homeomorphism (cbox a c) (cbox u w) f g" "f a = u" "f c = w" + "\x. x \ cbox a b \ f x = f1 x" "\x. x \ cbox b c \ f x = f2 x" +proof - + have le: "a \ b" "b \ c" "u \ v" "v \ w" + using assms by simp_all + then have ac: "cbox a c = cbox a b \ cbox b c" and uw: "cbox u w = cbox u v \ cbox v w" + by auto + define f where "f \ \x. if x \ b then f1 x else f2 x" + have "\g. homeomorphism (cbox a c) (cbox u w) f g" + proof (rule homeomorphism_compact) + have cf1: "continuous_on (cbox a b) f1" + using hom_ab homeomorphism_cont1 by blast + have cf2: "continuous_on (cbox b c) f2" + using hom_bc homeomorphism_cont1 by blast + show "continuous_on (cbox a c) f" + apply (simp add: f_def) + apply (rule continuous_on_cases_le [OF continuous_on_subset [OF cf1] continuous_on_subset [OF cf2]]) + using le eq apply (force simp: continuous_on_id)+ + done + have "f ` cbox a b = f1 ` cbox a b" "f ` cbox b c = f2 ` cbox b c" + unfolding f_def using eq by force+ + then show "f ` cbox a c = cbox u w" + apply (simp only: ac uw image_Un) + by (metis hom_ab hom_bc homeomorphism_def) + have neq12: "f1 x \ f2 y" if x: "a \ x" "x \ b" and y: "b < y" "y \ c" for x y + proof - + have "f1 x \ cbox u v" + by (metis hom_ab homeomorphism_def image_eqI mem_box_real(2) x) + moreover have "f2 y \ cbox v w" + by (metis (full_types) hom_bc homeomorphism_def image_subset_iff mem_box_real(2) not_le not_less_iff_gr_or_eq order_refl y) + moreover have "f2 y \ f2 b" + by (metis cancel_comm_monoid_add_class.diff_cancel diff_gt_0_iff_gt hom_bc homeomorphism_def le(2) less_imp_le less_numeral_extra(3) mem_box_real(2) order_refl y) + ultimately show ?thesis + using le eq by simp + qed + have "inj_on f1 (cbox a b)" + by (metis (full_types) hom_ab homeomorphism_def inj_onI) + moreover have "inj_on f2 (cbox b c)" + by (metis (full_types) hom_bc homeomorphism_def inj_onI) + ultimately show "inj_on f (cbox a c)" + apply (simp (no_asm) add: inj_on_def) + apply (simp add: f_def inj_on_eq_iff) + using neq12 apply force + done + qed auto + then obtain g where "homeomorphism (cbox a c) (cbox u w) f g" .. + then show ?thesis + apply (rule that) + using eq le by (auto simp: f_def) +qed + +lemma homeomorphism_grouping_point_3: + fixes a::real + assumes cbox_sub: "cbox c d \ box a b" "cbox u v \ box a b" + and box_ne: "box c d \ {}" "box u v \ {}" + obtains f g where "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" + "\x. x \ cbox c d \ f x \ cbox u v" +proof - + have less: "a < c" "a < u" "d < b" "v < b" "c < d" "u < v" "cbox c d \ {}" + using assms + by (simp_all add: cbox_sub subset_eq) + obtain f1 g1 where 1: "homeomorphism (cbox a c) (cbox a u) f1 g1" + and f1_eq: "f1 a = a" "f1 c = u" + using homeomorphism_grouping_point_1 [OF \a < c\ \a < u\] . + obtain f2 g2 where 2: "homeomorphism (cbox c d) (cbox u v) f2 g2" + and f2_eq: "f2 c = u" "f2 d = v" + using homeomorphism_grouping_point_1 [OF \c < d\ \u < v\] . + obtain f3 g3 where 3: "homeomorphism (cbox d b) (cbox v b) f3 g3" + and f3_eq: "f3 d = v" "f3 b = b" + using homeomorphism_grouping_point_1 [OF \d < b\ \v < b\] . + obtain f4 g4 where 4: "homeomorphism (cbox a d) (cbox a v) f4 g4" and "f4 a = a" "f4 d = v" + and f4_eq: "\x. x \ cbox a c \ f4 x = f1 x" "\x. x \ cbox c d \ f4 x = f2 x" + using homeomorphism_grouping_point_2 [OF 1 2] less by (auto simp: f1_eq f2_eq) + obtain f g where fg: "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" + and f_eq: "\x. x \ cbox a d \ f x = f4 x" "\x. x \ cbox d b \ f x = f3 x" + using homeomorphism_grouping_point_2 [OF 4 3] less by (auto simp: f4_eq f3_eq f2_eq f1_eq) + show ?thesis + apply (rule that [OF fg]) + using f4_eq f_eq homeomorphism_image1 [OF 2] + apply simp + by (metis atLeastAtMost_iff box_real(1) box_real(2) cbox_sub(1) greaterThanLessThan_iff imageI less_eq_real_def subset_eq) +qed + + +lemma homeomorphism_grouping_point_4: + fixes T :: "real set" + assumes "open U" "open S" "connected S" "U \ {}" "finite K" "K \ S" "U \ S" "S \ T" + obtains f g where "homeomorphism T T f g" + "\x. x \ K \ f x \ U" "{x. (\ (f x = x \ g x = x))} \ S" + "bounded {x. (\ (f x = x \ g x = x))}" +proof - + obtain c d where "box c d \ {}" "cbox c d \ U" + proof - + obtain u where "u \ U" + using \U \ {}\ by blast + then obtain e where "e > 0" "cball u e \ U" + using \open U\ open_contains_cball by blast + then show ?thesis + by (rule_tac c=u and d="u+e" in that) (auto simp: dist_norm subset_iff) + qed + have "compact K" + by (simp add: \finite K\ finite_imp_compact) + obtain a b where "box a b \ {}" "K \ cbox a b" "cbox a b \ S" + proof (cases "K = {}") + case True then show ?thesis + using \box c d \ {}\ \cbox c d \ U\ \U \ S\ that by blast + next + case False + then obtain a b where "a \ K" "b \ K" + and a: "\x. x \ K \ a \ x" and b: "\x. x \ K \ x \ b" + using compact_attains_inf compact_attains_sup by (metis \compact K\)+ + obtain e where "e > 0" "cball b e \ S" + using \open S\ open_contains_cball + by (metis \b \ K\ \K \ S\ subsetD) + show ?thesis + proof + show "box a (b + e) \ {}" + using \0 < e\ \b \ K\ a by force + show "K \ cbox a (b + e)" + using \0 < e\ a b by fastforce + have "a \ S" + using \a \ K\ assms(6) by blast + have "b + e \ S" + using \0 < e\ \cball b e \ S\ by (force simp: dist_norm) + show "cbox a (b + e) \ S" + using \a \ S\ \b + e \ S\ \connected S\ connected_contains_Icc by auto + qed + qed + obtain w z where "cbox w z \ S" and sub_wz: "cbox a b \ cbox c d \ box w z" + proof - + have "a \ S" "b \ S" + using \box a b \ {}\ \cbox a b \ S\ by auto + moreover have "c \ S" "d \ S" + using \box c d \ {}\ \cbox c d \ U\ \U \ S\ by force+ + ultimately have "min a c \ S" "max b d \ S" + by linarith+ + then obtain e1 e2 where "e1 > 0" "cball (min a c) e1 \ S" "e2 > 0" "cball (max b d) e2 \ S" + using \open S\ open_contains_cball by metis + then have *: "min a c - e1 \ S" "max b d + e2 \ S" + by (auto simp: dist_norm) + show ?thesis + proof + show "cbox (min a c - e1) (max b d+ e2) \ S" + using * \connected S\ connected_contains_Icc by auto + show "cbox a b \ cbox c d \ box (min a c - e1) (max b d + e2)" + using \0 < e1\ \0 < e2\ by auto + qed + qed + then + obtain f g where hom: "homeomorphism (cbox w z) (cbox w z) f g" + and "f w = w" "f z = z" + and fin: "\x. x \ cbox a b \ f x \ cbox c d" + using homeomorphism_grouping_point_3 [of a b w z c d] + using \box a b \ {}\ \box c d \ {}\ by blast + have contfg: "continuous_on (cbox w z) f" "continuous_on (cbox w z) g" + using hom homeomorphism_def by blast+ + define f' where "f' \ \x. if x \ cbox w z then f x else x" + define g' where "g' \ \x. if x \ cbox w z then g x else x" + show ?thesis + proof + have T: "cbox w z \ (T - box w z) = T" + using \cbox w z \ S\ \S \ T\ by auto + show "homeomorphism T T f' g'" + proof + have clo: "closedin (subtopology euclidean (cbox w z \ (T - box w z))) (T - box w z)" + by (metis Diff_Diff_Int Diff_subset T closedin_def open_box openin_open_Int topspace_euclidean_subtopology) + have "continuous_on (cbox w z \ (T - box w z)) f'" "continuous_on (cbox w z \ (T - box w z)) g'" + unfolding f'_def g'_def + apply (safe intro!: continuous_on_cases_local contfg continuous_on_id clo) + apply (simp_all add: closed_subset) + using \f w = w\ \f z = z\ apply force + by (metis \f w = w\ \f z = z\ hom homeomorphism_def less_eq_real_def mem_box_real(2)) + then show "continuous_on T f'" "continuous_on T g'" + by (simp_all only: T) + show "f' ` T \ T" + unfolding f'_def + by clarsimp (metis \cbox w z \ S\ \S \ T\ subsetD hom homeomorphism_def imageI mem_box_real(2)) + show "g' ` T \ T" + unfolding g'_def + by clarsimp (metis \cbox w z \ S\ \S \ T\ subsetD hom homeomorphism_def imageI mem_box_real(2)) + show "\x. x \ T \ g' (f' x) = x" + unfolding f'_def g'_def + using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] by fastforce + show "\y. y \ T \ f' (g' y) = y" + unfolding f'_def g'_def + using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] by fastforce + qed + show "\x. x \ K \ f' x \ U" + using fin sub_wz \K \ cbox a b\ \cbox c d \ U\ by (force simp: f'_def) + show "{x. \ (f' x = x \ g' x = x)} \ S" + using \cbox w z \ S\ by (auto simp: f'_def g'_def) + show "bounded {x. \ (f' x = x \ g' x = x)}" + apply (rule bounded_subset [of "cbox w z"]) + using bounded_cbox apply blast + apply (auto simp: f'_def g'_def) + done + qed +qed + +proposition%unimportant homeomorphism_grouping_points_exists: + fixes S :: "'a::euclidean_space set" + assumes "open U" "open S" "connected S" "U \ {}" "finite K" "K \ S" "U \ S" "S \ T" + obtains f g where "homeomorphism T T f g" "{x. (\ (f x = x \ g x = x))} \ S" + "bounded {x. (\ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" +proof (cases "2 \ DIM('a)") + case True + have TS: "T \ affine hull S" + using affine_hull_open assms by blast + have "infinite U" + using \open U\ \U \ {}\ finite_imp_not_open by blast + then obtain P where "P \ U" "finite P" "card K = card P" + using infinite_arbitrarily_large by metis + then obtain \ where \: "bij_betw \ K P" + using \finite K\ finite_same_card_bij by blast + obtain f g where "homeomorphism T T f g" "\i. i \ K \ f (id i) = \ i" "{x. \ (f x = x \ g x = x)} \ S" "bounded {x. \ (f x = x \ g x = x)}" + proof (rule homeomorphism_moving_points_exists [OF True \open S\ \connected S\ \S \ T\ \finite K\]) + show "\i. i \ K \ id i \ S \ \ i \ S" + using \P \ U\ \bij_betw \ K P\ \K \ S\ \U \ S\ bij_betwE by blast + show "pairwise (\i j. id i \ id j \ \ i \ \ j) K" + using \ by (auto simp: pairwise_def bij_betw_def inj_on_def) + qed (use affine_hull_open assms that in auto) + then show ?thesis + using \ \P \ U\ bij_betwE by (fastforce simp add: intro!: that) +next + case False + with DIM_positive have "DIM('a) = 1" + by (simp add: dual_order.antisym) + then obtain h::"'a \real" and j + where "linear h" "linear j" + and noh: "\x. norm(h x) = norm x" and noj: "\y. norm(j y) = norm y" + and hj: "\x. j(h x) = x" "\y. h(j y) = y" + and ranh: "surj h" + using isomorphisms_UNIV_UNIV + by (metis (mono_tags, hide_lams) DIM_real UNIV_eq_I range_eqI) + obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" + and f: "\x. x \ h ` K \ f x \ h ` U" + and sub: "{x. \ (f x = x \ g x = x)} \ h ` S" + and bou: "bounded {x. \ (f x = x \ g x = x)}" + apply (rule homeomorphism_grouping_point_4 [of "h ` U" "h ` S" "h ` K" "h ` T"]) + by (simp_all add: assms image_mono \linear h\ open_surjective_linear_image connected_linear_image ranh) + have jf: "j (f (h x)) = x \ f (h x) = h x" for x + by (metis hj) + have jg: "j (g (h x)) = x \ g (h x) = h x" for x + by (metis hj) + have cont_hj: "continuous_on X h" "continuous_on Y j" for X Y + by (simp_all add: \linear h\ \linear j\ linear_linear linear_continuous_on) + show ?thesis + proof + show "homeomorphism T T (j \ f \ h) (j \ g \ h)" + proof + show "continuous_on T (j \ f \ h)" "continuous_on T (j \ g \ h)" + using hom homeomorphism_def + by (blast intro: continuous_on_compose cont_hj)+ + show "(j \ f \ h) ` T \ T" "(j \ g \ h) ` T \ T" + by auto (metis (mono_tags, hide_lams) hj(1) hom homeomorphism_def imageE imageI)+ + show "\x. x \ T \ (j \ g \ h) ((j \ f \ h) x) = x" + using hj hom homeomorphism_apply1 by fastforce + show "\y. y \ T \ (j \ f \ h) ((j \ g \ h) y) = y" + using hj hom homeomorphism_apply2 by fastforce + qed + show "{x. \ ((j \ f \ h) x = x \ (j \ g \ h) x = x)} \ S" + apply (clarsimp simp: jf jg hj) + using sub hj + apply (drule_tac c="h x" in subsetD, force) + by (metis imageE) + have "bounded (j ` {x. (\ (f x = x \ g x = x))})" + by (rule bounded_linear_image [OF bou]) (use \linear j\ linear_conv_bounded_linear in auto) + moreover + have *: "{x. \((j \ f \ h) x = x \ (j \ g \ h) x = x)} = j ` {x. (\ (f x = x \ g x = x))}" + using hj by (auto simp: jf jg image_iff, metis+) + ultimately show "bounded {x. \ ((j \ f \ h) x = x \ (j \ g \ h) x = x)}" + by metis + show "\x. x \ K \ (j \ f \ h) x \ U" + using f hj by fastforce + qed +qed + + +proposition%unimportant homeomorphism_grouping_points_exists_gen: + fixes S :: "'a::euclidean_space set" + assumes opeU: "openin (subtopology euclidean S) U" + and opeS: "openin (subtopology euclidean (affine hull S)) S" + and "U \ {}" "finite K" "K \ S" and S: "S \ T" "T \ affine hull S" "connected S" + obtains f g where "homeomorphism T T f g" "{x. (\ (f x = x \ g x = x))} \ S" + "bounded {x. (\ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" +proof (cases "2 \ aff_dim S") + case True + have opeU': "openin (subtopology euclidean (affine hull S)) U" + using opeS opeU openin_trans by blast + obtain u where "u \ U" "u \ S" + using \U \ {}\ opeU openin_imp_subset by fastforce+ + have "infinite U" + apply (rule infinite_openin [OF opeU \u \ U\]) + apply (rule connected_imp_perfect_aff_dim [OF \connected S\ _ \u \ S\]) + using True apply simp + done + then obtain P where "P \ U" "finite P" "card K = card P" + using infinite_arbitrarily_large by metis + then obtain \ where \: "bij_betw \ K P" + using \finite K\ finite_same_card_bij by blast + have "\f g. homeomorphism T T f g \ (\i \ K. f(id i) = \ i) \ + {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" + proof (rule homeomorphism_moving_points_exists_gen [OF \finite K\ _ _ True opeS S]) + show "\i. i \ K \ id i \ S \ \ i \ S" + by (metis id_apply opeU openin_contains_cball subsetCE \P \ U\ \bij_betw \ K P\ \K \ S\ bij_betwE) + show "pairwise (\i j. id i \ id j \ \ i \ \ j) K" + using \ by (auto simp: pairwise_def bij_betw_def inj_on_def) + qed + then show ?thesis + using \ \P \ U\ bij_betwE by (fastforce simp add: intro!: that) +next + case False + with aff_dim_geq [of S] consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S = 1" by linarith + then show ?thesis + proof cases + assume "aff_dim S = -1" + then have "S = {}" + using aff_dim_empty by blast + then have "False" + using \U \ {}\ \K \ S\ openin_imp_subset [OF opeU] by blast + then show ?thesis .. + next + assume "aff_dim S = 0" + then obtain a where "S = {a}" + using aff_dim_eq_0 by blast + then have "K \ U" + using \U \ {}\ \K \ S\ openin_imp_subset [OF opeU] by blast + show ?thesis + apply (rule that [of id id]) + using \K \ U\ by (auto simp: continuous_on_id intro: homeomorphismI) + next + assume "aff_dim S = 1" + then have "affine hull S homeomorphic (UNIV :: real set)" + by (auto simp: homeomorphic_affine_sets) + then obtain h::"'a\real" and j where homhj: "homeomorphism (affine hull S) UNIV h j" + using homeomorphic_def by blast + then have h: "\x. x \ affine hull S \ j(h(x)) = x" and j: "\y. j y \ affine hull S \ h(j y) = y" + by (auto simp: homeomorphism_def) + have connh: "connected (h ` S)" + by (meson Topological_Spaces.connected_continuous_image \connected S\ homeomorphism_cont1 homeomorphism_of_subsets homhj hull_subset top_greatest) + have hUS: "h ` U \ h ` S" + by (meson homeomorphism_imp_open_map homeomorphism_of_subsets homhj hull_subset opeS opeU open_UNIV openin_open_eq) + have opn: "openin (subtopology euclidean (affine hull S)) U \ open (h ` U)" for U + using homeomorphism_imp_open_map [OF homhj] by simp + have "open (h ` U)" "open (h ` S)" + by (auto intro: opeS opeU openin_trans opn) + then obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" + and f: "\x. x \ h ` K \ f x \ h ` U" + and sub: "{x. \ (f x = x \ g x = x)} \ h ` S" + and bou: "bounded {x. \ (f x = x \ g x = x)}" + apply (rule homeomorphism_grouping_points_exists [of "h ` U" "h ` S" "h ` K" "h ` T"]) + using assms by (auto simp: connh hUS) + have jf: "\x. x \ affine hull S \ j (f (h x)) = x \ f (h x) = h x" + by (metis h j) + have jg: "\x. x \ affine hull S \ j (g (h x)) = x \ g (h x) = h x" + by (metis h j) + have cont_hj: "continuous_on T h" "continuous_on Y j" for Y + apply (rule continuous_on_subset [OF _ \T \ affine hull S\]) + using homeomorphism_def homhj apply blast + by (meson continuous_on_subset homeomorphism_def homhj top_greatest) + define f' where "f' \ \x. if x \ affine hull S then (j \ f \ h) x else x" + define g' where "g' \ \x. if x \ affine hull S then (j \ g \ h) x else x" + show ?thesis + proof + show "homeomorphism T T f' g'" + proof + have "continuous_on T (j \ f \ h)" + apply (intro continuous_on_compose cont_hj) + using hom homeomorphism_def by blast + then show "continuous_on T f'" + apply (rule continuous_on_eq) + using \T \ affine hull S\ f'_def by auto + have "continuous_on T (j \ g \ h)" + apply (intro continuous_on_compose cont_hj) + using hom homeomorphism_def by blast + then show "continuous_on T g'" + apply (rule continuous_on_eq) + using \T \ affine hull S\ g'_def by auto + show "f' ` T \ T" + proof (clarsimp simp: f'_def) + fix x assume "x \ T" + then have "f (h x) \ h ` T" + by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) + then show "j (f (h x)) \ T" + using \T \ affine hull S\ h by auto + qed + show "g' ` T \ T" + proof (clarsimp simp: g'_def) + fix x assume "x \ T" + then have "g (h x) \ h ` T" + by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) + then show "j (g (h x)) \ T" + using \T \ affine hull S\ h by auto + qed + show "\x. x \ T \ g' (f' x) = x" + using h j hom homeomorphism_apply1 by (fastforce simp add: f'_def g'_def) + show "\y. y \ T \ f' (g' y) = y" + using h j hom homeomorphism_apply2 by (fastforce simp add: f'_def g'_def) + qed + next + show "{x. \ (f' x = x \ g' x = x)} \ S" + apply (clarsimp simp: f'_def g'_def jf jg) + apply (rule imageE [OF subsetD [OF sub]], force) + by (metis h hull_inc) + next + have "compact (j ` closure {x. \ (f x = x \ g x = x)})" + using bou by (auto simp: compact_continuous_image cont_hj) + then have "bounded (j ` {x. \ (f x = x \ g x = x)})" + by (rule bounded_closure_image [OF compact_imp_bounded]) + moreover + have *: "{x \ affine hull S. j (f (h x)) \ x \ j (g (h x)) \ x} = j ` {x. (\ (f x = x \ g x = x))}" + using h j by (auto simp: image_iff; metis) + ultimately have "bounded {x \ affine hull S. j (f (h x)) \ x \ j (g (h x)) \ x}" + by metis + then show "bounded {x. \ (f' x = x \ g' x = x)}" + by (simp add: f'_def g'_def Collect_mono bounded_subset) + next + show "f' x \ U" if "x \ K" for x + proof - + have "U \ S" + using opeU openin_imp_subset by blast + then have "j (f (h x)) \ U" + using f h hull_subset that by fastforce + then show "f' x \ U" + using \K \ S\ S f'_def that by auto + qed + qed + qed +qed + + +subsection\Nullhomotopic mappings\ + +text\ A mapping out of a sphere is nullhomotopic iff it extends to the ball. +This even works out in the degenerate cases when the radius is \\\ 0, and +we also don't need to explicitly assume continuity since it's already implicit +in both sides of the equivalence.\ + +lemma nullhomotopic_from_lemma: + assumes contg: "continuous_on (cball a r - {a}) g" + and fa: "\e. 0 < e + \ \d. 0 < d \ (\x. x \ a \ norm(x - a) < d \ norm(g x - f a) < e)" + and r: "\x. x \ cball a r \ x \ a \ f x = g x" + shows "continuous_on (cball a r) f" +proof (clarsimp simp: continuous_on_eq_continuous_within Ball_def) + fix x + assume x: "dist a x \ r" + show "continuous (at x within cball a r) f" + proof (cases "x=a") + case True + then show ?thesis + by (metis continuous_within_eps_delta fa dist_norm dist_self r) + next + case False + show ?thesis + proof (rule continuous_transform_within [where f=g and d = "norm(x-a)"]) + have "\d>0. \x'\cball a r. + dist x' x < d \ dist (g x') (g x) < e" if "e>0" for e + proof - + obtain d where "d > 0" + and d: "\x'. \dist x' a \ r; x' \ a; dist x' x < d\ \ + dist (g x') (g x) < e" + using contg False x \e>0\ + unfolding continuous_on_iff by (fastforce simp add: dist_commute intro: that) + show ?thesis + using \d > 0\ \x \ a\ + by (rule_tac x="min d (norm(x - a))" in exI) + (auto simp: dist_commute dist_norm [symmetric] intro!: d) + qed + then show "continuous (at x within cball a r) g" + using contg False by (auto simp: continuous_within_eps_delta) + show "0 < norm (x - a)" + using False by force + show "x \ cball a r" + by (simp add: x) + show "\x'. \x' \ cball a r; dist x' x < norm (x - a)\ + \ g x' = f x'" + by (metis dist_commute dist_norm less_le r) + qed + qed +qed + +proposition nullhomotopic_from_sphere_extension: + fixes f :: "'M::euclidean_space \ 'a::real_normed_vector" + shows "(\c. homotopic_with (\x. True) (sphere a r) S f (\x. c)) \ + (\g. continuous_on (cball a r) g \ g ` (cball a r) \ S \ + (\x \ sphere a r. g x = f x))" + (is "?lhs = ?rhs") +proof (cases r "0::real" rule: linorder_cases) + case equal + then show ?thesis + apply (auto simp: homotopic_with) + apply (rule_tac x="\x. h (0, a)" in exI) + apply (fastforce simp add:) + using continuous_on_const by blast +next + case greater + let ?P = "continuous_on {x. norm(x - a) = r} f \ f ` {x. norm(x - a) = r} \ S" + have ?P if ?lhs using that + proof + fix c + assume c: "homotopic_with (\x. True) (sphere a r) S f (\x. c)" + then have contf: "continuous_on (sphere a r) f" and fim: "f ` sphere a r \ S" + by (auto simp: homotopic_with_imp_subset1 homotopic_with_imp_continuous) + show ?P + using contf fim by (auto simp: sphere_def dist_norm norm_minus_commute) + qed + moreover have ?P if ?rhs using that + proof + fix g + assume g: "continuous_on (cball a r) g \ g ` cball a r \ S \ (\xa\sphere a r. g xa = f xa)" + then + show ?P + apply (safe elim!: continuous_on_eq [OF continuous_on_subset]) + apply (auto simp: dist_norm norm_minus_commute) + by (metis dist_norm image_subset_iff mem_sphere norm_minus_commute sphere_cball subsetCE) + qed + moreover have ?thesis if ?P + proof + assume ?lhs + then obtain c where "homotopic_with (\x. True) (sphere a r) S (\x. c) f" + using homotopic_with_sym by blast + then obtain h where conth: "continuous_on ({0..1::real} \ sphere a r) h" + and him: "h ` ({0..1} \ sphere a r) \ S" + and h: "\x. h(0, x) = c" "\x. h(1, x) = f x" + by (auto simp: homotopic_with_def) + obtain b1::'M where "b1 \ Basis" + using SOME_Basis by auto + have "c \ S" + apply (rule him [THEN subsetD]) + apply (rule_tac x = "(0, a + r *\<^sub>R b1)" in image_eqI) + using h greater \b1 \ Basis\ + apply (auto simp: dist_norm) + done + have uconth: "uniformly_continuous_on ({0..1::real} \ (sphere a r)) h" + by (force intro: compact_Times conth compact_uniformly_continuous) + let ?g = "\x. h (norm (x - a)/r, + a + (if x = a then r *\<^sub>R b1 else (r / norm(x - a)) *\<^sub>R (x - a)))" + let ?g' = "\x. h (norm (x - a)/r, a + (r / norm(x - a)) *\<^sub>R (x - a))" + show ?rhs + proof (intro exI conjI) + have "continuous_on (cball a r - {a}) ?g'" + apply (rule continuous_on_compose2 [OF conth]) + apply (intro continuous_intros) + using greater apply (auto simp: dist_norm norm_minus_commute) + done + then show "continuous_on (cball a r) ?g" + proof (rule nullhomotopic_from_lemma) + show "\d>0. \x. x \ a \ norm (x - a) < d \ norm (?g' x - ?g a) < e" if "0 < e" for e + proof - + obtain d where "0 < d" + and d: "\x x'. \x \ {0..1} \ sphere a r; x' \ {0..1} \ sphere a r; dist x' x < d\ + \ dist (h x') (h x) < e" + using uniformly_continuous_onE [OF uconth \0 < e\] by auto + have *: "norm (h (norm (x - a) / r, + a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + r *\<^sub>R b1)) < e" + if "x \ a" "norm (x - a) < r" "norm (x - a) < d * r" for x + proof - + have "norm (h (norm (x - a) / r, a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + r *\<^sub>R b1)) = + norm (h (norm (x - a) / r, a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + (r / norm (x - a)) *\<^sub>R (x - a)))" + by (simp add: h) + also have "\ < e" + apply (rule d [unfolded dist_norm]) + using greater \0 < d\ \b1 \ Basis\ that + by (auto simp: dist_norm divide_simps) + finally show ?thesis . + qed + show ?thesis + apply (rule_tac x = "min r (d * r)" in exI) + using greater \0 < d\ by (auto simp: *) + qed + show "\x. x \ cball a r \ x \ a \ ?g x = ?g' x" + by auto + qed + next + show "?g ` cball a r \ S" + using greater him \c \ S\ + by (force simp: h dist_norm norm_minus_commute) + next + show "\x\sphere a r. ?g x = f x" + using greater by (auto simp: h dist_norm norm_minus_commute) + qed + next + assume ?rhs + then obtain g where contg: "continuous_on (cball a r) g" + and gim: "g ` cball a r \ S" + and gf: "\x \ sphere a r. g x = f x" + by auto + let ?h = "\y. g (a + (fst y) *\<^sub>R (snd y - a))" + have "continuous_on ({0..1} \ sphere a r) ?h" + apply (rule continuous_on_compose2 [OF contg]) + apply (intro continuous_intros) + apply (auto simp: dist_norm norm_minus_commute mult_left_le_one_le) + done + moreover + have "?h ` ({0..1} \ sphere a r) \ S" + by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gim [THEN subsetD]) + moreover + have "\x\sphere a r. ?h (0, x) = g a" "\x\sphere a r. ?h (1, x) = f x" + by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gf) + ultimately + show ?lhs + apply (subst homotopic_with_sym) + apply (rule_tac x="g a" in exI) + apply (auto simp: homotopic_with) + done + qed + ultimately + show ?thesis by meson +qed simp + +end \ No newline at end of file diff -r 3f7d8e05e0f2 -r 19d8a59481db src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy Mon Jan 07 14:06:54 2019 +0100 +++ b/src/HOL/Analysis/Path_Connected.thy Mon Jan 07 14:57:45 2019 +0100 @@ -2,10 +2,10 @@ Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light *) -section \Continuous Paths\ +section \Path-Connectedness\ theory Path_Connected - imports Continuous_Extension Continuum_Not_Denumerable + imports Starlike begin subsection \Paths and Arcs\ @@ -295,7 +295,7 @@ qed -section%unimportant \Path Images\ +subsection%unimportant \Path Images\ lemma bounded_path_image: "path g \ bounded(path_image g)" by (simp add: compact_imp_bounded compact_path_image) @@ -1421,8 +1421,6 @@ by (rule_tac x="e/2" in exI) auto qed -section "Path-Connectedness" (* TODO: separate theory? *) - subsection \Path component\ text \Original formalization by Tom Hales\ @@ -2531,7 +2529,7 @@ by (meson cobounded_unique_unbounded_components connected_eq_connected_components_eq assms) -section\The \inside\ and \outside\ of a Set\ +subsection\The \inside\ and \outside\ of a Set\ text%important\The inside comprises the points in a bounded connected component of the set's complement. The outside comprises the points in unbounded connected component of the complement.\ @@ -3386,5156 +3384,4 @@ by (metis dw_le norm_minus_commute not_less order_trans rle wy) qed - -section \Homotopy of Maps\ (* TODO separate theory? *) - - -definition%important homotopic_with :: - "[('a::topological_space \ 'b::topological_space) \ bool, 'a set, 'b set, 'a \ 'b, 'a \ 'b] \ bool" -where - "homotopic_with P X Y p q \ - (\h:: real \ 'a \ 'b. - continuous_on ({0..1} \ X) h \ - h ` ({0..1} \ X) \ Y \ - (\x. h(0, x) = p x) \ - (\x. h(1, x) = q x) \ - (\t \ {0..1}. P(\x. h(t, x))))" - -text\\p\, \q\ are functions \X \ Y\, and the property \P\ restricts all intermediate maps. -We often just want to require that \P\ fixes some subset, but to include the case of a loop homotopy, -it is convenient to have a general property \P\.\ - -text \We often want to just localize the ending function equality or whatever.\ -text%important \%whitespace\ -proposition homotopic_with: - fixes X :: "'a::topological_space set" and Y :: "'b::topological_space set" - assumes "\h k. (\x. x \ X \ h x = k x) \ (P h \ P k)" - shows "homotopic_with P X Y p q \ - (\h :: real \ 'a \ 'b. - continuous_on ({0..1} \ X) h \ - h ` ({0..1} \ X) \ Y \ - (\x \ X. h(0,x) = p x) \ - (\x \ X. h(1,x) = q x) \ - (\t \ {0..1}. P(\x. h(t, x))))" - unfolding homotopic_with_def - apply (rule iffI, blast, clarify) - apply (rule_tac x="\(u,v). if v \ X then h(u,v) else if u = 0 then p v else q v" in exI) - apply auto - apply (force elim: continuous_on_eq) - apply (drule_tac x=t in bspec, force) - apply (subst assms; simp) - done - -proposition homotopic_with_eq: - assumes h: "homotopic_with P X Y f g" - and f': "\x. x \ X \ f' x = f x" - and g': "\x. x \ X \ g' x = g x" - and P: "(\h k. (\x. x \ X \ h x = k x) \ (P h \ P k))" - shows "homotopic_with P X Y f' g'" - using h unfolding homotopic_with_def - apply safe - apply (rule_tac x="\(u,v). if v \ X then h(u,v) else if u = 0 then f' v else g' v" in exI) - apply (simp add: f' g', safe) - apply (fastforce intro: continuous_on_eq, fastforce) - apply (subst P; fastforce) - done - -proposition homotopic_with_equal: - assumes contf: "continuous_on X f" and fXY: "f ` X \ Y" - and gf: "\x. x \ X \ g x = f x" - and P: "P f" "P g" - shows "homotopic_with P X Y f g" - unfolding homotopic_with_def - apply (rule_tac x="\(u,v). if u = 1 then g v else f v" in exI) - using assms - apply (intro conjI) - apply (rule continuous_on_eq [where f = "f \ snd"]) - apply (rule continuous_intros | force)+ - apply clarify - apply (case_tac "t=1"; force) - done - - -lemma image_Pair_const: "(\x. (x, c)) ` A = A \ {c}" - by auto - -lemma homotopic_constant_maps: - "homotopic_with (\x. True) s t (\x. a) (\x. b) \ s = {} \ path_component t a b" -proof (cases "s = {} \ t = {}") - case True with continuous_on_const show ?thesis - by (auto simp: homotopic_with path_component_def) -next - case False - then obtain c where "c \ s" by blast - show ?thesis - proof - assume "homotopic_with (\x. True) s t (\x. a) (\x. b)" - then obtain h :: "real \ 'a \ 'b" - where conth: "continuous_on ({0..1} \ s) h" - and h: "h ` ({0..1} \ s) \ t" "(\x\s. h (0, x) = a)" "(\x\s. h (1, x) = b)" - by (auto simp: homotopic_with) - have "continuous_on {0..1} (h \ (\t. (t, c)))" - apply (rule continuous_intros conth | simp add: image_Pair_const)+ - apply (blast intro: \c \ s\ continuous_on_subset [OF conth]) - done - with \c \ s\ h show "s = {} \ path_component t a b" - apply (simp_all add: homotopic_with path_component_def, auto) - apply (drule_tac x="h \ (\t. (t, c))" in spec) - apply (auto simp: pathstart_def pathfinish_def path_image_def path_def) - done - next - assume "s = {} \ path_component t a b" - with False show "homotopic_with (\x. True) s t (\x. a) (\x. b)" - apply (clarsimp simp: homotopic_with path_component_def pathstart_def pathfinish_def path_image_def path_def) - apply (rule_tac x="g \ fst" in exI) - apply (rule conjI continuous_intros | force)+ - done - qed -qed - - -subsection%unimportant\Trivial properties\ - -lemma homotopic_with_imp_property: "homotopic_with P X Y f g \ P f \ P g" - unfolding homotopic_with_def Ball_def - apply clarify - apply (frule_tac x=0 in spec) - apply (drule_tac x=1 in spec, auto) - done - -lemma continuous_on_o_Pair: "\continuous_on (T \ X) h; t \ T\ \ continuous_on X (h \ Pair t)" - by (fast intro: continuous_intros elim!: continuous_on_subset) - -lemma homotopic_with_imp_continuous: - assumes "homotopic_with P X Y f g" - shows "continuous_on X f \ continuous_on X g" -proof - - obtain h :: "real \ 'a \ 'b" - where conth: "continuous_on ({0..1} \ X) h" - and h: "\x. h (0, x) = f x" "\x. h (1, x) = g x" - using assms by (auto simp: homotopic_with_def) - have *: "t \ {0..1} \ continuous_on X (h \ (\x. (t,x)))" for t - by (rule continuous_intros continuous_on_subset [OF conth] | force)+ - show ?thesis - using h *[of 0] *[of 1] by auto -qed - -proposition homotopic_with_imp_subset1: - "homotopic_with P X Y f g \ f ` X \ Y" - by (simp add: homotopic_with_def image_subset_iff) (metis atLeastAtMost_iff order_refl zero_le_one) - -proposition homotopic_with_imp_subset2: - "homotopic_with P X Y f g \ g ` X \ Y" - by (simp add: homotopic_with_def image_subset_iff) (metis atLeastAtMost_iff order_refl zero_le_one) - -proposition homotopic_with_mono: - assumes hom: "homotopic_with P X Y f g" - and Q: "\h. \continuous_on X h; image h X \ Y \ P h\ \ Q h" - shows "homotopic_with Q X Y f g" - using hom - apply (simp add: homotopic_with_def) - apply (erule ex_forward) - apply (force simp: intro!: Q dest: continuous_on_o_Pair) - done - -proposition homotopic_with_subset_left: - "\homotopic_with P X Y f g; Z \ X\ \ homotopic_with P Z Y f g" - apply (simp add: homotopic_with_def) - apply (fast elim!: continuous_on_subset ex_forward) - done - -proposition homotopic_with_subset_right: - "\homotopic_with P X Y f g; Y \ Z\ \ homotopic_with P X Z f g" - apply (simp add: homotopic_with_def) - apply (fast elim!: continuous_on_subset ex_forward) - done - -proposition homotopic_with_compose_continuous_right: - "\homotopic_with (\f. p (f \ h)) X Y f g; continuous_on W h; h ` W \ X\ - \ homotopic_with p W Y (f \ h) (g \ h)" - apply (clarsimp simp add: homotopic_with_def) - apply (rename_tac k) - apply (rule_tac x="k \ (\y. (fst y, h (snd y)))" in exI) - apply (rule conjI continuous_intros continuous_on_compose [where f=snd and g=h, unfolded o_def] | simp)+ - apply (erule continuous_on_subset) - apply (fastforce simp: o_def)+ - done - -proposition homotopic_compose_continuous_right: - "\homotopic_with (\f. True) X Y f g; continuous_on W h; h ` W \ X\ - \ homotopic_with (\f. True) W Y (f \ h) (g \ h)" - using homotopic_with_compose_continuous_right by fastforce - -proposition homotopic_with_compose_continuous_left: - "\homotopic_with (\f. p (h \ f)) X Y f g; continuous_on Y h; h ` Y \ Z\ - \ homotopic_with p X Z (h \ f) (h \ g)" - apply (clarsimp simp add: homotopic_with_def) - apply (rename_tac k) - apply (rule_tac x="h \ k" in exI) - apply (rule conjI continuous_intros continuous_on_compose [where f=snd and g=h, unfolded o_def] | simp)+ - apply (erule continuous_on_subset) - apply (fastforce simp: o_def)+ - done - -proposition homotopic_compose_continuous_left: - "\homotopic_with (\_. True) X Y f g; - continuous_on Y h; h ` Y \ Z\ - \ homotopic_with (\f. True) X Z (h \ f) (h \ g)" - using homotopic_with_compose_continuous_left by fastforce - -proposition homotopic_with_Pair: - assumes hom: "homotopic_with p s t f g" "homotopic_with p' s' t' f' g'" - and q: "\f g. \p f; p' g\ \ q(\(x,y). (f x, g y))" - shows "homotopic_with q (s \ s') (t \ t') - (\(x,y). (f x, f' y)) (\(x,y). (g x, g' y))" - using hom - apply (clarsimp simp add: homotopic_with_def) - apply (rename_tac k k') - apply (rule_tac x="\z. ((k \ (\x. (fst x, fst (snd x)))) z, (k' \ (\x. (fst x, snd (snd x)))) z)" in exI) - apply (rule conjI continuous_intros | erule continuous_on_subset | clarsimp)+ - apply (auto intro!: q [unfolded case_prod_unfold]) - done - -lemma homotopic_on_empty [simp]: "homotopic_with (\x. True) {} t f g" - by (metis continuous_on_def empty_iff homotopic_with_equal image_subset_iff) - - -text\Homotopy with P is an equivalence relation (on continuous functions mapping X into Y that satisfy P, - though this only affects reflexivity.\ - - -proposition homotopic_with_refl: - "homotopic_with P X Y f f \ continuous_on X f \ image f X \ Y \ P f" - apply (rule iffI) - using homotopic_with_imp_continuous homotopic_with_imp_property homotopic_with_imp_subset2 apply blast - apply (simp add: homotopic_with_def) - apply (rule_tac x="f \ snd" in exI) - apply (rule conjI continuous_intros | force)+ - done - -lemma homotopic_with_symD: - fixes X :: "'a::real_normed_vector set" - assumes "homotopic_with P X Y f g" - shows "homotopic_with P X Y g f" - using assms - apply (clarsimp simp add: homotopic_with_def) - apply (rename_tac h) - apply (rule_tac x="h \ (\y. (1 - fst y, snd y))" in exI) - apply (rule conjI continuous_intros | erule continuous_on_subset | force simp: image_subset_iff)+ - done - -proposition homotopic_with_sym: - fixes X :: "'a::real_normed_vector set" - shows "homotopic_with P X Y f g \ homotopic_with P X Y g f" - using homotopic_with_symD by blast - -lemma split_01: "{0..1::real} = {0..1/2} \ {1/2..1}" - by force - -lemma split_01_prod: "{0..1::real} \ X = ({0..1/2} \ X) \ ({1/2..1} \ X)" - by force - -proposition homotopic_with_trans: - fixes X :: "'a::real_normed_vector set" - assumes "homotopic_with P X Y f g" and "homotopic_with P X Y g h" - shows "homotopic_with P X Y f h" -proof - - have clo1: "closedin (subtopology euclidean ({0..1/2} \ X \ {1/2..1} \ X)) ({0..1/2::real} \ X)" - apply (simp add: closedin_closed split_01_prod [symmetric]) - apply (rule_tac x="{0..1/2} \ UNIV" in exI) - apply (force simp: closed_Times) - done - have clo2: "closedin (subtopology euclidean ({0..1/2} \ X \ {1/2..1} \ X)) ({1/2..1::real} \ X)" - apply (simp add: closedin_closed split_01_prod [symmetric]) - apply (rule_tac x="{1/2..1} \ UNIV" in exI) - apply (force simp: closed_Times) - done - { fix k1 k2:: "real \ 'a \ 'b" - assume cont: "continuous_on ({0..1} \ X) k1" "continuous_on ({0..1} \ X) k2" - and Y: "k1 ` ({0..1} \ X) \ Y" "k2 ` ({0..1} \ X) \ Y" - and geq: "\x. k1 (1, x) = g x" "\x. k2 (0, x) = g x" - and k12: "\x. k1 (0, x) = f x" "\x. k2 (1, x) = h x" - and P: "\t\{0..1}. P (\x. k1 (t, x))" "\t\{0..1}. P (\x. k2 (t, x))" - define k where "k y = - (if fst y \ 1 / 2 - then (k1 \ (\x. (2 *\<^sub>R fst x, snd x))) y - else (k2 \ (\x. (2 *\<^sub>R fst x -1, snd x))) y)" for y - have keq: "k1 (2 * u, v) = k2 (2 * u - 1, v)" if "u = 1/2" for u v - by (simp add: geq that) - have "continuous_on ({0..1} \ X) k" - using cont - apply (simp add: split_01_prod k_def) - apply (rule clo1 clo2 continuous_on_cases_local continuous_intros | erule continuous_on_subset | simp add: linear image_subset_iff)+ - apply (force simp: keq) - done - moreover have "k ` ({0..1} \ X) \ Y" - using Y by (force simp: k_def) - moreover have "\x. k (0, x) = f x" - by (simp add: k_def k12) - moreover have "(\x. k (1, x) = h x)" - by (simp add: k_def k12) - moreover have "\t\{0..1}. P (\x. k (t, x))" - using P - apply (clarsimp simp add: k_def) - apply (case_tac "t \ 1/2", auto) - done - ultimately have *: "\k :: real \ 'a \ 'b. - continuous_on ({0..1} \ X) k \ k ` ({0..1} \ X) \ Y \ - (\x. k (0, x) = f x) \ (\x. k (1, x) = h x) \ (\t\{0..1}. P (\x. k (t, x)))" - by blast - } note * = this - show ?thesis - using assms by (auto intro: * simp add: homotopic_with_def) -qed - -proposition homotopic_compose: - fixes s :: "'a::real_normed_vector set" - shows "\homotopic_with (\x. True) s t f f'; homotopic_with (\x. True) t u g g'\ - \ homotopic_with (\x. True) s u (g \ f) (g' \ f')" - apply (rule homotopic_with_trans [where g = "g \ f'"]) - apply (metis homotopic_compose_continuous_left homotopic_with_imp_continuous homotopic_with_imp_subset1) - by (metis homotopic_compose_continuous_right homotopic_with_imp_continuous homotopic_with_imp_subset2) - - -text\Homotopic triviality implicitly incorporates path-connectedness.\ -lemma homotopic_triviality: - fixes S :: "'a::real_normed_vector set" - shows "(\f g. continuous_on S f \ f ` S \ T \ - continuous_on S g \ g ` S \ T - \ homotopic_with (\x. True) S T f g) \ - (S = {} \ path_connected T) \ - (\f. continuous_on S f \ f ` S \ T \ (\c. homotopic_with (\x. True) S T f (\x. c)))" - (is "?lhs = ?rhs") -proof (cases "S = {} \ T = {}") - case True then show ?thesis by auto -next - case False show ?thesis - proof - assume LHS [rule_format]: ?lhs - have pab: "path_component T a b" if "a \ T" "b \ T" for a b - proof - - have "homotopic_with (\x. True) S T (\x. a) (\x. b)" - by (simp add: LHS continuous_on_const image_subset_iff that) - then show ?thesis - using False homotopic_constant_maps by blast - qed - moreover - have "\c. homotopic_with (\x. True) S T f (\x. c)" if "continuous_on S f" "f ` S \ T" for f - by (metis (full_types) False LHS equals0I homotopic_constant_maps homotopic_with_imp_continuous homotopic_with_imp_subset2 pab that) - ultimately show ?rhs - by (simp add: path_connected_component) - next - assume RHS: ?rhs - with False have T: "path_connected T" - by blast - show ?lhs - proof clarify - fix f g - assume "continuous_on S f" "f ` S \ T" "continuous_on S g" "g ` S \ T" - obtain c d where c: "homotopic_with (\x. True) S T f (\x. c)" and d: "homotopic_with (\x. True) S T g (\x. d)" - using False \continuous_on S f\ \f ` S \ T\ RHS \continuous_on S g\ \g ` S \ T\ by blast - then have "c \ T" "d \ T" - using False homotopic_with_imp_subset2 by fastforce+ - with T have "path_component T c d" - using path_connected_component by blast - then have "homotopic_with (\x. True) S T (\x. c) (\x. d)" - by (simp add: homotopic_constant_maps) - with c d show "homotopic_with (\x. True) S T f g" - by (meson homotopic_with_symD homotopic_with_trans) - qed - qed -qed - - -subsection\Homotopy of paths, maintaining the same endpoints\ - - -definition%important homotopic_paths :: "['a set, real \ 'a, real \ 'a::topological_space] \ bool" - where - "homotopic_paths s p q \ - homotopic_with (\r. pathstart r = pathstart p \ pathfinish r = pathfinish p) {0..1} s p q" - -lemma homotopic_paths: - "homotopic_paths s p q \ - (\h. continuous_on ({0..1} \ {0..1}) h \ - h ` ({0..1} \ {0..1}) \ s \ - (\x \ {0..1}. h(0,x) = p x) \ - (\x \ {0..1}. h(1,x) = q x) \ - (\t \ {0..1::real}. pathstart(h \ Pair t) = pathstart p \ - pathfinish(h \ Pair t) = pathfinish p))" - by (auto simp: homotopic_paths_def homotopic_with pathstart_def pathfinish_def) - -proposition homotopic_paths_imp_pathstart: - "homotopic_paths s p q \ pathstart p = pathstart q" - by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) - -proposition homotopic_paths_imp_pathfinish: - "homotopic_paths s p q \ pathfinish p = pathfinish q" - by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) - -lemma homotopic_paths_imp_path: - "homotopic_paths s p q \ path p \ path q" - using homotopic_paths_def homotopic_with_imp_continuous path_def by blast - -lemma homotopic_paths_imp_subset: - "homotopic_paths s p q \ path_image p \ s \ path_image q \ s" - by (simp add: homotopic_paths_def homotopic_with_imp_subset1 homotopic_with_imp_subset2 path_image_def) - -proposition homotopic_paths_refl [simp]: "homotopic_paths s p p \ path p \ path_image p \ s" -by (simp add: homotopic_paths_def homotopic_with_refl path_def path_image_def) - -proposition homotopic_paths_sym: "homotopic_paths s p q \ homotopic_paths s q p" - by (metis (mono_tags) homotopic_paths_def homotopic_paths_imp_pathfinish homotopic_paths_imp_pathstart homotopic_with_symD) - -proposition homotopic_paths_sym_eq: "homotopic_paths s p q \ homotopic_paths s q p" - by (metis homotopic_paths_sym) - -proposition homotopic_paths_trans [trans]: - "\homotopic_paths s p q; homotopic_paths s q r\ \ homotopic_paths s p r" - apply (simp add: homotopic_paths_def) - apply (rule homotopic_with_trans, assumption) - by (metis (mono_tags, lifting) homotopic_with_imp_property homotopic_with_mono) - -proposition homotopic_paths_eq: - "\path p; path_image p \ s; \t. t \ {0..1} \ p t = q t\ \ homotopic_paths s p q" - apply (simp add: homotopic_paths_def) - apply (rule homotopic_with_eq) - apply (auto simp: path_def homotopic_with_refl pathstart_def pathfinish_def path_image_def elim: continuous_on_eq) - done - -proposition homotopic_paths_reparametrize: - assumes "path p" - and pips: "path_image p \ s" - and contf: "continuous_on {0..1} f" - and f01:"f ` {0..1} \ {0..1}" - and [simp]: "f(0) = 0" "f(1) = 1" - and q: "\t. t \ {0..1} \ q(t) = p(f t)" - shows "homotopic_paths s p q" -proof - - have contp: "continuous_on {0..1} p" - by (metis \path p\ path_def) - then have "continuous_on {0..1} (p \ f)" - using contf continuous_on_compose continuous_on_subset f01 by blast - then have "path q" - by (simp add: path_def) (metis q continuous_on_cong) - have piqs: "path_image q \ s" - by (metis (no_types, hide_lams) pips f01 image_subset_iff path_image_def q) - have fb0: "\a b. \0 \ a; a \ 1; 0 \ b; b \ 1\ \ 0 \ (1 - a) * f b + a * b" - using f01 by force - have fb1: "\0 \ a; a \ 1; 0 \ b; b \ 1\ \ (1 - a) * f b + a * b \ 1" for a b - using f01 [THEN subsetD, of "f b"] by (simp add: convex_bound_le) - have "homotopic_paths s q p" - proof (rule homotopic_paths_trans) - show "homotopic_paths s q (p \ f)" - using q by (force intro: homotopic_paths_eq [OF \path q\ piqs]) - next - show "homotopic_paths s (p \ f) p" - apply (simp add: homotopic_paths_def homotopic_with_def) - apply (rule_tac x="p \ (\y. (1 - (fst y)) *\<^sub>R ((f \ snd) y) + (fst y) *\<^sub>R snd y)" in exI) - apply (rule conjI contf continuous_intros continuous_on_subset [OF contp] | simp)+ - using pips [unfolded path_image_def] - apply (auto simp: fb0 fb1 pathstart_def pathfinish_def) - done - qed - then show ?thesis - by (simp add: homotopic_paths_sym) -qed - -lemma homotopic_paths_subset: "\homotopic_paths s p q; s \ t\ \ homotopic_paths t p q" - using homotopic_paths_def homotopic_with_subset_right by blast - - -text\ A slightly ad-hoc but useful lemma in constructing homotopies.\ -lemma homotopic_join_lemma: - fixes q :: "[real,real] \ 'a::topological_space" - assumes p: "continuous_on ({0..1} \ {0..1}) (\y. p (fst y) (snd y))" - and q: "continuous_on ({0..1} \ {0..1}) (\y. q (fst y) (snd y))" - and pf: "\t. t \ {0..1} \ pathfinish(p t) = pathstart(q t)" - shows "continuous_on ({0..1} \ {0..1}) (\y. (p(fst y) +++ q(fst y)) (snd y))" -proof - - have 1: "(\y. p (fst y) (2 * snd y)) = (\y. p (fst y) (snd y)) \ (\y. (fst y, 2 * snd y))" - by (rule ext) (simp) - have 2: "(\y. q (fst y) (2 * snd y - 1)) = (\y. q (fst y) (snd y)) \ (\y. (fst y, 2 * snd y - 1))" - by (rule ext) (simp) - show ?thesis - apply (simp add: joinpaths_def) - apply (rule continuous_on_cases_le) - apply (simp_all only: 1 2) - apply (rule continuous_intros continuous_on_subset [OF p] continuous_on_subset [OF q] | force)+ - using pf - apply (auto simp: mult.commute pathstart_def pathfinish_def) - done -qed - -text\ Congruence properties of homotopy w.r.t. path-combining operations.\ - -lemma homotopic_paths_reversepath_D: - assumes "homotopic_paths s p q" - shows "homotopic_paths s (reversepath p) (reversepath q)" - using assms - apply (simp add: homotopic_paths_def homotopic_with_def, clarify) - apply (rule_tac x="h \ (\x. (fst x, 1 - snd x))" in exI) - apply (rule conjI continuous_intros)+ - apply (auto simp: reversepath_def pathstart_def pathfinish_def elim!: continuous_on_subset) - done - -proposition homotopic_paths_reversepath: - "homotopic_paths s (reversepath p) (reversepath q) \ homotopic_paths s p q" - using homotopic_paths_reversepath_D by force - - -proposition homotopic_paths_join: - "\homotopic_paths s p p'; homotopic_paths s q q'; pathfinish p = pathstart q\ \ homotopic_paths s (p +++ q) (p' +++ q')" - apply (simp add: homotopic_paths_def homotopic_with_def, clarify) - apply (rename_tac k1 k2) - apply (rule_tac x="(\y. ((k1 \ Pair (fst y)) +++ (k2 \ Pair (fst y))) (snd y))" in exI) - apply (rule conjI continuous_intros homotopic_join_lemma)+ - apply (auto simp: joinpaths_def pathstart_def pathfinish_def path_image_def) - done - -proposition homotopic_paths_continuous_image: - "\homotopic_paths s f g; continuous_on s h; h ` s \ t\ \ homotopic_paths t (h \ f) (h \ g)" - unfolding homotopic_paths_def - apply (rule homotopic_with_compose_continuous_left [of _ _ _ s]) - apply (auto simp: pathstart_def pathfinish_def elim!: homotopic_with_mono) - done - - -subsection\Group properties for homotopy of paths\ - -text%important\So taking equivalence classes under homotopy would give the fundamental group\ - -proposition homotopic_paths_rid: - "\path p; path_image p \ s\ \ homotopic_paths s (p +++ linepath (pathfinish p) (pathfinish p)) p" - apply (subst homotopic_paths_sym) - apply (rule homotopic_paths_reparametrize [where f = "\t. if t \ 1 / 2 then 2 *\<^sub>R t else 1"]) - apply (simp_all del: le_divide_eq_numeral1) - apply (subst split_01) - apply (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+ - done - -proposition homotopic_paths_lid: - "\path p; path_image p \ s\ \ homotopic_paths s (linepath (pathstart p) (pathstart p) +++ p) p" - using homotopic_paths_rid [of "reversepath p" s] - by (metis homotopic_paths_reversepath path_image_reversepath path_reversepath pathfinish_linepath - pathfinish_reversepath reversepath_joinpaths reversepath_linepath) - -proposition homotopic_paths_assoc: - "\path p; path_image p \ s; path q; path_image q \ s; path r; path_image r \ s; pathfinish p = pathstart q; - pathfinish q = pathstart r\ - \ homotopic_paths s (p +++ (q +++ r)) ((p +++ q) +++ r)" - apply (subst homotopic_paths_sym) - apply (rule homotopic_paths_reparametrize - [where f = "\t. if t \ 1 / 2 then inverse 2 *\<^sub>R t - else if t \ 3 / 4 then t - (1 / 4) - else 2 *\<^sub>R t - 1"]) - apply (simp_all del: le_divide_eq_numeral1) - apply (simp add: subset_path_image_join) - apply (rule continuous_on_cases_1 continuous_intros)+ - apply (auto simp: joinpaths_def) - done - -proposition homotopic_paths_rinv: - assumes "path p" "path_image p \ s" - shows "homotopic_paths s (p +++ reversepath p) (linepath (pathstart p) (pathstart p))" -proof - - have "continuous_on ({0..1} \ {0..1}) (\x. (subpath 0 (fst x) p +++ reversepath (subpath 0 (fst x) p)) (snd x))" - using assms - apply (simp add: joinpaths_def subpath_def reversepath_def path_def del: le_divide_eq_numeral1) - apply (rule continuous_on_cases_le) - apply (rule_tac [2] continuous_on_compose [of _ _ p, unfolded o_def]) - apply (rule continuous_on_compose [of _ _ p, unfolded o_def]) - apply (auto intro!: continuous_intros simp del: eq_divide_eq_numeral1) - apply (force elim!: continuous_on_subset simp add: mult_le_one)+ - done - then show ?thesis - using assms - apply (subst homotopic_paths_sym_eq) - unfolding homotopic_paths_def homotopic_with_def - apply (rule_tac x="(\y. (subpath 0 (fst y) p +++ reversepath(subpath 0 (fst y) p)) (snd y))" in exI) - apply (simp add: path_defs joinpaths_def subpath_def reversepath_def) - apply (force simp: mult_le_one) - done -qed - -proposition homotopic_paths_linv: - assumes "path p" "path_image p \ s" - shows "homotopic_paths s (reversepath p +++ p) (linepath (pathfinish p) (pathfinish p))" - using homotopic_paths_rinv [of "reversepath p" s] assms by simp - - -subsection\Homotopy of loops without requiring preservation of endpoints\ - -definition%important homotopic_loops :: "'a::topological_space set \ (real \ 'a) \ (real \ 'a) \ bool" where - "homotopic_loops s p q \ - homotopic_with (\r. pathfinish r = pathstart r) {0..1} s p q" - -lemma homotopic_loops: - "homotopic_loops s p q \ - (\h. continuous_on ({0..1::real} \ {0..1}) h \ - image h ({0..1} \ {0..1}) \ s \ - (\x \ {0..1}. h(0,x) = p x) \ - (\x \ {0..1}. h(1,x) = q x) \ - (\t \ {0..1}. pathfinish(h \ Pair t) = pathstart(h \ Pair t)))" - by (simp add: homotopic_loops_def pathstart_def pathfinish_def homotopic_with) - -proposition homotopic_loops_imp_loop: - "homotopic_loops s p q \ pathfinish p = pathstart p \ pathfinish q = pathstart q" -using homotopic_with_imp_property homotopic_loops_def by blast - -proposition homotopic_loops_imp_path: - "homotopic_loops s p q \ path p \ path q" - unfolding homotopic_loops_def path_def - using homotopic_with_imp_continuous by blast - -proposition homotopic_loops_imp_subset: - "homotopic_loops s p q \ path_image p \ s \ path_image q \ s" - unfolding homotopic_loops_def path_image_def - by (metis homotopic_with_imp_subset1 homotopic_with_imp_subset2) - -proposition homotopic_loops_refl: - "homotopic_loops s p p \ - path p \ path_image p \ s \ pathfinish p = pathstart p" - by (simp add: homotopic_loops_def homotopic_with_refl path_image_def path_def) - -proposition homotopic_loops_sym: "homotopic_loops s p q \ homotopic_loops s q p" - by (simp add: homotopic_loops_def homotopic_with_sym) - -proposition homotopic_loops_sym_eq: "homotopic_loops s p q \ homotopic_loops s q p" - by (metis homotopic_loops_sym) - -proposition homotopic_loops_trans: - "\homotopic_loops s p q; homotopic_loops s q r\ \ homotopic_loops s p r" - unfolding homotopic_loops_def by (blast intro: homotopic_with_trans) - -proposition homotopic_loops_subset: - "\homotopic_loops s p q; s \ t\ \ homotopic_loops t p q" - by (simp add: homotopic_loops_def homotopic_with_subset_right) - -proposition homotopic_loops_eq: - "\path p; path_image p \ s; pathfinish p = pathstart p; \t. t \ {0..1} \ p(t) = q(t)\ - \ homotopic_loops s p q" - unfolding homotopic_loops_def - apply (rule homotopic_with_eq) - apply (rule homotopic_with_refl [where f = p, THEN iffD2]) - apply (simp_all add: path_image_def path_def pathstart_def pathfinish_def) - done - -proposition homotopic_loops_continuous_image: - "\homotopic_loops s f g; continuous_on s h; h ` s \ t\ \ homotopic_loops t (h \ f) (h \ g)" - unfolding homotopic_loops_def - apply (rule homotopic_with_compose_continuous_left) - apply (erule homotopic_with_mono) - by (simp add: pathfinish_def pathstart_def) - - -subsection\Relations between the two variants of homotopy\ - -proposition homotopic_paths_imp_homotopic_loops: - "\homotopic_paths s p q; pathfinish p = pathstart p; pathfinish q = pathstart p\ \ homotopic_loops s p q" - by (auto simp: homotopic_paths_def homotopic_loops_def intro: homotopic_with_mono) - -proposition homotopic_loops_imp_homotopic_paths_null: - assumes "homotopic_loops s p (linepath a a)" - shows "homotopic_paths s p (linepath (pathstart p) (pathstart p))" -proof - - have "path p" by (metis assms homotopic_loops_imp_path) - have ploop: "pathfinish p = pathstart p" by (metis assms homotopic_loops_imp_loop) - have pip: "path_image p \ s" by (metis assms homotopic_loops_imp_subset) - obtain h where conth: "continuous_on ({0..1::real} \ {0..1}) h" - and hs: "h ` ({0..1} \ {0..1}) \ s" - and [simp]: "\x. x \ {0..1} \ h(0,x) = p x" - and [simp]: "\x. x \ {0..1} \ h(1,x) = a" - and ends: "\t. t \ {0..1} \ pathfinish (h \ Pair t) = pathstart (h \ Pair t)" - using assms by (auto simp: homotopic_loops homotopic_with) - have conth0: "path (\u. h (u, 0))" - unfolding path_def - apply (rule continuous_on_compose [of _ _ h, unfolded o_def]) - apply (force intro: continuous_intros continuous_on_subset [OF conth])+ - done - have pih0: "path_image (\u. h (u, 0)) \ s" - using hs by (force simp: path_image_def) - have c1: "continuous_on ({0..1} \ {0..1}) (\x. h (fst x * snd x, 0))" - apply (rule continuous_on_compose [of _ _ h, unfolded o_def]) - apply (force simp: mult_le_one intro: continuous_intros continuous_on_subset [OF conth])+ - done - have c2: "continuous_on ({0..1} \ {0..1}) (\x. h (fst x - fst x * snd x, 0))" - apply (rule continuous_on_compose [of _ _ h, unfolded o_def]) - apply (force simp: mult_left_le mult_le_one intro: continuous_intros continuous_on_subset [OF conth])+ - apply (rule continuous_on_subset [OF conth]) - apply (auto simp: algebra_simps add_increasing2 mult_left_le) - done - have [simp]: "\t. \0 \ t \ t \ 1\ \ h (t, 1) = h (t, 0)" - using ends by (simp add: pathfinish_def pathstart_def) - have adhoc_le: "c * 4 \ 1 + c * (d * 4)" if "\ d * 4 \ 3" "0 \ c" "c \ 1" for c d::real - proof - - have "c * 3 \ c * (d * 4)" using that less_eq_real_def by auto - with \c \ 1\ show ?thesis by fastforce - qed - have *: "\p x. (path p \ path(reversepath p)) \ - (path_image p \ s \ path_image(reversepath p) \ s) \ - (pathfinish p = pathstart(linepath a a +++ reversepath p) \ - pathstart(reversepath p) = a) \ pathstart p = x - \ homotopic_paths s (p +++ linepath a a +++ reversepath p) (linepath x x)" - by (metis homotopic_paths_lid homotopic_paths_join - homotopic_paths_trans homotopic_paths_sym homotopic_paths_rinv) - have 1: "homotopic_paths s p (p +++ linepath (pathfinish p) (pathfinish p))" - using \path p\ homotopic_paths_rid homotopic_paths_sym pip by blast - moreover have "homotopic_paths s (p +++ linepath (pathfinish p) (pathfinish p)) - (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p))" - apply (rule homotopic_paths_sym) - using homotopic_paths_lid [of "p +++ linepath (pathfinish p) (pathfinish p)" s] - by (metis 1 homotopic_paths_imp_path homotopic_paths_imp_pathstart homotopic_paths_imp_subset) - moreover have "homotopic_paths s (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p)) - ((\u. h (u, 0)) +++ linepath a a +++ reversepath (\u. h (u, 0)))" - apply (simp add: homotopic_paths_def homotopic_with_def) - apply (rule_tac x="\y. (subpath 0 (fst y) (\u. h (u, 0)) +++ (\u. h (Pair (fst y) u)) +++ subpath (fst y) 0 (\u. h (u, 0))) (snd y)" in exI) - apply (simp add: subpath_reversepath) - apply (intro conjI homotopic_join_lemma) - using ploop - apply (simp_all add: path_defs joinpaths_def o_def subpath_def conth c1 c2) - apply (force simp: algebra_simps mult_le_one mult_left_le intro: hs [THEN subsetD] adhoc_le) - done - moreover have "homotopic_paths s ((\u. h (u, 0)) +++ linepath a a +++ reversepath (\u. h (u, 0))) - (linepath (pathstart p) (pathstart p))" - apply (rule *) - apply (simp add: pih0 pathstart_def pathfinish_def conth0) - apply (simp add: reversepath_def joinpaths_def) - done - ultimately show ?thesis - by (blast intro: homotopic_paths_trans) -qed - -proposition homotopic_loops_conjugate: - fixes s :: "'a::real_normed_vector set" - assumes "path p" "path q" and pip: "path_image p \ s" and piq: "path_image q \ s" - and papp: "pathfinish p = pathstart q" and qloop: "pathfinish q = pathstart q" - shows "homotopic_loops s (p +++ q +++ reversepath p) q" -proof - - have contp: "continuous_on {0..1} p" using \path p\ [unfolded path_def] by blast - have contq: "continuous_on {0..1} q" using \path q\ [unfolded path_def] by blast - have c1: "continuous_on ({0..1} \ {0..1}) (\x. p ((1 - fst x) * snd x + fst x))" - apply (rule continuous_on_compose [of _ _ p, unfolded o_def]) - apply (force simp: mult_le_one intro!: continuous_intros) - apply (rule continuous_on_subset [OF contp]) - apply (auto simp: algebra_simps add_increasing2 mult_right_le_one_le sum_le_prod1) - done - have c2: "continuous_on ({0..1} \ {0..1}) (\x. p ((fst x - 1) * snd x + 1))" - apply (rule continuous_on_compose [of _ _ p, unfolded o_def]) - apply (force simp: mult_le_one intro!: continuous_intros) - apply (rule continuous_on_subset [OF contp]) - apply (auto simp: algebra_simps add_increasing2 mult_left_le_one_le) - done - have ps1: "\a b. \b * 2 \ 1; 0 \ b; 0 \ a; a \ 1\ \ p ((1 - a) * (2 * b) + a) \ s" - using sum_le_prod1 - by (force simp: algebra_simps add_increasing2 mult_left_le intro: pip [unfolded path_image_def, THEN subsetD]) - have ps2: "\a b. \\ 4 * b \ 3; b \ 1; 0 \ a; a \ 1\ \ p ((a - 1) * (4 * b - 3) + 1) \ s" - apply (rule pip [unfolded path_image_def, THEN subsetD]) - apply (rule image_eqI, blast) - apply (simp add: algebra_simps) - by (metis add_mono_thms_linordered_semiring(1) affine_ineq linear mult.commute mult.left_neutral mult_right_mono not_le - add.commute zero_le_numeral) - have qs: "\a b. \4 * b \ 3; \ b * 2 \ 1\ \ q (4 * b - 2) \ s" - using path_image_def piq by fastforce - have "homotopic_loops s (p +++ q +++ reversepath p) - (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q))" - apply (simp add: homotopic_loops_def homotopic_with_def) - apply (rule_tac x="(\y. (subpath (fst y) 1 p +++ q +++ subpath 1 (fst y) p) (snd y))" in exI) - apply (simp add: subpath_refl subpath_reversepath) - apply (intro conjI homotopic_join_lemma) - using papp qloop - apply (simp_all add: path_defs joinpaths_def o_def subpath_def c1 c2) - apply (force simp: contq intro: continuous_on_compose [of _ _ q, unfolded o_def] continuous_on_id continuous_on_snd) - apply (auto simp: ps1 ps2 qs) - done - moreover have "homotopic_loops s (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) q" - proof - - have "homotopic_paths s (linepath (pathfinish q) (pathfinish q) +++ q) q" - using \path q\ homotopic_paths_lid qloop piq by auto - hence 1: "\f. homotopic_paths s f q \ \ homotopic_paths s f (linepath (pathfinish q) (pathfinish q) +++ q)" - using homotopic_paths_trans by blast - hence "homotopic_paths s (linepath (pathfinish q) (pathfinish q) +++ q +++ linepath (pathfinish q) (pathfinish q)) q" - proof - - have "homotopic_paths s (q +++ linepath (pathfinish q) (pathfinish q)) q" - by (simp add: \path q\ homotopic_paths_rid piq) - thus ?thesis - by (metis (no_types) 1 \path q\ homotopic_paths_join homotopic_paths_rinv homotopic_paths_sym - homotopic_paths_trans qloop pathfinish_linepath piq) - qed - thus ?thesis - by (metis (no_types) qloop homotopic_loops_sym homotopic_paths_imp_homotopic_loops homotopic_paths_imp_pathfinish homotopic_paths_sym) - qed - ultimately show ?thesis - by (blast intro: homotopic_loops_trans) -qed - -lemma homotopic_paths_loop_parts: - assumes loops: "homotopic_loops S (p +++ reversepath q) (linepath a a)" and "path q" - shows "homotopic_paths S p q" -proof - - have paths: "homotopic_paths S (p +++ reversepath q) (linepath (pathstart p) (pathstart p))" - using homotopic_loops_imp_homotopic_paths_null [OF loops] by simp - then have "path p" - using \path q\ homotopic_loops_imp_path loops path_join path_join_path_ends path_reversepath by blast - show ?thesis - proof (cases "pathfinish p = pathfinish q") - case True - have pipq: "path_image p \ S" "path_image q \ S" - by (metis Un_subset_iff paths \path p\ \path q\ homotopic_loops_imp_subset homotopic_paths_imp_path loops - path_image_join path_image_reversepath path_imp_reversepath path_join_eq)+ - have "homotopic_paths S p (p +++ (linepath (pathfinish p) (pathfinish p)))" - using \path p\ \path_image p \ S\ homotopic_paths_rid homotopic_paths_sym by blast - moreover have "homotopic_paths S (p +++ (linepath (pathfinish p) (pathfinish p))) (p +++ (reversepath q +++ q))" - by (simp add: True \path p\ \path q\ pipq homotopic_paths_join homotopic_paths_linv homotopic_paths_sym) - moreover have "homotopic_paths S (p +++ (reversepath q +++ q)) ((p +++ reversepath q) +++ q)" - by (simp add: True \path p\ \path q\ homotopic_paths_assoc pipq) - moreover have "homotopic_paths S ((p +++ reversepath q) +++ q) (linepath (pathstart p) (pathstart p) +++ q)" - by (simp add: \path q\ homotopic_paths_join paths pipq) - moreover then have "homotopic_paths S (linepath (pathstart p) (pathstart p) +++ q) q" - by (metis \path q\ homotopic_paths_imp_path homotopic_paths_lid linepath_trivial path_join_path_ends pathfinish_def pipq(2)) - ultimately show ?thesis - using homotopic_paths_trans by metis - next - case False - then show ?thesis - using \path q\ homotopic_loops_imp_path loops path_join_path_ends by fastforce - qed -qed - - -subsection%unimportant\Homotopy of "nearby" function, paths and loops\ - -lemma homotopic_with_linear: - fixes f g :: "_ \ 'b::real_normed_vector" - assumes contf: "continuous_on s f" - and contg:"continuous_on s g" - and sub: "\x. x \ s \ closed_segment (f x) (g x) \ t" - shows "homotopic_with (\z. True) s t f g" - apply (simp add: homotopic_with_def) - apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R f(snd y) + (fst y) *\<^sub>R g(snd y))" in exI) - apply (intro conjI) - apply (rule subset_refl continuous_intros continuous_on_subset [OF contf] continuous_on_compose2 [where g=f] - continuous_on_subset [OF contg] continuous_on_compose2 [where g=g]| simp)+ - using sub closed_segment_def apply fastforce+ - done - -lemma homotopic_paths_linear: - fixes g h :: "real \ 'a::real_normed_vector" - assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" - "\t. t \ {0..1} \ closed_segment (g t) (h t) \ s" - shows "homotopic_paths s g h" - using assms - unfolding path_def - apply (simp add: closed_segment_def pathstart_def pathfinish_def homotopic_paths_def homotopic_with_def) - apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R (g \ snd) y + (fst y) *\<^sub>R (h \ snd) y)" in exI) - apply (intro conjI subsetI continuous_intros; force) - done - -lemma homotopic_loops_linear: - fixes g h :: "real \ 'a::real_normed_vector" - assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" - "\t x. t \ {0..1} \ closed_segment (g t) (h t) \ s" - shows "homotopic_loops s g h" - using assms - unfolding path_def - apply (simp add: pathstart_def pathfinish_def homotopic_loops_def homotopic_with_def) - apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R g(snd y) + (fst y) *\<^sub>R h(snd y))" in exI) - apply (auto intro!: continuous_intros intro: continuous_on_compose2 [where g=g] continuous_on_compose2 [where g=h]) - apply (force simp: closed_segment_def) - done - -lemma homotopic_paths_nearby_explicit: - assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" - and no: "\t x. \t \ {0..1}; x \ s\ \ norm(h t - g t) < norm(g t - x)" - shows "homotopic_paths s g h" - apply (rule homotopic_paths_linear [OF assms(1-4)]) - by (metis no segment_bound(1) subsetI norm_minus_commute not_le) - -lemma homotopic_loops_nearby_explicit: - assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" - and no: "\t x. \t \ {0..1}; x \ s\ \ norm(h t - g t) < norm(g t - x)" - shows "homotopic_loops s g h" - apply (rule homotopic_loops_linear [OF assms(1-4)]) - by (metis no segment_bound(1) subsetI norm_minus_commute not_le) - -lemma homotopic_nearby_paths: - fixes g h :: "real \ 'a::euclidean_space" - assumes "path g" "open s" "path_image g \ s" - shows "\e. 0 < e \ - (\h. path h \ - pathstart h = pathstart g \ pathfinish h = pathfinish g \ - (\t \ {0..1}. norm(h t - g t) < e) \ homotopic_paths s g h)" -proof - - obtain e where "e > 0" and e: "\x y. x \ path_image g \ y \ - s \ e \ dist x y" - using separate_compact_closed [of "path_image g" "-s"] assms by force - show ?thesis - apply (intro exI conjI) - using e [unfolded dist_norm] - apply (auto simp: intro!: homotopic_paths_nearby_explicit assms \e > 0\) - by (metis atLeastAtMost_iff imageI le_less_trans not_le path_image_def) -qed - -lemma homotopic_nearby_loops: - fixes g h :: "real \ 'a::euclidean_space" - assumes "path g" "open s" "path_image g \ s" "pathfinish g = pathstart g" - shows "\e. 0 < e \ - (\h. path h \ pathfinish h = pathstart h \ - (\t \ {0..1}. norm(h t - g t) < e) \ homotopic_loops s g h)" -proof - - obtain e where "e > 0" and e: "\x y. x \ path_image g \ y \ - s \ e \ dist x y" - using separate_compact_closed [of "path_image g" "-s"] assms by force - show ?thesis - apply (intro exI conjI) - using e [unfolded dist_norm] - apply (auto simp: intro!: homotopic_loops_nearby_explicit assms \e > 0\) - by (metis atLeastAtMost_iff imageI le_less_trans not_le path_image_def) -qed - - -subsection\ Homotopy and subpaths\ - -lemma homotopic_join_subpaths1: - assumes "path g" and pag: "path_image g \ s" - and u: "u \ {0..1}" and v: "v \ {0..1}" and w: "w \ {0..1}" "u \ v" "v \ w" - shows "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" -proof - - have 1: "t * 2 \ 1 \ u + t * (v * 2) \ v + t * (u * 2)" for t - using affine_ineq \u \ v\ by fastforce - have 2: "t * 2 > 1 \ u + (2*t - 1) * v \ v + (2*t - 1) * w" for t - by (metis add_mono_thms_linordered_semiring(1) diff_gt_0_iff_gt less_eq_real_def mult.commute mult_right_mono \u \ v\ \v \ w\) - have t2: "\t::real. t*2 = 1 \ t = 1/2" by auto - show ?thesis - apply (rule homotopic_paths_subset [OF _ pag]) - using assms - apply (cases "w = u") - using homotopic_paths_rinv [of "subpath u v g" "path_image g"] - apply (force simp: closed_segment_eq_real_ivl image_mono path_image_def subpath_refl) - apply (rule homotopic_paths_sym) - apply (rule homotopic_paths_reparametrize - [where f = "\t. if t \ 1 / 2 - then inverse((w - u)) *\<^sub>R (2 * (v - u)) *\<^sub>R t - else inverse((w - u)) *\<^sub>R ((v - u) + (w - v) *\<^sub>R (2 *\<^sub>R t - 1))"]) - using \path g\ path_subpath u w apply blast - using \path g\ path_image_subpath_subset u w(1) apply blast - apply simp_all - apply (subst split_01) - apply (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+ - apply (simp_all add: field_simps not_le) - apply (force dest!: t2) - apply (force simp: algebra_simps mult_left_mono affine_ineq dest!: 1 2) - apply (simp add: joinpaths_def subpath_def) - apply (force simp: algebra_simps) - done -qed - -lemma homotopic_join_subpaths2: - assumes "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" - shows "homotopic_paths s (subpath w v g +++ subpath v u g) (subpath w u g)" -by (metis assms homotopic_paths_reversepath_D pathfinish_subpath pathstart_subpath reversepath_joinpaths reversepath_subpath) - -lemma homotopic_join_subpaths3: - assumes hom: "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" - and "path g" and pag: "path_image g \ s" - and u: "u \ {0..1}" and v: "v \ {0..1}" and w: "w \ {0..1}" - shows "homotopic_paths s (subpath v w g +++ subpath w u g) (subpath v u g)" -proof - - have "homotopic_paths s (subpath u w g +++ subpath w v g) ((subpath u v g +++ subpath v w g) +++ subpath w v g)" - apply (rule homotopic_paths_join) - using hom homotopic_paths_sym_eq apply blast - apply (metis \path g\ homotopic_paths_eq pag path_image_subpath_subset path_subpath subset_trans v w, simp) - done - also have "homotopic_paths s ((subpath u v g +++ subpath v w g) +++ subpath w v g) (subpath u v g +++ subpath v w g +++ subpath w v g)" - apply (rule homotopic_paths_sym [OF homotopic_paths_assoc]) - using assms by (simp_all add: path_image_subpath_subset [THEN order_trans]) - also have "homotopic_paths s (subpath u v g +++ subpath v w g +++ subpath w v g) - (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))" - apply (rule homotopic_paths_join) - apply (metis \path g\ homotopic_paths_eq order.trans pag path_image_subpath_subset path_subpath u v) - apply (metis (no_types, lifting) \path g\ homotopic_paths_linv order_trans pag path_image_subpath_subset path_subpath pathfinish_subpath reversepath_subpath v w) - apply simp - done - also have "homotopic_paths s (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g))) (subpath u v g)" - apply (rule homotopic_paths_rid) - using \path g\ path_subpath u v apply blast - apply (meson \path g\ order.trans pag path_image_subpath_subset u v) - done - finally have "homotopic_paths s (subpath u w g +++ subpath w v g) (subpath u v g)" . - then show ?thesis - using homotopic_join_subpaths2 by blast -qed - -proposition homotopic_join_subpaths: - "\path g; path_image g \ s; u \ {0..1}; v \ {0..1}; w \ {0..1}\ - \ homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" - apply (rule le_cases3 [of u v w]) -using homotopic_join_subpaths1 homotopic_join_subpaths2 homotopic_join_subpaths3 by metis+ - -text\Relating homotopy of trivial loops to path-connectedness.\ - -lemma path_component_imp_homotopic_points: - "path_component S a b \ homotopic_loops S (linepath a a) (linepath b b)" -apply (simp add: path_component_def homotopic_loops_def homotopic_with_def - pathstart_def pathfinish_def path_image_def path_def, clarify) -apply (rule_tac x="g \ fst" in exI) -apply (intro conjI continuous_intros continuous_on_compose)+ -apply (auto elim!: continuous_on_subset) -done - -lemma homotopic_loops_imp_path_component_value: - "\homotopic_loops S p q; 0 \ t; t \ 1\ - \ path_component S (p t) (q t)" -apply (simp add: path_component_def homotopic_loops_def homotopic_with_def - pathstart_def pathfinish_def path_image_def path_def, clarify) -apply (rule_tac x="h \ (\u. (u, t))" in exI) -apply (intro conjI continuous_intros continuous_on_compose)+ -apply (auto elim!: continuous_on_subset) -done - -lemma homotopic_points_eq_path_component: - "homotopic_loops S (linepath a a) (linepath b b) \ - path_component S a b" -by (auto simp: path_component_imp_homotopic_points - dest: homotopic_loops_imp_path_component_value [where t=1]) - -lemma path_connected_eq_homotopic_points: - "path_connected S \ - (\a b. a \ S \ b \ S \ homotopic_loops S (linepath a a) (linepath b b))" -by (auto simp: path_connected_def path_component_def homotopic_points_eq_path_component) - - -subsection\Simply connected sets\ - -text%important\defined as "all loops are homotopic (as loops)\ - -definition%important simply_connected where - "simply_connected S \ - \p q. path p \ pathfinish p = pathstart p \ path_image p \ S \ - path q \ pathfinish q = pathstart q \ path_image q \ S - \ homotopic_loops S p q" - -lemma simply_connected_empty [iff]: "simply_connected {}" - by (simp add: simply_connected_def) - -lemma simply_connected_imp_path_connected: - fixes S :: "_::real_normed_vector set" - shows "simply_connected S \ path_connected S" -by (simp add: simply_connected_def path_connected_eq_homotopic_points) - -lemma simply_connected_imp_connected: - fixes S :: "_::real_normed_vector set" - shows "simply_connected S \ connected S" -by (simp add: path_connected_imp_connected simply_connected_imp_path_connected) - -lemma simply_connected_eq_contractible_loop_any: - fixes S :: "_::real_normed_vector set" - shows "simply_connected S \ - (\p a. path p \ path_image p \ S \ - pathfinish p = pathstart p \ a \ S - \ homotopic_loops S p (linepath a a))" -apply (simp add: simply_connected_def) -apply (rule iffI, force, clarify) -apply (rule_tac q = "linepath (pathstart p) (pathstart p)" in homotopic_loops_trans) -apply (fastforce simp add:) -using homotopic_loops_sym apply blast -done - -lemma simply_connected_eq_contractible_loop_some: - fixes S :: "_::real_normed_vector set" - shows "simply_connected S \ - path_connected S \ - (\p. path p \ path_image p \ S \ pathfinish p = pathstart p - \ (\a. a \ S \ homotopic_loops S p (linepath a a)))" -apply (rule iffI) - apply (fastforce simp: simply_connected_imp_path_connected simply_connected_eq_contractible_loop_any) -apply (clarsimp simp add: simply_connected_eq_contractible_loop_any) -apply (drule_tac x=p in spec) -using homotopic_loops_trans path_connected_eq_homotopic_points - apply blast -done - -lemma simply_connected_eq_contractible_loop_all: - fixes S :: "_::real_normed_vector set" - shows "simply_connected S \ - S = {} \ - (\a \ S. \p. path p \ path_image p \ S \ pathfinish p = pathstart p - \ homotopic_loops S p (linepath a a))" - (is "?lhs = ?rhs") -proof (cases "S = {}") - case True then show ?thesis by force -next - case False - then obtain a where "a \ S" by blast - show ?thesis - proof - assume "simply_connected S" - then show ?rhs - using \a \ S\ \simply_connected S\ simply_connected_eq_contractible_loop_any - by blast - next - assume ?rhs - then show "simply_connected S" - apply (simp add: simply_connected_eq_contractible_loop_any False) - by (meson homotopic_loops_refl homotopic_loops_sym homotopic_loops_trans - path_component_imp_homotopic_points path_component_refl) - qed -qed - -lemma simply_connected_eq_contractible_path: - fixes S :: "_::real_normed_vector set" - shows "simply_connected S \ - path_connected S \ - (\p. path p \ path_image p \ S \ pathfinish p = pathstart p - \ homotopic_paths S p (linepath (pathstart p) (pathstart p)))" -apply (rule iffI) - apply (simp add: simply_connected_imp_path_connected) - apply (metis simply_connected_eq_contractible_loop_some homotopic_loops_imp_homotopic_paths_null) -by (meson homotopic_paths_imp_homotopic_loops pathfinish_linepath pathstart_in_path_image - simply_connected_eq_contractible_loop_some subset_iff) - -lemma simply_connected_eq_homotopic_paths: - fixes S :: "_::real_normed_vector set" - shows "simply_connected S \ - path_connected S \ - (\p q. path p \ path_image p \ S \ - path q \ path_image q \ S \ - pathstart q = pathstart p \ pathfinish q = pathfinish p - \ homotopic_paths S p q)" - (is "?lhs = ?rhs") -proof - assume ?lhs - then have pc: "path_connected S" - and *: "\p. \path p; path_image p \ S; - pathfinish p = pathstart p\ - \ homotopic_paths S p (linepath (pathstart p) (pathstart p))" - by (auto simp: simply_connected_eq_contractible_path) - have "homotopic_paths S p q" - if "path p" "path_image p \ S" "path q" - "path_image q \ S" "pathstart q = pathstart p" - "pathfinish q = pathfinish p" for p q - proof - - have "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))" - by (simp add: homotopic_paths_rid homotopic_paths_sym that) - also have "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p)) - (p +++ reversepath q +++ q)" - using that - by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl homotopic_paths_sym_eq pathstart_linepath) - also have "homotopic_paths S (p +++ reversepath q +++ q) - ((p +++ reversepath q) +++ q)" - by (simp add: that homotopic_paths_assoc) - also have "homotopic_paths S ((p +++ reversepath q) +++ q) - (linepath (pathstart q) (pathstart q) +++ q)" - using * [of "p +++ reversepath q"] that - by (simp add: homotopic_paths_join path_image_join) - also have "homotopic_paths S (linepath (pathstart q) (pathstart q) +++ q) q" - using that homotopic_paths_lid by blast - finally show ?thesis . - qed - then show ?rhs - by (blast intro: pc *) -next - assume ?rhs - then show ?lhs - by (force simp: simply_connected_eq_contractible_path) -qed - -proposition simply_connected_Times: - fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" - assumes S: "simply_connected S" and T: "simply_connected T" - shows "simply_connected(S \ T)" -proof - - have "homotopic_loops (S \ T) p (linepath (a, b) (a, b))" - if "path p" "path_image p \ S \ T" "p 1 = p 0" "a \ S" "b \ T" - for p a b - proof - - have "path (fst \ p)" - apply (rule Path_Connected.path_continuous_image [OF \path p\]) - apply (rule continuous_intros)+ - done - moreover have "path_image (fst \ p) \ S" - using that apply (simp add: path_image_def) by force - ultimately have p1: "homotopic_loops S (fst \ p) (linepath a a)" - using S that - apply (simp add: simply_connected_eq_contractible_loop_any) - apply (drule_tac x="fst \ p" in spec) - apply (drule_tac x=a in spec) - apply (auto simp: pathstart_def pathfinish_def) - done - have "path (snd \ p)" - apply (rule Path_Connected.path_continuous_image [OF \path p\]) - apply (rule continuous_intros)+ - done - moreover have "path_image (snd \ p) \ T" - using that apply (simp add: path_image_def) by force - ultimately have p2: "homotopic_loops T (snd \ p) (linepath b b)" - using T that - apply (simp add: simply_connected_eq_contractible_loop_any) - apply (drule_tac x="snd \ p" in spec) - apply (drule_tac x=b in spec) - apply (auto simp: pathstart_def pathfinish_def) - done - show ?thesis - using p1 p2 - apply (simp add: homotopic_loops, clarify) - apply (rename_tac h k) - apply (rule_tac x="\z. Pair (h z) (k z)" in exI) - apply (intro conjI continuous_intros | assumption)+ - apply (auto simp: pathstart_def pathfinish_def) - done - qed - with assms show ?thesis - by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) -qed - - -subsection\Contractible sets\ - -definition%important contractible where - "contractible S \ \a. homotopic_with (\x. True) S S id (\x. a)" - -proposition contractible_imp_simply_connected: - fixes S :: "_::real_normed_vector set" - assumes "contractible S" shows "simply_connected S" -proof (cases "S = {}") - case True then show ?thesis by force -next - case False - obtain a where a: "homotopic_with (\x. True) S S id (\x. a)" - using assms by (force simp: contractible_def) - then have "a \ S" - by (metis False homotopic_constant_maps homotopic_with_symD homotopic_with_trans path_component_mem(2)) - show ?thesis - apply (simp add: simply_connected_eq_contractible_loop_all False) - apply (rule bexI [OF _ \a \ S\]) - using a apply (simp add: homotopic_loops_def homotopic_with_def path_def path_image_def pathfinish_def pathstart_def, clarify) - apply (rule_tac x="(h \ (\y. (fst y, (p \ snd) y)))" in exI) - apply (intro conjI continuous_on_compose continuous_intros) - apply (erule continuous_on_subset | force)+ - done -qed - -corollary contractible_imp_connected: - fixes S :: "_::real_normed_vector set" - shows "contractible S \ connected S" -by (simp add: contractible_imp_simply_connected simply_connected_imp_connected) - -lemma contractible_imp_path_connected: - fixes S :: "_::real_normed_vector set" - shows "contractible S \ path_connected S" -by (simp add: contractible_imp_simply_connected simply_connected_imp_path_connected) - -lemma nullhomotopic_through_contractible: - fixes S :: "_::topological_space set" - assumes f: "continuous_on S f" "f ` S \ T" - and g: "continuous_on T g" "g ` T \ U" - and T: "contractible T" - obtains c where "homotopic_with (\h. True) S U (g \ f) (\x. c)" -proof - - obtain b where b: "homotopic_with (\x. True) T T id (\x. b)" - using assms by (force simp: contractible_def) - have "homotopic_with (\f. True) T U (g \ id) (g \ (\x. b))" - by (rule homotopic_compose_continuous_left [OF b g]) - then have "homotopic_with (\f. True) S U (g \ id \ f) (g \ (\x. b) \ f)" - by (rule homotopic_compose_continuous_right [OF _ f]) - then show ?thesis - by (simp add: comp_def that) -qed - -lemma nullhomotopic_into_contractible: - assumes f: "continuous_on S f" "f ` S \ T" - and T: "contractible T" - obtains c where "homotopic_with (\h. True) S T f (\x. c)" -apply (rule nullhomotopic_through_contractible [OF f, of id T]) -using assms -apply (auto simp: continuous_on_id) -done - -lemma nullhomotopic_from_contractible: - assumes f: "continuous_on S f" "f ` S \ T" - and S: "contractible S" - obtains c where "homotopic_with (\h. True) S T f (\x. c)" -apply (rule nullhomotopic_through_contractible [OF continuous_on_id _ f S, of S]) -using assms -apply (auto simp: comp_def) -done - -lemma homotopic_through_contractible: - fixes S :: "_::real_normed_vector set" - assumes "continuous_on S f1" "f1 ` S \ T" - "continuous_on T g1" "g1 ` T \ U" - "continuous_on S f2" "f2 ` S \ T" - "continuous_on T g2" "g2 ` T \ U" - "contractible T" "path_connected U" - shows "homotopic_with (\h. True) S U (g1 \ f1) (g2 \ f2)" -proof - - obtain c1 where c1: "homotopic_with (\h. True) S U (g1 \ f1) (\x. c1)" - apply (rule nullhomotopic_through_contractible [of S f1 T g1 U]) - using assms apply auto - done - obtain c2 where c2: "homotopic_with (\h. True) S U (g2 \ f2) (\x. c2)" - apply (rule nullhomotopic_through_contractible [of S f2 T g2 U]) - using assms apply auto - done - have *: "S = {} \ (\t. path_connected t \ t \ U \ c2 \ t \ c1 \ t)" - proof (cases "S = {}") - case True then show ?thesis by force - next - case False - with c1 c2 have "c1 \ U" "c2 \ U" - using homotopic_with_imp_subset2 all_not_in_conv image_subset_iff by blast+ - with \path_connected U\ show ?thesis by blast - qed - show ?thesis - apply (rule homotopic_with_trans [OF c1]) - apply (rule homotopic_with_symD) - apply (rule homotopic_with_trans [OF c2]) - apply (simp add: path_component homotopic_constant_maps *) - done -qed - -lemma homotopic_into_contractible: - fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" - assumes f: "continuous_on S f" "f ` S \ T" - and g: "continuous_on S g" "g ` S \ T" - and T: "contractible T" - shows "homotopic_with (\h. True) S T f g" -using homotopic_through_contractible [of S f T id T g id] -by (simp add: assms contractible_imp_path_connected continuous_on_id) - -lemma homotopic_from_contractible: - fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" - assumes f: "continuous_on S f" "f ` S \ T" - and g: "continuous_on S g" "g ` S \ T" - and "contractible S" "path_connected T" - shows "homotopic_with (\h. True) S T f g" -using homotopic_through_contractible [of S id S f T id g] -by (simp add: assms contractible_imp_path_connected continuous_on_id) - -lemma starlike_imp_contractible_gen: - fixes S :: "'a::real_normed_vector set" - assumes S: "starlike S" - and P: "\a T. \a \ S; 0 \ T; T \ 1\ \ P(\x. (1 - T) *\<^sub>R x + T *\<^sub>R a)" - obtains a where "homotopic_with P S S (\x. x) (\x. a)" -proof - - obtain a where "a \ S" and a: "\x. x \ S \ closed_segment a x \ S" - using S by (auto simp: starlike_def) - have "(\y. (1 - fst y) *\<^sub>R snd y + fst y *\<^sub>R a) ` ({0..1} \ S) \ S" - apply clarify - apply (erule a [unfolded closed_segment_def, THEN subsetD], simp) - apply (metis add_diff_cancel_right' diff_ge_0_iff_ge le_add_diff_inverse pth_c(1)) - done - then show ?thesis - apply (rule_tac a=a in that) - using \a \ S\ - apply (simp add: homotopic_with_def) - apply (rule_tac x="\y. (1 - (fst y)) *\<^sub>R snd y + (fst y) *\<^sub>R a" in exI) - apply (intro conjI ballI continuous_on_compose continuous_intros) - apply (simp_all add: P) - done -qed - -lemma starlike_imp_contractible: - fixes S :: "'a::real_normed_vector set" - shows "starlike S \ contractible S" -using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def) - -lemma contractible_UNIV [simp]: "contractible (UNIV :: 'a::real_normed_vector set)" - by (simp add: starlike_imp_contractible) - -lemma starlike_imp_simply_connected: - fixes S :: "'a::real_normed_vector set" - shows "starlike S \ simply_connected S" -by (simp add: contractible_imp_simply_connected starlike_imp_contractible) - -lemma convex_imp_simply_connected: - fixes S :: "'a::real_normed_vector set" - shows "convex S \ simply_connected S" -using convex_imp_starlike starlike_imp_simply_connected by blast - -lemma starlike_imp_path_connected: - fixes S :: "'a::real_normed_vector set" - shows "starlike S \ path_connected S" -by (simp add: simply_connected_imp_path_connected starlike_imp_simply_connected) - -lemma starlike_imp_connected: - fixes S :: "'a::real_normed_vector set" - shows "starlike S \ connected S" -by (simp add: path_connected_imp_connected starlike_imp_path_connected) - -lemma is_interval_simply_connected_1: - fixes S :: "real set" - shows "is_interval S \ simply_connected S" -using convex_imp_simply_connected is_interval_convex_1 is_interval_path_connected_1 simply_connected_imp_path_connected by auto - -lemma contractible_empty [simp]: "contractible {}" - by (simp add: contractible_def homotopic_with) - -lemma contractible_convex_tweak_boundary_points: - fixes S :: "'a::euclidean_space set" - assumes "convex S" and TS: "rel_interior S \ T" "T \ closure S" - shows "contractible T" -proof (cases "S = {}") - case True - with assms show ?thesis - by (simp add: subsetCE) -next - case False - show ?thesis - apply (rule starlike_imp_contractible) - apply (rule starlike_convex_tweak_boundary_points [OF \convex S\ False TS]) - done -qed - -lemma convex_imp_contractible: - fixes S :: "'a::real_normed_vector set" - shows "convex S \ contractible S" - using contractible_empty convex_imp_starlike starlike_imp_contractible by blast - -lemma contractible_sing [simp]: - fixes a :: "'a::real_normed_vector" - shows "contractible {a}" -by (rule convex_imp_contractible [OF convex_singleton]) - -lemma is_interval_contractible_1: - fixes S :: "real set" - shows "is_interval S \ contractible S" -using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1 - is_interval_simply_connected_1 by auto - -lemma contractible_Times: - fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" - assumes S: "contractible S" and T: "contractible T" - shows "contractible (S \ T)" -proof - - obtain a h where conth: "continuous_on ({0..1} \ S) h" - and hsub: "h ` ({0..1} \ S) \ S" - and [simp]: "\x. x \ S \ h (0, x) = x" - and [simp]: "\x. x \ S \ h (1::real, x) = a" - using S by (auto simp: contractible_def homotopic_with) - obtain b k where contk: "continuous_on ({0..1} \ T) k" - and ksub: "k ` ({0..1} \ T) \ T" - and [simp]: "\x. x \ T \ k (0, x) = x" - and [simp]: "\x. x \ T \ k (1::real, x) = b" - using T by (auto simp: contractible_def homotopic_with) - show ?thesis - apply (simp add: contractible_def homotopic_with) - apply (rule exI [where x=a]) - apply (rule exI [where x=b]) - apply (rule exI [where x = "\z. (h (fst z, fst(snd z)), k (fst z, snd(snd z)))"]) - apply (intro conjI ballI continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF contk]) - using hsub ksub - apply auto - done -qed - -lemma homotopy_dominated_contractibility: - fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" - assumes S: "contractible S" - and f: "continuous_on S f" "image f S \ T" - and g: "continuous_on T g" "image g T \ S" - and hom: "homotopic_with (\x. True) T T (f \ g) id" - shows "contractible T" -proof - - obtain b where "homotopic_with (\h. True) S T f (\x. b)" - using nullhomotopic_from_contractible [OF f S] . - then have homg: "homotopic_with (\x. True) T T ((\x. b) \ g) (f \ g)" - by (rule homotopic_with_compose_continuous_right [OF homotopic_with_symD g]) - show ?thesis - apply (simp add: contractible_def) - apply (rule exI [where x = b]) - apply (rule homotopic_with_symD) - apply (rule homotopic_with_trans [OF _ hom]) - using homg apply (simp add: o_def) - done -qed - - -subsection\Local versions of topological properties in general\ - -definition%important locally :: "('a::topological_space set \ bool) \ 'a set \ bool" -where - "locally P S \ - \w x. openin (subtopology euclidean S) w \ x \ w - \ (\u v. openin (subtopology euclidean S) u \ P v \ - x \ u \ u \ v \ v \ w)" - -lemma locallyI: - assumes "\w x. \openin (subtopology euclidean S) w; x \ w\ - \ \u v. openin (subtopology euclidean S) u \ P v \ - x \ u \ u \ v \ v \ w" - shows "locally P S" -using assms by (force simp: locally_def) - -lemma locallyE: - assumes "locally P S" "openin (subtopology euclidean S) w" "x \ w" - obtains u v where "openin (subtopology euclidean S) u" - "P v" "x \ u" "u \ v" "v \ w" - using assms unfolding locally_def by meson - -lemma locally_mono: - assumes "locally P S" "\t. P t \ Q t" - shows "locally Q S" -by (metis assms locally_def) - -lemma locally_open_subset: - assumes "locally P S" "openin (subtopology euclidean S) t" - shows "locally P t" -using assms -apply (simp add: locally_def) -apply (erule all_forward)+ -apply (rule impI) -apply (erule impCE) - using openin_trans apply blast -apply (erule ex_forward) -by (metis (no_types, hide_lams) Int_absorb1 Int_lower1 Int_subset_iff openin_open openin_subtopology_Int_subset) - -lemma locally_diff_closed: - "\locally P S; closedin (subtopology euclidean S) t\ \ locally P (S - t)" - using locally_open_subset closedin_def by fastforce - -lemma locally_empty [iff]: "locally P {}" - by (simp add: locally_def openin_subtopology) - -lemma locally_singleton [iff]: - fixes a :: "'a::metric_space" - shows "locally P {a} \ P {a}" -apply (simp add: locally_def openin_euclidean_subtopology_iff subset_singleton_iff conj_disj_distribR cong: conj_cong) -using zero_less_one by blast - -lemma locally_iff: - "locally P S \ - (\T x. open T \ x \ S \ T \ (\U. open U \ (\v. P v \ x \ S \ U \ S \ U \ v \ v \ S \ T)))" -apply (simp add: le_inf_iff locally_def openin_open, safe) -apply (metis IntE IntI le_inf_iff) -apply (metis IntI Int_subset_iff) -done - -lemma locally_Int: - assumes S: "locally P S" and t: "locally P t" - and P: "\S t. P S \ P t \ P(S \ t)" - shows "locally P (S \ t)" -using S t unfolding locally_iff -apply clarify -apply (drule_tac x=T in spec)+ -apply (drule_tac x=x in spec)+ -apply clarsimp -apply (rename_tac U1 U2 V1 V2) -apply (rule_tac x="U1 \ U2" in exI) -apply (simp add: open_Int) -apply (rule_tac x="V1 \ V2" in exI) -apply (auto intro: P) -done - -lemma locally_Times: - fixes S :: "('a::metric_space) set" and T :: "('b::metric_space) set" - assumes PS: "locally P S" and QT: "locally Q T" and R: "\S T. P S \ Q T \ R(S \ T)" - shows "locally R (S \ T)" - unfolding locally_def -proof (clarify) - fix W x y - assume W: "openin (subtopology euclidean (S \ T)) W" and xy: "(x, y) \ W" - then obtain U V where "openin (subtopology euclidean S) U" "x \ U" - "openin (subtopology euclidean T) V" "y \ V" "U \ V \ W" - using Times_in_interior_subtopology by metis - then obtain U1 U2 V1 V2 - where opeS: "openin (subtopology euclidean S) U1 \ P U2 \ x \ U1 \ U1 \ U2 \ U2 \ U" - and opeT: "openin (subtopology euclidean T) V1 \ Q V2 \ y \ V1 \ V1 \ V2 \ V2 \ V" - by (meson PS QT locallyE) - with \U \ V \ W\ show "\u v. openin (subtopology euclidean (S \ T)) u \ R v \ (x,y) \ u \ u \ v \ v \ W" - apply (rule_tac x="U1 \ V1" in exI) - apply (rule_tac x="U2 \ V2" in exI) - apply (auto simp: openin_Times R) - done -qed - - -proposition homeomorphism_locally_imp: - fixes S :: "'a::metric_space set" and t :: "'b::t2_space set" - assumes S: "locally P S" and hom: "homeomorphism S t f g" - and Q: "\S S'. \P S; homeomorphism S S' f g\ \ Q S'" - shows "locally Q t" -proof (clarsimp simp: locally_def) - fix W y - assume "y \ W" and "openin (subtopology euclidean t) W" - then obtain T where T: "open T" "W = t \ T" - by (force simp: openin_open) - then have "W \ t" by auto - have f: "\x. x \ S \ g(f x) = x" "f ` S = t" "continuous_on S f" - and g: "\y. y \ t \ f(g y) = y" "g ` t = S" "continuous_on t g" - using hom by (auto simp: homeomorphism_def) - have gw: "g ` W = S \ f -` W" - using \W \ t\ - apply auto - using \g ` t = S\ \W \ t\ apply blast - using g \W \ t\ apply auto[1] - by (simp add: f rev_image_eqI) - have \: "openin (subtopology euclidean S) (g ` W)" - proof - - have "continuous_on S f" - using f(3) by blast - then show "openin (subtopology euclidean S) (g ` W)" - by (simp add: gw Collect_conj_eq \openin (subtopology euclidean t) W\ continuous_on_open f(2)) - qed - then obtain u v - where osu: "openin (subtopology euclidean S) u" and uv: "P v" "g y \ u" "u \ v" "v \ g ` W" - using S [unfolded locally_def, rule_format, of "g ` W" "g y"] \y \ W\ by force - have "v \ S" using uv by (simp add: gw) - have fv: "f ` v = t \ {x. g x \ v}" - using \f ` S = t\ f \v \ S\ by auto - have "f ` v \ W" - using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto - have contvf: "continuous_on v f" - using \v \ S\ continuous_on_subset f(3) by blast - have contvg: "continuous_on (f ` v) g" - using \f ` v \ W\ \W \ t\ continuous_on_subset [OF g(3)] by blast - have homv: "homeomorphism v (f ` v) f g" - using \v \ S\ \W \ t\ f - apply (simp add: homeomorphism_def contvf contvg, auto) - by (metis f(1) rev_image_eqI rev_subsetD) - have 1: "openin (subtopology euclidean t) (t \ g -` u)" - apply (rule continuous_on_open [THEN iffD1, rule_format]) - apply (rule \continuous_on t g\) - using \g ` t = S\ apply (simp add: osu) - done - have 2: "\V. Q V \ y \ (t \ g -` u) \ (t \ g -` u) \ V \ V \ W" - apply (rule_tac x="f ` v" in exI) - apply (intro conjI Q [OF \P v\ homv]) - using \W \ t\ \y \ W\ \f ` v \ W\ uv apply (auto simp: fv) - done - show "\U. openin (subtopology euclidean t) U \ (\v. Q v \ y \ U \ U \ v \ v \ W)" - by (meson 1 2) -qed - -lemma homeomorphism_locally: - fixes f:: "'a::metric_space \ 'b::metric_space" - assumes hom: "homeomorphism S t f g" - and eq: "\S t. homeomorphism S t f g \ (P S \ Q t)" - shows "locally P S \ locally Q t" -apply (rule iffI) -apply (erule homeomorphism_locally_imp [OF _ hom]) -apply (simp add: eq) -apply (erule homeomorphism_locally_imp) -using eq homeomorphism_sym homeomorphism_symD [OF hom] apply blast+ -done - -lemma homeomorphic_locally: - fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" - assumes hom: "S homeomorphic T" - and iff: "\X Y. X homeomorphic Y \ (P X \ Q Y)" - shows "locally P S \ locally Q T" -proof - - obtain f g where hom: "homeomorphism S T f g" - using assms by (force simp: homeomorphic_def) - then show ?thesis - using homeomorphic_def local.iff - by (blast intro!: homeomorphism_locally) -qed - -lemma homeomorphic_local_compactness: - fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" - shows "S homeomorphic T \ locally compact S \ locally compact T" -by (simp add: homeomorphic_compactness homeomorphic_locally) - -lemma locally_translation: - fixes P :: "'a :: real_normed_vector set \ bool" - shows - "(\S. P (image (\x. a + x) S) \ P S) - \ locally P (image (\x. a + x) S) \ locally P S" -apply (rule homeomorphism_locally [OF homeomorphism_translation]) -apply (simp add: homeomorphism_def) -by metis - -lemma locally_injective_linear_image: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes f: "linear f" "inj f" and iff: "\S. P (f ` S) \ Q S" - shows "locally P (f ` S) \ locally Q S" -apply (rule linear_homeomorphism_image [OF f]) -apply (rule_tac f=g and g = f in homeomorphism_locally, assumption) -by (metis iff homeomorphism_def) - -lemma locally_open_map_image: - fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" - assumes P: "locally P S" - and f: "continuous_on S f" - and oo: "\t. openin (subtopology euclidean S) t - \ openin (subtopology euclidean (f ` S)) (f ` t)" - and Q: "\t. \t \ S; P t\ \ Q(f ` t)" - shows "locally Q (f ` S)" -proof (clarsimp simp add: locally_def) - fix W y - assume oiw: "openin (subtopology euclidean (f ` S)) W" and "y \ W" - then have "W \ f ` S" by (simp add: openin_euclidean_subtopology_iff) - have oivf: "openin (subtopology euclidean S) (S \ f -` W)" - by (rule continuous_on_open [THEN iffD1, rule_format, OF f oiw]) - then obtain x where "x \ S" "f x = y" - using \W \ f ` S\ \y \ W\ by blast - then obtain U V - where "openin (subtopology euclidean S) U" "P V" "x \ U" "U \ V" "V \ S \ f -` W" - using P [unfolded locally_def, rule_format, of "(S \ f -` W)" x] oivf \y \ W\ - by auto - then show "\X. openin (subtopology euclidean (f ` S)) X \ (\Y. Q Y \ y \ X \ X \ Y \ Y \ W)" - apply (rule_tac x="f ` U" in exI) - apply (rule conjI, blast intro!: oo) - apply (rule_tac x="f ` V" in exI) - apply (force simp: \f x = y\ rev_image_eqI intro: Q) - done -qed - - -subsection\An induction principle for connected sets\ - -proposition connected_induction: - assumes "connected S" - and opD: "\T a. \openin (subtopology euclidean S) T; a \ T\ \ \z. z \ T \ P z" - and opI: "\a. a \ S - \ \T. openin (subtopology euclidean S) T \ a \ T \ - (\x \ T. \y \ T. P x \ P y \ Q x \ Q y)" - and etc: "a \ S" "b \ S" "P a" "P b" "Q a" - shows "Q b" -proof - - have 1: "openin (subtopology euclidean S) - {b. \T. openin (subtopology euclidean S) T \ - b \ T \ (\x\T. P x \ Q x)}" - apply (subst openin_subopen, clarify) - apply (rule_tac x=T in exI, auto) - done - have 2: "openin (subtopology euclidean S) - {b. \T. openin (subtopology euclidean S) T \ - b \ T \ (\x\T. P x \ \ Q x)}" - apply (subst openin_subopen, clarify) - apply (rule_tac x=T in exI, auto) - done - show ?thesis - using \connected S\ - apply (simp only: connected_openin HOL.not_ex HOL.de_Morgan_conj) - apply (elim disjE allE) - apply (blast intro: 1) - apply (blast intro: 2, simp_all) - apply clarify apply (metis opI) - using opD apply (blast intro: etc elim: dest:) - using opI etc apply meson+ - done -qed - -lemma connected_equivalence_relation_gen: - assumes "connected S" - and etc: "a \ S" "b \ S" "P a" "P b" - and trans: "\x y z. \R x y; R y z\ \ R x z" - and opD: "\T a. \openin (subtopology euclidean S) T; a \ T\ \ \z. z \ T \ P z" - and opI: "\a. a \ S - \ \T. openin (subtopology euclidean S) T \ a \ T \ - (\x \ T. \y \ T. P x \ P y \ R x y)" - shows "R a b" -proof - - have "\a b c. \a \ S; P a; b \ S; c \ S; P b; P c; R a b\ \ R a c" - apply (rule connected_induction [OF \connected S\ opD], simp_all) - by (meson trans opI) - then show ?thesis by (metis etc opI) -qed - -lemma connected_induction_simple: - assumes "connected S" - and etc: "a \ S" "b \ S" "P a" - and opI: "\a. a \ S - \ \T. openin (subtopology euclidean S) T \ a \ T \ - (\x \ T. \y \ T. P x \ P y)" - shows "P b" -apply (rule connected_induction [OF \connected S\ _, where P = "\x. True"], blast) -apply (frule opI) -using etc apply simp_all -done - -lemma connected_equivalence_relation: - assumes "connected S" - and etc: "a \ S" "b \ S" - and sym: "\x y. \R x y; x \ S; y \ S\ \ R y x" - and trans: "\x y z. \R x y; R y z; x \ S; y \ S; z \ S\ \ R x z" - and opI: "\a. a \ S \ \T. openin (subtopology euclidean S) T \ a \ T \ (\x \ T. R a x)" - shows "R a b" -proof - - have "\a b c. \a \ S; b \ S; c \ S; R a b\ \ R a c" - apply (rule connected_induction_simple [OF \connected S\], simp_all) - by (meson local.sym local.trans opI openin_imp_subset subsetCE) - then show ?thesis by (metis etc opI) -qed - -lemma locally_constant_imp_constant: - assumes "connected S" - and opI: "\a. a \ S - \ \T. openin (subtopology euclidean S) T \ a \ T \ (\x \ T. f x = f a)" - shows "f constant_on S" -proof - - have "\x y. x \ S \ y \ S \ f x = f y" - apply (rule connected_equivalence_relation [OF \connected S\], simp_all) - by (metis opI) - then show ?thesis - by (metis constant_on_def) -qed - -lemma locally_constant: - "connected S \ locally (\U. f constant_on U) S \ f constant_on S" -apply (simp add: locally_def) -apply (rule iffI) - apply (rule locally_constant_imp_constant, assumption) - apply (metis (mono_tags, hide_lams) constant_on_def constant_on_subset openin_subtopology_self) -by (meson constant_on_subset openin_imp_subset order_refl) - - -subsection\Basic properties of local compactness\ - -proposition locally_compact: - fixes s :: "'a :: metric_space set" - shows - "locally compact s \ - (\x \ s. \u v. x \ u \ u \ v \ v \ s \ - openin (subtopology euclidean s) u \ compact v)" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - apply clarify - apply (erule_tac w = "s \ ball x 1" in locallyE) - by auto -next - assume r [rule_format]: ?rhs - have *: "\u v. - openin (subtopology euclidean s) u \ - compact v \ x \ u \ u \ v \ v \ s \ T" - if "open T" "x \ s" "x \ T" for x T - proof - - obtain u v where uv: "x \ u" "u \ v" "v \ s" "compact v" "openin (subtopology euclidean s) u" - using r [OF \x \ s\] by auto - obtain e where "e>0" and e: "cball x e \ T" - using open_contains_cball \open T\ \x \ T\ by blast - show ?thesis - apply (rule_tac x="(s \ ball x e) \ u" in exI) - apply (rule_tac x="cball x e \ v" in exI) - using that \e > 0\ e uv - apply auto - done - qed - show ?lhs - apply (rule locallyI) - apply (subst (asm) openin_open) - apply (blast intro: *) - done -qed - -lemma locally_compactE: - fixes s :: "'a :: metric_space set" - assumes "locally compact s" - obtains u v where "\x. x \ s \ x \ u x \ u x \ v x \ v x \ s \ - openin (subtopology euclidean s) (u x) \ compact (v x)" -using assms -unfolding locally_compact by metis - -lemma locally_compact_alt: - fixes s :: "'a :: heine_borel set" - shows "locally compact s \ - (\x \ s. \u. x \ u \ - openin (subtopology euclidean s) u \ compact(closure u) \ closure u \ s)" -apply (simp add: locally_compact) -apply (intro ball_cong ex_cong refl iffI) -apply (metis bounded_subset closure_eq closure_mono compact_eq_bounded_closed dual_order.trans) -by (meson closure_subset compact_closure) - -lemma locally_compact_Int_cball: - fixes s :: "'a :: heine_borel set" - shows "locally compact s \ (\x \ s. \e. 0 < e \ closed(cball x e \ s))" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - apply (simp add: locally_compact openin_contains_cball) - apply (clarify | assumption | drule bspec)+ - by (metis (no_types, lifting) compact_cball compact_imp_closed compact_Int inf.absorb_iff2 inf.orderE inf_sup_aci(2)) -next - assume ?rhs - then show ?lhs - apply (simp add: locally_compact openin_contains_cball) - apply (clarify | assumption | drule bspec)+ - apply (rule_tac x="ball x e \ s" in exI, simp) - apply (rule_tac x="cball x e \ s" in exI) - using compact_eq_bounded_closed - apply auto - apply (metis open_ball le_infI1 mem_ball open_contains_cball_eq) - done -qed - -lemma locally_compact_compact: - fixes s :: "'a :: heine_borel set" - shows "locally compact s \ - (\k. k \ s \ compact k - \ (\u v. k \ u \ u \ v \ v \ s \ - openin (subtopology euclidean s) u \ compact v))" - (is "?lhs = ?rhs") -proof - assume ?lhs - then obtain u v where - uv: "\x. x \ s \ x \ u x \ u x \ v x \ v x \ s \ - openin (subtopology euclidean s) (u x) \ compact (v x)" - by (metis locally_compactE) - have *: "\u v. k \ u \ u \ v \ v \ s \ openin (subtopology euclidean s) u \ compact v" - if "k \ s" "compact k" for k - proof - - have "\C. (\c\C. openin (subtopology euclidean k) c) \ k \ \C \ - \D\C. finite D \ k \ \D" - using that by (simp add: compact_eq_openin_cover) - moreover have "\c \ (\x. k \ u x) ` k. openin (subtopology euclidean k) c" - using that by clarify (metis subsetD inf.absorb_iff2 openin_subset openin_subtopology_Int_subset topspace_euclidean_subtopology uv) - moreover have "k \ \((\x. k \ u x) ` k)" - using that by clarsimp (meson subsetCE uv) - ultimately obtain D where "D \ (\x. k \ u x) ` k" "finite D" "k \ \D" - by metis - then obtain T where T: "T \ k" "finite T" "k \ \((\x. k \ u x) ` T)" - by (metis finite_subset_image) - have Tuv: "\(u ` T) \ \(v ` T)" - using T that by (force simp: dest!: uv) - show ?thesis - apply (rule_tac x="\(u ` T)" in exI) - apply (rule_tac x="\(v ` T)" in exI) - apply (simp add: Tuv) - using T that - apply (auto simp: dest!: uv) - done - qed - show ?rhs - by (blast intro: *) -next - assume ?rhs - then show ?lhs - apply (clarsimp simp add: locally_compact) - apply (drule_tac x="{x}" in spec, simp) - done -qed - -lemma open_imp_locally_compact: - fixes s :: "'a :: heine_borel set" - assumes "open s" - shows "locally compact s" -proof - - have *: "\u v. x \ u \ u \ v \ v \ s \ openin (subtopology euclidean s) u \ compact v" - if "x \ s" for x - proof - - obtain e where "e>0" and e: "cball x e \ s" - using open_contains_cball assms \x \ s\ by blast - have ope: "openin (subtopology euclidean s) (ball x e)" - by (meson e open_ball ball_subset_cball dual_order.trans open_subset) - show ?thesis - apply (rule_tac x="ball x e" in exI) - apply (rule_tac x="cball x e" in exI) - using \e > 0\ e apply (auto simp: ope) - done - qed - show ?thesis - unfolding locally_compact - by (blast intro: *) -qed - -lemma closed_imp_locally_compact: - fixes s :: "'a :: heine_borel set" - assumes "closed s" - shows "locally compact s" -proof - - have *: "\u v. x \ u \ u \ v \ v \ s \ - openin (subtopology euclidean s) u \ compact v" - if "x \ s" for x - proof - - show ?thesis - apply (rule_tac x = "s \ ball x 1" in exI) - apply (rule_tac x = "s \ cball x 1" in exI) - using \x \ s\ assms apply auto - done - qed - show ?thesis - unfolding locally_compact - by (blast intro: *) -qed - -lemma locally_compact_UNIV: "locally compact (UNIV :: 'a :: heine_borel set)" - by (simp add: closed_imp_locally_compact) - -lemma locally_compact_Int: - fixes s :: "'a :: t2_space set" - shows "\locally compact s; locally compact t\ \ locally compact (s \ t)" -by (simp add: compact_Int locally_Int) - -lemma locally_compact_closedin: - fixes s :: "'a :: heine_borel set" - shows "\closedin (subtopology euclidean s) t; locally compact s\ - \ locally compact t" -unfolding closedin_closed -using closed_imp_locally_compact locally_compact_Int by blast - -lemma locally_compact_delete: - fixes s :: "'a :: t1_space set" - shows "locally compact s \ locally compact (s - {a})" - by (auto simp: openin_delete locally_open_subset) - -lemma locally_closed: - fixes s :: "'a :: heine_borel set" - shows "locally closed s \ locally compact s" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - apply (simp only: locally_def) - apply (erule all_forward imp_forward asm_rl exE)+ - apply (rule_tac x = "u \ ball x 1" in exI) - apply (rule_tac x = "v \ cball x 1" in exI) - apply (force intro: openin_trans) - done -next - assume ?rhs then show ?lhs - using compact_eq_bounded_closed locally_mono by blast -qed - -lemma locally_compact_openin_Un: - fixes S :: "'a::euclidean_space set" - assumes LCS: "locally compact S" and LCT:"locally compact T" - and opS: "openin (subtopology euclidean (S \ T)) S" - and opT: "openin (subtopology euclidean (S \ T)) T" - shows "locally compact (S \ T)" -proof - - have "\e>0. closed (cball x e \ (S \ T))" if "x \ S" for x - proof - - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" - using LCS \x \ S\ unfolding locally_compact_Int_cball by blast - moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \ (S \ T) \ S" - by (meson \x \ S\ opS openin_contains_cball) - then have "cball x e2 \ (S \ T) = cball x e2 \ S" - by force - ultimately show ?thesis - apply (rule_tac x="min e1 e2" in exI) - apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int) - by (metis closed_Int closed_cball inf_left_commute) - qed - moreover have "\e>0. closed (cball x e \ (S \ T))" if "x \ T" for x - proof - - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ T)" - using LCT \x \ T\ unfolding locally_compact_Int_cball by blast - moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \ (S \ T) \ T" - by (meson \x \ T\ opT openin_contains_cball) - then have "cball x e2 \ (S \ T) = cball x e2 \ T" - by force - ultimately show ?thesis - apply (rule_tac x="min e1 e2" in exI) - apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int) - by (metis closed_Int closed_cball inf_left_commute) - qed - ultimately show ?thesis - by (force simp: locally_compact_Int_cball) -qed - -lemma locally_compact_closedin_Un: - fixes S :: "'a::euclidean_space set" - assumes LCS: "locally compact S" and LCT:"locally compact T" - and clS: "closedin (subtopology euclidean (S \ T)) S" - and clT: "closedin (subtopology euclidean (S \ T)) T" - shows "locally compact (S \ T)" -proof - - have "\e>0. closed (cball x e \ (S \ T))" if "x \ S" "x \ T" for x - proof - - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" - using LCS \x \ S\ unfolding locally_compact_Int_cball by blast - moreover - obtain e2 where "e2 > 0" and e2: "closed (cball x e2 \ T)" - using LCT \x \ T\ unfolding locally_compact_Int_cball by blast - ultimately show ?thesis - apply (rule_tac x="min e1 e2" in exI) - apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int Int_Un_distrib) - by (metis closed_Int closed_Un closed_cball inf_left_commute) - qed - moreover - have "\e>0. closed (cball x e \ (S \ T))" if x: "x \ S" "x \ T" for x - proof - - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" - using LCS \x \ S\ unfolding locally_compact_Int_cball by blast - moreover - obtain e2 where "e2>0" and "cball x e2 \ (S \ T) \ S - T" - using clT x by (fastforce simp: openin_contains_cball closedin_def) - then have "closed (cball x e2 \ T)" - proof - - have "{} = T - (T - cball x e2)" - using Diff_subset Int_Diff \cball x e2 \ (S \ T) \ S - T\ by auto - then show ?thesis - by (simp add: Diff_Diff_Int inf_commute) - qed - ultimately show ?thesis - apply (rule_tac x="min e1 e2" in exI) - apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int Int_Un_distrib) - by (metis closed_Int closed_Un closed_cball inf_left_commute) - qed - moreover - have "\e>0. closed (cball x e \ (S \ T))" if x: "x \ S" "x \ T" for x - proof - - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ T)" - using LCT \x \ T\ unfolding locally_compact_Int_cball by blast - moreover - obtain e2 where "e2>0" and "cball x e2 \ (S \ T) \ S \ T - S" - using clS x by (fastforce simp: openin_contains_cball closedin_def) - then have "closed (cball x e2 \ S)" - by (metis Diff_disjoint Int_empty_right closed_empty inf.left_commute inf.orderE inf_sup_absorb) - ultimately show ?thesis - apply (rule_tac x="min e1 e2" in exI) - apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int Int_Un_distrib) - by (metis closed_Int closed_Un closed_cball inf_left_commute) - qed - ultimately show ?thesis - by (auto simp: locally_compact_Int_cball) -qed - -lemma locally_compact_Times: - fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" - shows "\locally compact S; locally compact T\ \ locally compact (S \ T)" - by (auto simp: compact_Times locally_Times) - -lemma locally_compact_compact_subopen: - fixes S :: "'a :: heine_borel set" - shows - "locally compact S \ - (\K T. K \ S \ compact K \ open T \ K \ T - \ (\U V. K \ U \ U \ V \ U \ T \ V \ S \ - openin (subtopology euclidean S) U \ compact V))" - (is "?lhs = ?rhs") -proof - assume L: ?lhs - show ?rhs - proof clarify - fix K :: "'a set" and T :: "'a set" - assume "K \ S" and "compact K" and "open T" and "K \ T" - obtain U V where "K \ U" "U \ V" "V \ S" "compact V" - and ope: "openin (subtopology euclidean S) U" - using L unfolding locally_compact_compact by (meson \K \ S\ \compact K\) - show "\U V. K \ U \ U \ V \ U \ T \ V \ S \ - openin (subtopology euclidean S) U \ compact V" - proof (intro exI conjI) - show "K \ U \ T" - by (simp add: \K \ T\ \K \ U\) - show "U \ T \ closure(U \ T)" - by (rule closure_subset) - show "closure (U \ T) \ S" - by (metis \U \ V\ \V \ S\ \compact V\ closure_closed closure_mono compact_imp_closed inf.cobounded1 subset_trans) - show "openin (subtopology euclidean S) (U \ T)" - by (simp add: \open T\ ope openin_Int_open) - show "compact (closure (U \ T))" - by (meson Int_lower1 \U \ V\ \compact V\ bounded_subset compact_closure compact_eq_bounded_closed) - qed auto - qed -next - assume ?rhs then show ?lhs - unfolding locally_compact_compact - by (metis open_openin openin_topspace subtopology_superset top.extremum topspace_euclidean_subtopology) -qed - - -subsection\Sura-Bura's results about compact components of sets\ - -proposition Sura_Bura_compact: - fixes S :: "'a::euclidean_space set" - assumes "compact S" and C: "C \ components S" - shows "C = \{T. C \ T \ openin (subtopology euclidean S) T \ - closedin (subtopology euclidean S) T}" - (is "C = \?\") -proof - obtain x where x: "C = connected_component_set S x" and "x \ S" - using C by (auto simp: components_def) - have "C \ S" - by (simp add: C in_components_subset) - have "\?\ \ connected_component_set S x" - proof (rule connected_component_maximal) - have "x \ C" - by (simp add: \x \ S\ x) - then show "x \ \?\" - by blast - have clo: "closed (\?\)" - by (simp add: \compact S\ closed_Inter closedin_compact_eq compact_imp_closed) - have False - if K1: "closedin (subtopology euclidean (\?\)) K1" and - K2: "closedin (subtopology euclidean (\?\)) K2" and - K12_Int: "K1 \ K2 = {}" and K12_Un: "K1 \ K2 = \?\" and "K1 \ {}" "K2 \ {}" - for K1 K2 - proof - - have "closed K1" "closed K2" - using closedin_closed_trans clo K1 K2 by blast+ - then obtain V1 V2 where "open V1" "open V2" "K1 \ V1" "K2 \ V2" and V12: "V1 \ V2 = {}" - using separation_normal \K1 \ K2 = {}\ by metis - have SV12_ne: "(S - (V1 \ V2)) \ (\?\) \ {}" - proof (rule compact_imp_fip) - show "compact (S - (V1 \ V2))" - by (simp add: \open V1\ \open V2\ \compact S\ compact_diff open_Un) - show clo\: "closed T" if "T \ ?\" for T - using that \compact S\ - by (force intro: closedin_closed_trans simp add: compact_imp_closed) - show "(S - (V1 \ V2)) \ \\ \ {}" if "finite \" and \: "\ \ ?\" for \ - proof - assume djo: "(S - (V1 \ V2)) \ \\ = {}" - obtain D where opeD: "openin (subtopology euclidean S) D" - and cloD: "closedin (subtopology euclidean S) D" - and "C \ D" and DV12: "D \ V1 \ V2" - proof (cases "\ = {}") - case True - with \C \ S\ djo that show ?thesis - by force - next - case False show ?thesis - proof - show ope: "openin (subtopology euclidean S) (\\)" - using openin_Inter \finite \\ False \ by blast - then show "closedin (subtopology euclidean S) (\\)" - by (meson clo\ \ closed_Inter closed_subset openin_imp_subset subset_eq) - show "C \ \\" - using \ by auto - show "\\ \ V1 \ V2" - using ope djo openin_imp_subset by fastforce - qed - qed - have "connected C" - by (simp add: x) - have "closed D" - using \compact S\ cloD closedin_closed_trans compact_imp_closed by blast - have cloV1: "closedin (subtopology euclidean D) (D \ closure V1)" - and cloV2: "closedin (subtopology euclidean D) (D \ closure V2)" - by (simp_all add: closedin_closed_Int) - moreover have "D \ closure V1 = D \ V1" "D \ closure V2 = D \ V2" - apply safe - using \D \ V1 \ V2\ \open V1\ \open V2\ V12 - apply (simp_all add: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+) - done - ultimately have cloDV1: "closedin (subtopology euclidean D) (D \ V1)" - and cloDV2: "closedin (subtopology euclidean D) (D \ V2)" - by metis+ - then obtain U1 U2 where "closed U1" "closed U2" - and D1: "D \ V1 = D \ U1" and D2: "D \ V2 = D \ U2" - by (auto simp: closedin_closed) - have "D \ U1 \ C \ {}" - proof - assume "D \ U1 \ C = {}" - then have *: "C \ D \ V2" - using D1 DV12 \C \ D\ by auto - have "\?\ \ D \ V2" - apply (rule Inter_lower) - using * apply simp - by (meson cloDV2 \open V2\ cloD closedin_trans le_inf_iff opeD openin_Int_open) - then show False - using K1 V12 \K1 \ {}\ \K1 \ V1\ closedin_imp_subset by blast - qed - moreover have "D \ U2 \ C \ {}" - proof - assume "D \ U2 \ C = {}" - then have *: "C \ D \ V1" - using D2 DV12 \C \ D\ by auto - have "\?\ \ D \ V1" - apply (rule Inter_lower) - using * apply simp - by (meson cloDV1 \open V1\ cloD closedin_trans le_inf_iff opeD openin_Int_open) - then show False - using K2 V12 \K2 \ {}\ \K2 \ V2\ closedin_imp_subset by blast - qed - ultimately show False - using \connected C\ unfolding connected_closed - apply (simp only: not_ex) - apply (drule_tac x="D \ U1" in spec) - apply (drule_tac x="D \ U2" in spec) - using \C \ D\ D1 D2 V12 DV12 \closed U1\ \closed U2\ \closed D\ - by blast - qed - qed - show False - by (metis (full_types) DiffE UnE Un_upper2 SV12_ne \K1 \ V1\ \K2 \ V2\ disjoint_iff_not_equal subsetCE sup_ge1 K12_Un) - qed - then show "connected (\?\)" - by (auto simp: connected_closedin_eq) - show "\?\ \ S" - by (fastforce simp: C in_components_subset) - qed - with x show "\?\ \ C" by simp -qed auto - - -corollary Sura_Bura_clopen_subset: - fixes S :: "'a::euclidean_space set" - assumes S: "locally compact S" and C: "C \ components S" and "compact C" - and U: "open U" "C \ U" - obtains K where "openin (subtopology euclidean S) K" "compact K" "C \ K" "K \ U" -proof (rule ccontr) - assume "\ thesis" - with that have neg: "\K. openin (subtopology euclidean S) K \ compact K \ C \ K \ K \ U" - by metis - obtain V K where "C \ V" "V \ U" "V \ K" "K \ S" "compact K" - and opeSV: "openin (subtopology euclidean S) V" - using S U \compact C\ - apply (simp add: locally_compact_compact_subopen) - by (meson C in_components_subset) - let ?\ = "{T. C \ T \ openin (subtopology euclidean K) T \ compact T \ T \ K}" - have CK: "C \ components K" - by (meson C \C \ V\ \K \ S\ \V \ K\ components_intermediate_subset subset_trans) - with \compact K\ - have "C = \{T. C \ T \ openin (subtopology euclidean K) T \ closedin (subtopology euclidean K) T}" - by (simp add: Sura_Bura_compact) - then have Ceq: "C = \?\" - by (simp add: closedin_compact_eq \compact K\) - obtain W where "open W" and W: "V = S \ W" - using opeSV by (auto simp: openin_open) - have "-(U \ W) \ \?\ \ {}" - proof (rule closed_imp_fip_compact) - show "- (U \ W) \ \\ \ {}" - if "finite \" and \: "\ \ ?\" for \ - proof (cases "\ = {}") - case True - have False if "U = UNIV" "W = UNIV" - proof - - have "V = S" - by (simp add: W \W = UNIV\) - with neg show False - using \C \ V\ \K \ S\ \V \ K\ \V \ U\ \compact K\ by auto - qed - with True show ?thesis - by auto - next - case False - show ?thesis - proof - assume "- (U \ W) \ \\ = {}" - then have FUW: "\\ \ U \ W" - by blast - have "C \ \\" - using \ by auto - moreover have "compact (\\)" - by (metis (no_types, lifting) compact_Inter False mem_Collect_eq subsetCE \) - moreover have "\\ \ K" - using False that(2) by fastforce - moreover have opeKF: "openin (subtopology euclidean K) (\\)" - using False \ \finite \\ by blast - then have opeVF: "openin (subtopology euclidean V) (\\)" - using W \K \ S\ \V \ K\ opeKF \\\ \ K\ FUW openin_subset_trans by fastforce - then have "openin (subtopology euclidean S) (\\)" - by (metis opeSV openin_trans) - moreover have "\\ \ U" - by (meson \V \ U\ opeVF dual_order.trans openin_imp_subset) - ultimately show False - using neg by blast - qed - qed - qed (use \open W\ \open U\ in auto) - with W Ceq \C \ V\ \C \ U\ show False - by auto -qed - - -corollary Sura_Bura_clopen_subset_alt: - fixes S :: "'a::euclidean_space set" - assumes S: "locally compact S" and C: "C \ components S" and "compact C" - and opeSU: "openin (subtopology euclidean S) U" and "C \ U" - obtains K where "openin (subtopology euclidean S) K" "compact K" "C \ K" "K \ U" -proof - - obtain V where "open V" "U = S \ V" - using opeSU by (auto simp: openin_open) - with \C \ U\ have "C \ V" - by auto - then show ?thesis - using Sura_Bura_clopen_subset [OF S C \compact C\ \open V\] - by (metis \U = S \ V\ inf.bounded_iff openin_imp_subset that) -qed - -corollary Sura_Bura: - fixes S :: "'a::euclidean_space set" - assumes "locally compact S" "C \ components S" "compact C" - shows "C = \ {K. C \ K \ compact K \ openin (subtopology euclidean S) K}" - (is "C = ?rhs") -proof - show "?rhs \ C" - proof (clarsimp, rule ccontr) - fix x - assume *: "\X. C \ X \ compact X \ openin (subtopology euclidean S) X \ x \ X" - and "x \ C" - obtain U V where "open U" "open V" "{x} \ U" "C \ V" "U \ V = {}" - using separation_normal [of "{x}" C] - by (metis Int_empty_left \x \ C\ \compact C\ closed_empty closed_insert compact_imp_closed insert_disjoint(1)) - have "x \ V" - using \U \ V = {}\ \{x} \ U\ by blast - then show False - by (meson "*" Sura_Bura_clopen_subset \C \ V\ \open V\ assms(1) assms(2) assms(3) subsetCE) - qed -qed blast - - -subsection\Special cases of local connectedness and path connectedness\ - -lemma locally_connected_1: - assumes - "\v x. \openin (subtopology euclidean S) v; x \ v\ - \ \u. openin (subtopology euclidean S) u \ - connected u \ x \ u \ u \ v" - shows "locally connected S" -apply (clarsimp simp add: locally_def) -apply (drule assms; blast) -done - -lemma locally_connected_2: - assumes "locally connected S" - "openin (subtopology euclidean S) t" - "x \ t" - shows "openin (subtopology euclidean S) (connected_component_set t x)" -proof - - { fix y :: 'a - let ?SS = "subtopology euclidean S" - assume 1: "openin ?SS t" - "\w x. openin ?SS w \ x \ w \ (\u. openin ?SS u \ (\v. connected v \ x \ u \ u \ v \ v \ w))" - and "connected_component t x y" - then have "y \ t" and y: "y \ connected_component_set t x" - using connected_component_subset by blast+ - obtain F where - "\x y. (\w. openin ?SS w \ (\u. connected u \ x \ w \ w \ u \ u \ y)) = (openin ?SS (F x y) \ (\u. connected u \ x \ F x y \ F x y \ u \ u \ y))" - by moura - then obtain G where - "\a A. (\U. openin ?SS U \ (\V. connected V \ a \ U \ U \ V \ V \ A)) = (openin ?SS (F a A) \ connected (G a A) \ a \ F a A \ F a A \ G a A \ G a A \ A)" - by moura - then have *: "openin ?SS (F y t) \ connected (G y t) \ y \ F y t \ F y t \ G y t \ G y t \ t" - using 1 \y \ t\ by presburger - have "G y t \ connected_component_set t y" - by (metis (no_types) * connected_component_eq_self connected_component_mono contra_subsetD) - then have "\A. openin ?SS A \ y \ A \ A \ connected_component_set t x" - by (metis (no_types) * connected_component_eq dual_order.trans y) - } - then show ?thesis - using assms openin_subopen by (force simp: locally_def) -qed - -lemma locally_connected_3: - assumes "\t x. \openin (subtopology euclidean S) t; x \ t\ - \ openin (subtopology euclidean S) - (connected_component_set t x)" - "openin (subtopology euclidean S) v" "x \ v" - shows "\u. openin (subtopology euclidean S) u \ connected u \ x \ u \ u \ v" -using assms connected_component_subset by fastforce - -lemma locally_connected: - "locally connected S \ - (\v x. openin (subtopology euclidean S) v \ x \ v - \ (\u. openin (subtopology euclidean S) u \ connected u \ x \ u \ u \ v))" -by (metis locally_connected_1 locally_connected_2 locally_connected_3) - -lemma locally_connected_open_connected_component: - "locally connected S \ - (\t x. openin (subtopology euclidean S) t \ x \ t - \ openin (subtopology euclidean S) (connected_component_set t x))" -by (metis locally_connected_1 locally_connected_2 locally_connected_3) - -lemma locally_path_connected_1: - assumes - "\v x. \openin (subtopology euclidean S) v; x \ v\ - \ \u. openin (subtopology euclidean S) u \ path_connected u \ x \ u \ u \ v" - shows "locally path_connected S" -apply (clarsimp simp add: locally_def) -apply (drule assms; blast) -done - -lemma locally_path_connected_2: - assumes "locally path_connected S" - "openin (subtopology euclidean S) t" - "x \ t" - shows "openin (subtopology euclidean S) (path_component_set t x)" -proof - - { fix y :: 'a - let ?SS = "subtopology euclidean S" - assume 1: "openin ?SS t" - "\w x. openin ?SS w \ x \ w \ (\u. openin ?SS u \ (\v. path_connected v \ x \ u \ u \ v \ v \ w))" - and "path_component t x y" - then have "y \ t" and y: "y \ path_component_set t x" - using path_component_mem(2) by blast+ - obtain F where - "\x y. (\w. openin ?SS w \ (\u. path_connected u \ x \ w \ w \ u \ u \ y)) = (openin ?SS (F x y) \ (\u. path_connected u \ x \ F x y \ F x y \ u \ u \ y))" - by moura - then obtain G where - "\a A. (\U. openin ?SS U \ (\V. path_connected V \ a \ U \ U \ V \ V \ A)) = (openin ?SS (F a A) \ path_connected (G a A) \ a \ F a A \ F a A \ G a A \ G a A \ A)" - by moura - then have *: "openin ?SS (F y t) \ path_connected (G y t) \ y \ F y t \ F y t \ G y t \ G y t \ t" - using 1 \y \ t\ by presburger - have "G y t \ path_component_set t y" - using * path_component_maximal set_rev_mp by blast - then have "\A. openin ?SS A \ y \ A \ A \ path_component_set t x" - by (metis "*" \G y t \ path_component_set t y\ dual_order.trans path_component_eq y) - } - then show ?thesis - using assms openin_subopen by (force simp: locally_def) -qed - -lemma locally_path_connected_3: - assumes "\t x. \openin (subtopology euclidean S) t; x \ t\ - \ openin (subtopology euclidean S) (path_component_set t x)" - "openin (subtopology euclidean S) v" "x \ v" - shows "\u. openin (subtopology euclidean S) u \ path_connected u \ x \ u \ u \ v" -proof - - have "path_component v x x" - by (meson assms(3) path_component_refl) - then show ?thesis - by (metis assms(1) assms(2) assms(3) mem_Collect_eq path_component_subset path_connected_path_component) -qed - -proposition locally_path_connected: - "locally path_connected S \ - (\v x. openin (subtopology euclidean S) v \ x \ v - \ (\u. openin (subtopology euclidean S) u \ path_connected u \ x \ u \ u \ v))" - by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) - -proposition locally_path_connected_open_path_component: - "locally path_connected S \ - (\t x. openin (subtopology euclidean S) t \ x \ t - \ openin (subtopology euclidean S) (path_component_set t x))" - by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) - -lemma locally_connected_open_component: - "locally connected S \ - (\t c. openin (subtopology euclidean S) t \ c \ components t - \ openin (subtopology euclidean S) c)" -by (metis components_iff locally_connected_open_connected_component) - -proposition locally_connected_im_kleinen: - "locally connected S \ - (\v x. openin (subtopology euclidean S) v \ x \ v - \ (\u. openin (subtopology euclidean S) u \ - x \ u \ u \ v \ - (\y. y \ u \ (\c. connected c \ c \ v \ x \ c \ y \ c))))" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - by (fastforce simp add: locally_connected) -next - assume ?rhs - have *: "\T. openin (subtopology euclidean S) T \ x \ T \ T \ c" - if "openin (subtopology euclidean S) t" and c: "c \ components t" and "x \ c" for t c x - proof - - from that \?rhs\ [rule_format, of t x] - obtain u where u: - "openin (subtopology euclidean S) u \ x \ u \ u \ t \ - (\y. y \ u \ (\c. connected c \ c \ t \ x \ c \ y \ c))" - using in_components_subset by auto - obtain F :: "'a set \ 'a set \ 'a" where - "\x y. (\z. z \ x \ y = connected_component_set x z) = (F x y \ x \ y = connected_component_set x (F x y))" - by moura - then have F: "F t c \ t \ c = connected_component_set t (F t c)" - by (meson components_iff c) - obtain G :: "'a set \ 'a set \ 'a" where - G: "\x y. (\z. z \ y \ z \ x) = (G x y \ y \ G x y \ x)" - by moura - have "G c u \ u \ G c u \ c" - using F by (metis (full_types) u connected_componentI connected_component_eq mem_Collect_eq that(3)) - then show ?thesis - using G u by auto - qed - show ?lhs - apply (clarsimp simp add: locally_connected_open_component) - apply (subst openin_subopen) - apply (blast intro: *) - done -qed - -proposition locally_path_connected_im_kleinen: - "locally path_connected S \ - (\v x. openin (subtopology euclidean S) v \ x \ v - \ (\u. openin (subtopology euclidean S) u \ - x \ u \ u \ v \ - (\y. y \ u \ (\p. path p \ path_image p \ v \ - pathstart p = x \ pathfinish p = y))))" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - apply (simp add: locally_path_connected path_connected_def) - apply (erule all_forward ex_forward imp_forward conjE | simp)+ - by (meson dual_order.trans) -next - assume ?rhs - have *: "\T. openin (subtopology euclidean S) T \ - x \ T \ T \ path_component_set u z" - if "openin (subtopology euclidean S) u" and "z \ u" and c: "path_component u z x" for u z x - proof - - have "x \ u" - by (meson c path_component_mem(2)) - with that \?rhs\ [rule_format, of u x] - obtain U where U: - "openin (subtopology euclidean S) U \ x \ U \ U \ u \ - (\y. y \ U \ (\p. path p \ path_image p \ u \ pathstart p = x \ pathfinish p = y))" - by blast - show ?thesis - apply (rule_tac x=U in exI) - apply (auto simp: U) - apply (metis U c path_component_trans path_component_def) - done - qed - show ?lhs - apply (clarsimp simp add: locally_path_connected_open_path_component) - apply (subst openin_subopen) - apply (blast intro: *) - done -qed - -lemma locally_path_connected_imp_locally_connected: - "locally path_connected S \ locally connected S" -using locally_mono path_connected_imp_connected by blast - -lemma locally_connected_components: - "\locally connected S; c \ components S\ \ locally connected c" -by (meson locally_connected_open_component locally_open_subset openin_subtopology_self) - -lemma locally_path_connected_components: - "\locally path_connected S; c \ components S\ \ locally path_connected c" -by (meson locally_connected_open_component locally_open_subset locally_path_connected_imp_locally_connected openin_subtopology_self) - -lemma locally_path_connected_connected_component: - "locally path_connected S \ locally path_connected (connected_component_set S x)" -by (metis components_iff connected_component_eq_empty locally_empty locally_path_connected_components) - -lemma open_imp_locally_path_connected: - fixes S :: "'a :: real_normed_vector set" - shows "open S \ locally path_connected S" -apply (rule locally_mono [of convex]) -apply (simp_all add: locally_def openin_open_eq convex_imp_path_connected) -apply (meson open_ball centre_in_ball convex_ball openE order_trans) -done - -lemma open_imp_locally_connected: - fixes S :: "'a :: real_normed_vector set" - shows "open S \ locally connected S" -by (simp add: locally_path_connected_imp_locally_connected open_imp_locally_path_connected) - -lemma locally_path_connected_UNIV: "locally path_connected (UNIV::'a :: real_normed_vector set)" - by (simp add: open_imp_locally_path_connected) - -lemma locally_connected_UNIV: "locally connected (UNIV::'a :: real_normed_vector set)" - by (simp add: open_imp_locally_connected) - -lemma openin_connected_component_locally_connected: - "locally connected S - \ openin (subtopology euclidean S) (connected_component_set S x)" -apply (simp add: locally_connected_open_connected_component) -by (metis connected_component_eq_empty connected_component_subset open_empty open_subset openin_subtopology_self) - -lemma openin_components_locally_connected: - "\locally connected S; c \ components S\ \ openin (subtopology euclidean S) c" - using locally_connected_open_component openin_subtopology_self by blast - -lemma openin_path_component_locally_path_connected: - "locally path_connected S - \ openin (subtopology euclidean S) (path_component_set S x)" -by (metis (no_types) empty_iff locally_path_connected_2 openin_subopen openin_subtopology_self path_component_eq_empty) - -lemma closedin_path_component_locally_path_connected: - "locally path_connected S - \ closedin (subtopology euclidean S) (path_component_set S x)" -apply (simp add: closedin_def path_component_subset complement_path_component_Union) -apply (rule openin_Union) -using openin_path_component_locally_path_connected by auto - -lemma convex_imp_locally_path_connected: - fixes S :: "'a:: real_normed_vector set" - shows "convex S \ locally path_connected S" -apply (clarsimp simp add: locally_path_connected) -apply (subst (asm) openin_open) -apply clarify -apply (erule (1) openE) -apply (rule_tac x = "S \ ball x e" in exI) -apply (force simp: convex_Int convex_imp_path_connected) -done - -lemma convex_imp_locally_connected: - fixes S :: "'a:: real_normed_vector set" - shows "convex S \ locally connected S" - by (simp add: locally_path_connected_imp_locally_connected convex_imp_locally_path_connected) - - -subsection\Relations between components and path components\ - -lemma path_component_eq_connected_component: - assumes "locally path_connected S" - shows "(path_component S x = connected_component S x)" -proof (cases "x \ S") - case True - have "openin (subtopology euclidean (connected_component_set S x)) (path_component_set S x)" - apply (rule openin_subset_trans [of S]) - apply (intro conjI openin_path_component_locally_path_connected [OF assms]) - using path_component_subset_connected_component apply (auto simp: connected_component_subset) - done - moreover have "closedin (subtopology euclidean (connected_component_set S x)) (path_component_set S x)" - apply (rule closedin_subset_trans [of S]) - apply (intro conjI closedin_path_component_locally_path_connected [OF assms]) - using path_component_subset_connected_component apply (auto simp: connected_component_subset) - done - ultimately have *: "path_component_set S x = connected_component_set S x" - by (metis connected_connected_component connected_clopen True path_component_eq_empty) - then show ?thesis - by blast -next - case False then show ?thesis - by (metis Collect_empty_eq_bot connected_component_eq_empty path_component_eq_empty) -qed - -lemma path_component_eq_connected_component_set: - "locally path_connected S \ (path_component_set S x = connected_component_set S x)" -by (simp add: path_component_eq_connected_component) - -lemma locally_path_connected_path_component: - "locally path_connected S \ locally path_connected (path_component_set S x)" -using locally_path_connected_connected_component path_component_eq_connected_component by fastforce - -lemma open_path_connected_component: - fixes S :: "'a :: real_normed_vector set" - shows "open S \ path_component S x = connected_component S x" -by (simp add: path_component_eq_connected_component open_imp_locally_path_connected) - -lemma open_path_connected_component_set: - fixes S :: "'a :: real_normed_vector set" - shows "open S \ path_component_set S x = connected_component_set S x" -by (simp add: open_path_connected_component) - -proposition locally_connected_quotient_image: - assumes lcS: "locally connected S" - and oo: "\T. T \ f ` S - \ openin (subtopology euclidean S) (S \ f -` T) \ - openin (subtopology euclidean (f ` S)) T" - shows "locally connected (f ` S)" -proof (clarsimp simp: locally_connected_open_component) - fix U C - assume opefSU: "openin (subtopology euclidean (f ` S)) U" and "C \ components U" - then have "C \ U" "U \ f ` S" - by (meson in_components_subset openin_imp_subset)+ - then have "openin (subtopology euclidean (f ` S)) C \ - openin (subtopology euclidean S) (S \ f -` C)" - by (auto simp: oo) - moreover have "openin (subtopology euclidean S) (S \ f -` C)" - proof (subst openin_subopen, clarify) - fix x - assume "x \ S" "f x \ C" - show "\T. openin (subtopology euclidean S) T \ x \ T \ T \ (S \ f -` C)" - proof (intro conjI exI) - show "openin (subtopology euclidean S) (connected_component_set (S \ f -` U) x)" - proof (rule ccontr) - assume **: "\ openin (subtopology euclidean S) (connected_component_set (S \ f -` U) x)" - then have "x \ (S \ f -` U)" - using \U \ f ` S\ opefSU lcS locally_connected_2 oo by blast - with ** show False - by (metis (no_types) connected_component_eq_empty empty_iff openin_subopen) - qed - next - show "x \ connected_component_set (S \ f -` U) x" - using \C \ U\ \f x \ C\ \x \ S\ by auto - next - have contf: "continuous_on S f" - by (simp add: continuous_on_open oo openin_imp_subset) - then have "continuous_on (connected_component_set (S \ f -` U) x) f" - apply (rule continuous_on_subset) - using connected_component_subset apply blast - done - then have "connected (f ` connected_component_set (S \ f -` U) x)" - by (rule connected_continuous_image [OF _ connected_connected_component]) - moreover have "f ` connected_component_set (S \ f -` U) x \ U" - using connected_component_in by blast - moreover have "C \ f ` connected_component_set (S \ f -` U) x \ {}" - using \C \ U\ \f x \ C\ \x \ S\ by fastforce - ultimately have fC: "f ` (connected_component_set (S \ f -` U) x) \ C" - by (rule components_maximal [OF \C \ components U\]) - have cUC: "connected_component_set (S \ f -` U) x \ (S \ f -` C)" - using connected_component_subset fC by blast - have "connected_component_set (S \ f -` U) x \ connected_component_set (S \ f -` C) x" - proof - - { assume "x \ connected_component_set (S \ f -` U) x" - then have ?thesis - using cUC connected_component_idemp connected_component_mono by blast } - then show ?thesis - using connected_component_eq_empty by auto - qed - also have "\ \ (S \ f -` C)" - by (rule connected_component_subset) - finally show "connected_component_set (S \ f -` U) x \ (S \ f -` C)" . - qed - qed - ultimately show "openin (subtopology euclidean (f ` S)) C" - by metis -qed - -text\The proof resembles that above but is not identical!\ -proposition locally_path_connected_quotient_image: - assumes lcS: "locally path_connected S" - and oo: "\T. T \ f ` S - \ openin (subtopology euclidean S) (S \ f -` T) \ openin (subtopology euclidean (f ` S)) T" - shows "locally path_connected (f ` S)" -proof (clarsimp simp: locally_path_connected_open_path_component) - fix U y - assume opefSU: "openin (subtopology euclidean (f ` S)) U" and "y \ U" - then have "path_component_set U y \ U" "U \ f ` S" - by (meson path_component_subset openin_imp_subset)+ - then have "openin (subtopology euclidean (f ` S)) (path_component_set U y) \ - openin (subtopology euclidean S) (S \ f -` path_component_set U y)" - proof - - have "path_component_set U y \ f ` S" - using \U \ f ` S\ \path_component_set U y \ U\ by blast - then show ?thesis - using oo by blast - qed - moreover have "openin (subtopology euclidean S) (S \ f -` path_component_set U y)" - proof (subst openin_subopen, clarify) - fix x - assume "x \ S" and Uyfx: "path_component U y (f x)" - then have "f x \ U" - using path_component_mem by blast - show "\T. openin (subtopology euclidean S) T \ x \ T \ T \ (S \ f -` path_component_set U y)" - proof (intro conjI exI) - show "openin (subtopology euclidean S) (path_component_set (S \ f -` U) x)" - proof (rule ccontr) - assume **: "\ openin (subtopology euclidean S) (path_component_set (S \ f -` U) x)" - then have "x \ (S \ f -` U)" - by (metis (no_types, lifting) \U \ f ` S\ opefSU lcS oo locally_path_connected_open_path_component) - then show False - using ** \path_component_set U y \ U\ \x \ S\ \path_component U y (f x)\ by blast - qed - next - show "x \ path_component_set (S \ f -` U) x" - by (simp add: \f x \ U\ \x \ S\ path_component_refl) - next - have contf: "continuous_on S f" - by (simp add: continuous_on_open oo openin_imp_subset) - then have "continuous_on (path_component_set (S \ f -` U) x) f" - apply (rule continuous_on_subset) - using path_component_subset apply blast - done - then have "path_connected (f ` path_component_set (S \ f -` U) x)" - by (simp add: path_connected_continuous_image) - moreover have "f ` path_component_set (S \ f -` U) x \ U" - using path_component_mem by fastforce - moreover have "f x \ f ` path_component_set (S \ f -` U) x" - by (force simp: \x \ S\ \f x \ U\ path_component_refl_eq) - ultimately have "f ` (path_component_set (S \ f -` U) x) \ path_component_set U (f x)" - by (meson path_component_maximal) - also have "\ \ path_component_set U y" - by (simp add: Uyfx path_component_maximal path_component_subset path_component_sym) - finally have fC: "f ` (path_component_set (S \ f -` U) x) \ path_component_set U y" . - have cUC: "path_component_set (S \ f -` U) x \ (S \ f -` path_component_set U y)" - using path_component_subset fC by blast - have "path_component_set (S \ f -` U) x \ path_component_set (S \ f -` path_component_set U y) x" - proof - - have "\a. path_component_set (path_component_set (S \ f -` U) x) a \ path_component_set (S \ f -` path_component_set U y) a" - using cUC path_component_mono by blast - then show ?thesis - using path_component_path_component by blast - qed - also have "\ \ (S \ f -` path_component_set U y)" - by (rule path_component_subset) - finally show "path_component_set (S \ f -` U) x \ (S \ f -` path_component_set U y)" . - qed - qed - ultimately show "openin (subtopology euclidean (f ` S)) (path_component_set U y)" - by metis -qed - -subsection%unimportant\Components, continuity, openin, closedin\ - -lemma continuous_on_components_gen: - fixes f :: "'a::topological_space \ 'b::topological_space" - assumes "\c. c \ components S \ - openin (subtopology euclidean S) c \ continuous_on c f" - shows "continuous_on S f" -proof (clarsimp simp: continuous_openin_preimage_eq) - fix t :: "'b set" - assume "open t" - have *: "S \ f -` t = (\c \ components S. c \ f -` t)" - by auto - show "openin (subtopology euclidean S) (S \ f -` t)" - unfolding * using \open t\ assms continuous_openin_preimage_gen openin_trans openin_Union by blast -qed - -lemma continuous_on_components: - fixes f :: "'a::topological_space \ 'b::topological_space" - assumes "locally connected S " - "\c. c \ components S \ continuous_on c f" - shows "continuous_on S f" -apply (rule continuous_on_components_gen) -apply (auto simp: assms intro: openin_components_locally_connected) -done - -lemma continuous_on_components_eq: - "locally connected S - \ (continuous_on S f \ (\c \ components S. continuous_on c f))" -by (meson continuous_on_components continuous_on_subset in_components_subset) - -lemma continuous_on_components_open: - fixes S :: "'a::real_normed_vector set" - assumes "open S " - "\c. c \ components S \ continuous_on c f" - shows "continuous_on S f" -using continuous_on_components open_imp_locally_connected assms by blast - -lemma continuous_on_components_open_eq: - fixes S :: "'a::real_normed_vector set" - shows "open S \ (continuous_on S f \ (\c \ components S. continuous_on c f))" -using continuous_on_subset in_components_subset -by (blast intro: continuous_on_components_open) - -lemma closedin_union_complement_components: - assumes u: "locally connected u" - and S: "closedin (subtopology euclidean u) S" - and cuS: "c \ components(u - S)" - shows "closedin (subtopology euclidean u) (S \ \c)" -proof - - have di: "(\S t. S \ c \ t \ c' \ disjnt S t) \ disjnt (\ c) (\ c')" for c' - by (simp add: disjnt_def) blast - have "S \ u" - using S closedin_imp_subset by blast - moreover have "u - S = \c \ \(components (u - S) - c)" - by (metis Diff_partition Union_components Union_Un_distrib assms(3)) - moreover have "disjnt (\c) (\(components (u - S) - c))" - apply (rule di) - by (metis DiffD1 DiffD2 assms(3) components_nonoverlap disjnt_def subsetCE) - ultimately have eq: "S \ \c = u - (\(components(u - S) - c))" - by (auto simp: disjnt_def) - have *: "openin (subtopology euclidean u) (\(components (u - S) - c))" - apply (rule openin_Union) - apply (rule openin_trans [of "u - S"]) - apply (simp add: u S locally_diff_closed openin_components_locally_connected) - apply (simp add: openin_diff S) - done - have "openin (subtopology euclidean u) (u - (u - \(components (u - S) - c)))" - apply (rule openin_diff, simp) - apply (metis closedin_diff closedin_topspace topspace_euclidean_subtopology *) - done - then show ?thesis - by (force simp: eq closedin_def) -qed - -lemma closed_union_complement_components: - fixes S :: "'a::real_normed_vector set" - assumes S: "closed S" and c: "c \ components(- S)" - shows "closed(S \ \ c)" -proof - - have "closedin (subtopology euclidean UNIV) (S \ \c)" - apply (rule closedin_union_complement_components [OF locally_connected_UNIV]) - using S c apply (simp_all add: Compl_eq_Diff_UNIV) - done - then show ?thesis by simp -qed - -lemma closedin_Un_complement_component: - fixes S :: "'a::real_normed_vector set" - assumes u: "locally connected u" - and S: "closedin (subtopology euclidean u) S" - and c: " c \ components(u - S)" - shows "closedin (subtopology euclidean u) (S \ c)" -proof - - have "closedin (subtopology euclidean u) (S \ \{c})" - using c by (blast intro: closedin_union_complement_components [OF u S]) - then show ?thesis - by simp -qed - -lemma closed_Un_complement_component: - fixes S :: "'a::real_normed_vector set" - assumes S: "closed S" and c: " c \ components(-S)" - shows "closed (S \ c)" - by (metis Compl_eq_Diff_UNIV S c closed_closedin closedin_Un_complement_component - locally_connected_UNIV subtopology_UNIV) - - -subsection\Existence of isometry between subspaces of same dimension\ - -lemma isometry_subset_subspace: - fixes S :: "'a::euclidean_space set" - and T :: "'b::euclidean_space set" - assumes S: "subspace S" - and T: "subspace T" - and d: "dim S \ dim T" - obtains f where "linear f" "f ` S \ T" "\x. x \ S \ norm(f x) = norm x" -proof - - obtain B where "B \ S" and Borth: "pairwise orthogonal B" - and B1: "\x. x \ B \ norm x = 1" - and "independent B" "finite B" "card B = dim S" "span B = S" - by (metis orthonormal_basis_subspace [OF S] independent_finite) - obtain C where "C \ T" and Corth: "pairwise orthogonal C" - and C1:"\x. x \ C \ norm x = 1" - and "independent C" "finite C" "card C = dim T" "span C = T" - by (metis orthonormal_basis_subspace [OF T] independent_finite) - obtain fb where "fb ` B \ C" "inj_on fb B" - by (metis \card B = dim S\ \card C = dim T\ \finite B\ \finite C\ card_le_inj d) - then have pairwise_orth_fb: "pairwise (\v j. orthogonal (fb v) (fb j)) B" - using Corth - apply (auto simp: pairwise_def orthogonal_clauses) - by (meson subsetD image_eqI inj_on_def) - obtain f where "linear f" and ffb: "\x. x \ B \ f x = fb x" - using linear_independent_extend \independent B\ by fastforce - have "span (f ` B) \ span C" - by (metis \fb ` B \ C\ ffb image_cong span_mono) - then have "f ` S \ T" - unfolding \span B = S\ \span C = T\ span_linear_image[OF \linear f\] . - have [simp]: "\x. x \ B \ norm (fb x) = norm x" - using B1 C1 \fb ` B \ C\ by auto - have "norm (f x) = norm x" if "x \ S" for x - proof - - interpret linear f by fact - obtain a where x: "x = (\v \ B. a v *\<^sub>R v)" - using \finite B\ \span B = S\ \x \ S\ span_finite by fastforce - have "norm (f x)^2 = norm (\v\B. a v *\<^sub>R fb v)^2" by (simp add: sum scale ffb x) - also have "\ = (\v\B. norm ((a v *\<^sub>R fb v))^2)" - apply (rule norm_sum_Pythagorean [OF \finite B\]) - apply (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) - done - also have "\ = norm x ^2" - by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \finite B\]) - finally show ?thesis - by (simp add: norm_eq_sqrt_inner) - qed - then show ?thesis - by (rule that [OF \linear f\ \f ` S \ T\]) -qed - -proposition isometries_subspaces: - fixes S :: "'a::euclidean_space set" - and T :: "'b::euclidean_space set" - assumes S: "subspace S" - and T: "subspace T" - and d: "dim S = dim T" - obtains f g where "linear f" "linear g" "f ` S = T" "g ` T = S" - "\x. x \ S \ norm(f x) = norm x" - "\x. x \ T \ norm(g x) = norm x" - "\x. x \ S \ g(f x) = x" - "\x. x \ T \ f(g x) = x" -proof - - obtain B where "B \ S" and Borth: "pairwise orthogonal B" - and B1: "\x. x \ B \ norm x = 1" - and "independent B" "finite B" "card B = dim S" "span B = S" - by (metis orthonormal_basis_subspace [OF S] independent_finite) - obtain C where "C \ T" and Corth: "pairwise orthogonal C" - and C1:"\x. x \ C \ norm x = 1" - and "independent C" "finite C" "card C = dim T" "span C = T" - by (metis orthonormal_basis_subspace [OF T] independent_finite) - obtain fb where "bij_betw fb B C" - by (metis \finite B\ \finite C\ bij_betw_iff_card \card B = dim S\ \card C = dim T\ d) - then have pairwise_orth_fb: "pairwise (\v j. orthogonal (fb v) (fb j)) B" - using Corth - apply (auto simp: pairwise_def orthogonal_clauses bij_betw_def) - by (meson subsetD image_eqI inj_on_def) - obtain f where "linear f" and ffb: "\x. x \ B \ f x = fb x" - using linear_independent_extend \independent B\ by fastforce - interpret f: linear f by fact - define gb where "gb \ inv_into B fb" - then have pairwise_orth_gb: "pairwise (\v j. orthogonal (gb v) (gb j)) C" - using Borth - apply (auto simp: pairwise_def orthogonal_clauses bij_betw_def) - by (metis \bij_betw fb B C\ bij_betw_imp_surj_on bij_betw_inv_into_right inv_into_into) - obtain g where "linear g" and ggb: "\x. x \ C \ g x = gb x" - using linear_independent_extend \independent C\ by fastforce - interpret g: linear g by fact - have "span (f ` B) \ span C" - by (metis \bij_betw fb B C\ bij_betw_imp_surj_on eq_iff ffb image_cong) - then have "f ` S \ T" - unfolding \span B = S\ \span C = T\ - span_linear_image[OF \linear f\] . - have [simp]: "\x. x \ B \ norm (fb x) = norm x" - using B1 C1 \bij_betw fb B C\ bij_betw_imp_surj_on by fastforce - have f [simp]: "norm (f x) = norm x" "g (f x) = x" if "x \ S" for x - proof - - obtain a where x: "x = (\v \ B. a v *\<^sub>R v)" - using \finite B\ \span B = S\ \x \ S\ span_finite by fastforce - have "f x = (\v \ B. f (a v *\<^sub>R v))" - using linear_sum [OF \linear f\] x by auto - also have "\ = (\v \ B. a v *\<^sub>R f v)" - by (simp add: f.sum f.scale) - also have "\ = (\v \ B. a v *\<^sub>R fb v)" - by (simp add: ffb cong: sum.cong) - finally have *: "f x = (\v\B. a v *\<^sub>R fb v)" . - then have "(norm (f x))\<^sup>2 = (norm (\v\B. a v *\<^sub>R fb v))\<^sup>2" by simp - also have "\ = (\v\B. norm ((a v *\<^sub>R fb v))^2)" - apply (rule norm_sum_Pythagorean [OF \finite B\]) - apply (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) - done - also have "\ = (norm x)\<^sup>2" - by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \finite B\]) - finally show "norm (f x) = norm x" - by (simp add: norm_eq_sqrt_inner) - have "g (f x) = g (\v\B. a v *\<^sub>R fb v)" by (simp add: *) - also have "\ = (\v\B. g (a v *\<^sub>R fb v))" - by (simp add: g.sum g.scale) - also have "\ = (\v\B. a v *\<^sub>R g (fb v))" - by (simp add: g.scale) - also have "\ = (\v\B. a v *\<^sub>R v)" - apply (rule sum.cong [OF refl]) - using \bij_betw fb B C\ gb_def bij_betwE bij_betw_inv_into_left gb_def ggb by fastforce - also have "\ = x" - using x by blast - finally show "g (f x) = x" . - qed - have [simp]: "\x. x \ C \ norm (gb x) = norm x" - by (metis B1 C1 \bij_betw fb B C\ bij_betw_imp_surj_on gb_def inv_into_into) - have g [simp]: "f (g x) = x" if "x \ T" for x - proof - - obtain a where x: "x = (\v \ C. a v *\<^sub>R v)" - using \finite C\ \span C = T\ \x \ T\ span_finite by fastforce - have "g x = (\v \ C. g (a v *\<^sub>R v))" - by (simp add: x g.sum) - also have "\ = (\v \ C. a v *\<^sub>R g v)" - by (simp add: g.scale) - also have "\ = (\v \ C. a v *\<^sub>R gb v)" - by (simp add: ggb cong: sum.cong) - finally have "f (g x) = f (\v\C. a v *\<^sub>R gb v)" by simp - also have "\ = (\v\C. f (a v *\<^sub>R gb v))" - by (simp add: f.scale f.sum) - also have "\ = (\v\C. a v *\<^sub>R f (gb v))" - by (simp add: f.scale f.sum) - also have "\ = (\v\C. a v *\<^sub>R v)" - using \bij_betw fb B C\ - by (simp add: bij_betw_def gb_def bij_betw_inv_into_right ffb inv_into_into) - also have "\ = x" - using x by blast - finally show "f (g x) = x" . - qed - have gim: "g ` T = S" - by (metis (full_types) S T \f ` S \ T\ d dim_eq_span dim_image_le f(2) g.linear_axioms - image_iff linear_subspace_image span_eq_iff subset_iff) - have fim: "f ` S = T" - using \g ` T = S\ image_iff by fastforce - have [simp]: "norm (g x) = norm x" if "x \ T" for x - using fim that by auto - show ?thesis - apply (rule that [OF \linear f\ \linear g\]) - apply (simp_all add: fim gim) - done -qed - -corollary isometry_subspaces: - fixes S :: "'a::euclidean_space set" - and T :: "'b::euclidean_space set" - assumes S: "subspace S" - and T: "subspace T" - and d: "dim S = dim T" - obtains f where "linear f" "f ` S = T" "\x. x \ S \ norm(f x) = norm x" -using isometries_subspaces [OF assms] -by metis - -corollary isomorphisms_UNIV_UNIV: - assumes "DIM('M) = DIM('N)" - obtains f::"'M::euclidean_space \'N::euclidean_space" and g - where "linear f" "linear g" - "\x. norm(f x) = norm x" "\y. norm(g y) = norm y" - "\x. g (f x) = x" "\y. f(g y) = y" - using assms by (auto intro: isometries_subspaces [of "UNIV::'M set" "UNIV::'N set"]) - -lemma homeomorphic_subspaces: - fixes S :: "'a::euclidean_space set" - and T :: "'b::euclidean_space set" - assumes S: "subspace S" - and T: "subspace T" - and d: "dim S = dim T" - shows "S homeomorphic T" -proof - - obtain f g where "linear f" "linear g" "f ` S = T" "g ` T = S" - "\x. x \ S \ g(f x) = x" "\x. x \ T \ f(g x) = x" - by (blast intro: isometries_subspaces [OF assms]) - then show ?thesis - apply (simp add: homeomorphic_def homeomorphism_def) - apply (rule_tac x=f in exI) - apply (rule_tac x=g in exI) - apply (auto simp: linear_continuous_on linear_conv_bounded_linear) - done -qed - -lemma homeomorphic_affine_sets: - assumes "affine S" "affine T" "aff_dim S = aff_dim T" - shows "S homeomorphic T" -proof (cases "S = {} \ T = {}") - case True with assms aff_dim_empty homeomorphic_empty show ?thesis - by metis -next - case False - then obtain a b where ab: "a \ S" "b \ T" by auto - then have ss: "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)" - using affine_diffs_subspace assms by blast+ - have dd: "dim ((+) (- a) ` S) = dim ((+) (- b) ` T)" - using assms ab by (simp add: aff_dim_eq_dim [OF hull_inc] image_def) - have "S homeomorphic ((+) (- a) ` S)" - by (simp add: homeomorphic_translation) - also have "\ homeomorphic ((+) (- b) ` T)" - by (rule homeomorphic_subspaces [OF ss dd]) - also have "\ homeomorphic T" - using homeomorphic_sym homeomorphic_translation by auto - finally show ?thesis . -qed - - -subsection\Retracts, in a general sense, preserve (co)homotopic triviality)\ - -locale%important Retracts = - fixes s h t k - assumes conth: "continuous_on s h" - and imh: "h ` s = t" - and contk: "continuous_on t k" - and imk: "k ` t \ s" - and idhk: "\y. y \ t \ h(k y) = y" - -begin - -lemma homotopically_trivial_retraction_gen: - assumes P: "\f. \continuous_on u f; f ` u \ t; Q f\ \ P(k \ f)" - and Q: "\f. \continuous_on u f; f ` u \ s; P f\ \ Q(h \ f)" - and Qeq: "\h k. (\x. x \ u \ h x = k x) \ Q h = Q k" - and hom: "\f g. \continuous_on u f; f ` u \ s; P f; - continuous_on u g; g ` u \ s; P g\ - \ homotopic_with P u s f g" - and contf: "continuous_on u f" and imf: "f ` u \ t" and Qf: "Q f" - and contg: "continuous_on u g" and img: "g ` u \ t" and Qg: "Q g" - shows "homotopic_with Q u t f g" -proof - - have feq: "\x. x \ u \ (h \ (k \ f)) x = f x" using idhk imf by auto - have geq: "\x. x \ u \ (h \ (k \ g)) x = g x" using idhk img by auto - have "continuous_on u (k \ f)" - using contf continuous_on_compose continuous_on_subset contk imf by blast - moreover have "(k \ f) ` u \ s" - using imf imk by fastforce - moreover have "P (k \ f)" - by (simp add: P Qf contf imf) - moreover have "continuous_on u (k \ g)" - using contg continuous_on_compose continuous_on_subset contk img by blast - moreover have "(k \ g) ` u \ s" - using img imk by fastforce - moreover have "P (k \ g)" - by (simp add: P Qg contg img) - ultimately have "homotopic_with P u s (k \ f) (k \ g)" - by (rule hom) - then have "homotopic_with Q u t (h \ (k \ f)) (h \ (k \ g))" - apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) - using Q by (auto simp: conth imh) - then show ?thesis - apply (rule homotopic_with_eq) - apply (metis feq) - apply (metis geq) - apply (metis Qeq) - done -qed - -lemma homotopically_trivial_retraction_null_gen: - assumes P: "\f. \continuous_on u f; f ` u \ t; Q f\ \ P(k \ f)" - and Q: "\f. \continuous_on u f; f ` u \ s; P f\ \ Q(h \ f)" - and Qeq: "\h k. (\x. x \ u \ h x = k x) \ Q h = Q k" - and hom: "\f. \continuous_on u f; f ` u \ s; P f\ - \ \c. homotopic_with P u s f (\x. c)" - and contf: "continuous_on u f" and imf:"f ` u \ t" and Qf: "Q f" - obtains c where "homotopic_with Q u t f (\x. c)" -proof - - have feq: "\x. x \ u \ (h \ (k \ f)) x = f x" using idhk imf by auto - have "continuous_on u (k \ f)" - using contf continuous_on_compose continuous_on_subset contk imf by blast - moreover have "(k \ f) ` u \ s" - using imf imk by fastforce - moreover have "P (k \ f)" - by (simp add: P Qf contf imf) - ultimately obtain c where "homotopic_with P u s (k \ f) (\x. c)" - by (metis hom) - then have "homotopic_with Q u t (h \ (k \ f)) (h \ (\x. c))" - apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) - using Q by (auto simp: conth imh) - then show ?thesis - apply (rule_tac c = "h c" in that) - apply (erule homotopic_with_eq) - apply (metis feq, simp) - apply (metis Qeq) - done -qed - -lemma cohomotopically_trivial_retraction_gen: - assumes P: "\f. \continuous_on t f; f ` t \ u; Q f\ \ P(f \ h)" - and Q: "\f. \continuous_on s f; f ` s \ u; P f\ \ Q(f \ k)" - and Qeq: "\h k. (\x. x \ t \ h x = k x) \ Q h = Q k" - and hom: "\f g. \continuous_on s f; f ` s \ u; P f; - continuous_on s g; g ` s \ u; P g\ - \ homotopic_with P s u f g" - and contf: "continuous_on t f" and imf: "f ` t \ u" and Qf: "Q f" - and contg: "continuous_on t g" and img: "g ` t \ u" and Qg: "Q g" - shows "homotopic_with Q t u f g" -proof - - have feq: "\x. x \ t \ (f \ h \ k) x = f x" using idhk imf by auto - have geq: "\x. x \ t \ (g \ h \ k) x = g x" using idhk img by auto - have "continuous_on s (f \ h)" - using contf conth continuous_on_compose imh by blast - moreover have "(f \ h) ` s \ u" - using imf imh by fastforce - moreover have "P (f \ h)" - by (simp add: P Qf contf imf) - moreover have "continuous_on s (g \ h)" - using contg continuous_on_compose continuous_on_subset conth imh by blast - moreover have "(g \ h) ` s \ u" - using img imh by fastforce - moreover have "P (g \ h)" - by (simp add: P Qg contg img) - ultimately have "homotopic_with P s u (f \ h) (g \ h)" - by (rule hom) - then have "homotopic_with Q t u (f \ h \ k) (g \ h \ k)" - apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) - using Q by (auto simp: contk imk) - then show ?thesis - apply (rule homotopic_with_eq) - apply (metis feq) - apply (metis geq) - apply (metis Qeq) - done -qed - -lemma cohomotopically_trivial_retraction_null_gen: - assumes P: "\f. \continuous_on t f; f ` t \ u; Q f\ \ P(f \ h)" - and Q: "\f. \continuous_on s f; f ` s \ u; P f\ \ Q(f \ k)" - and Qeq: "\h k. (\x. x \ t \ h x = k x) \ Q h = Q k" - and hom: "\f g. \continuous_on s f; f ` s \ u; P f\ - \ \c. homotopic_with P s u f (\x. c)" - and contf: "continuous_on t f" and imf: "f ` t \ u" and Qf: "Q f" - obtains c where "homotopic_with Q t u f (\x. c)" -proof - - have feq: "\x. x \ t \ (f \ h \ k) x = f x" using idhk imf by auto - have "continuous_on s (f \ h)" - using contf conth continuous_on_compose imh by blast - moreover have "(f \ h) ` s \ u" - using imf imh by fastforce - moreover have "P (f \ h)" - by (simp add: P Qf contf imf) - ultimately obtain c where "homotopic_with P s u (f \ h) (\x. c)" - by (metis hom) - then have "homotopic_with Q t u (f \ h \ k) ((\x. c) \ k)" - apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) - using Q by (auto simp: contk imk) - then show ?thesis - apply (rule_tac c = c in that) - apply (erule homotopic_with_eq) - apply (metis feq, simp) - apply (metis Qeq) - done -qed - end - -lemma simply_connected_retraction_gen: - shows "\simply_connected S; continuous_on S h; h ` S = T; - continuous_on T k; k ` T \ S; \y. y \ T \ h(k y) = y\ - \ simply_connected T" -apply (simp add: simply_connected_def path_def path_image_def homotopic_loops_def, clarify) -apply (rule Retracts.homotopically_trivial_retraction_gen - [of S h _ k _ "\p. pathfinish p = pathstart p" "\p. pathfinish p = pathstart p"]) -apply (simp_all add: Retracts_def pathfinish_def pathstart_def) -done - -lemma homeomorphic_simply_connected: - "\S homeomorphic T; simply_connected S\ \ simply_connected T" - by (auto simp: homeomorphic_def homeomorphism_def intro: simply_connected_retraction_gen) - -lemma homeomorphic_simply_connected_eq: - "S homeomorphic T \ (simply_connected S \ simply_connected T)" - by (metis homeomorphic_simply_connected homeomorphic_sym) - - -subsection\Homotopy equivalence\ - -definition%important homotopy_eqv :: "'a::topological_space set \ 'b::topological_space set \ bool" - (infix "homotopy'_eqv" 50) - where "S homotopy_eqv T \ - \f g. continuous_on S f \ f ` S \ T \ - continuous_on T g \ g ` T \ S \ - homotopic_with (\x. True) S S (g \ f) id \ - homotopic_with (\x. True) T T (f \ g) id" - -lemma homeomorphic_imp_homotopy_eqv: "S homeomorphic T \ S homotopy_eqv T" - unfolding homeomorphic_def homotopy_eqv_def homeomorphism_def - by (fastforce intro!: homotopic_with_equal continuous_on_compose) - -lemma homotopy_eqv_refl: "S homotopy_eqv S" - by (rule homeomorphic_imp_homotopy_eqv homeomorphic_refl)+ - -lemma homotopy_eqv_sym: "S homotopy_eqv T \ T homotopy_eqv S" - by (auto simp: homotopy_eqv_def) - -lemma homotopy_eqv_trans [trans]: - fixes S :: "'a::real_normed_vector set" and U :: "'c::real_normed_vector set" - assumes ST: "S homotopy_eqv T" and TU: "T homotopy_eqv U" - shows "S homotopy_eqv U" -proof - - obtain f1 g1 where f1: "continuous_on S f1" "f1 ` S \ T" - and g1: "continuous_on T g1" "g1 ` T \ S" - and hom1: "homotopic_with (\x. True) S S (g1 \ f1) id" - "homotopic_with (\x. True) T T (f1 \ g1) id" - using ST by (auto simp: homotopy_eqv_def) - obtain f2 g2 where f2: "continuous_on T f2" "f2 ` T \ U" - and g2: "continuous_on U g2" "g2 ` U \ T" - and hom2: "homotopic_with (\x. True) T T (g2 \ f2) id" - "homotopic_with (\x. True) U U (f2 \ g2) id" - using TU by (auto simp: homotopy_eqv_def) - have "homotopic_with (\f. True) S T (g2 \ f2 \ f1) (id \ f1)" - by (rule homotopic_with_compose_continuous_right hom2 f1)+ - then have "homotopic_with (\f. True) S T (g2 \ (f2 \ f1)) (id \ f1)" - by (simp add: o_assoc) - then have "homotopic_with (\x. True) S S - (g1 \ (g2 \ (f2 \ f1))) (g1 \ (id \ f1))" - by (simp add: g1 homotopic_with_compose_continuous_left) - moreover have "homotopic_with (\x. True) S S (g1 \ id \ f1) id" - using hom1 by simp - ultimately have SS: "homotopic_with (\x. True) S S (g1 \ g2 \ (f2 \ f1)) id" - apply (simp add: o_assoc) - apply (blast intro: homotopic_with_trans) - done - have "homotopic_with (\f. True) U T (f1 \ g1 \ g2) (id \ g2)" - by (rule homotopic_with_compose_continuous_right hom1 g2)+ - then have "homotopic_with (\f. True) U T (f1 \ (g1 \ g2)) (id \ g2)" - by (simp add: o_assoc) - then have "homotopic_with (\x. True) U U - (f2 \ (f1 \ (g1 \ g2))) (f2 \ (id \ g2))" - by (simp add: f2 homotopic_with_compose_continuous_left) - moreover have "homotopic_with (\x. True) U U (f2 \ id \ g2) id" - using hom2 by simp - ultimately have UU: "homotopic_with (\x. True) U U (f2 \ f1 \ (g1 \ g2)) id" - apply (simp add: o_assoc) - apply (blast intro: homotopic_with_trans) - done - show ?thesis - unfolding homotopy_eqv_def - apply (rule_tac x = "f2 \ f1" in exI) - apply (rule_tac x = "g1 \ g2" in exI) - apply (intro conjI continuous_on_compose SS UU) - using f1 f2 g1 g2 apply (force simp: elim!: continuous_on_subset)+ - done -qed - -lemma homotopy_eqv_inj_linear_image: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "linear f" "inj f" - shows "(f ` S) homotopy_eqv S" -apply (rule homeomorphic_imp_homotopy_eqv) -using assms homeomorphic_sym linear_homeomorphic_image by auto - -lemma homotopy_eqv_translation: - fixes S :: "'a::real_normed_vector set" - shows "(+) a ` S homotopy_eqv S" - apply (rule homeomorphic_imp_homotopy_eqv) - using homeomorphic_translation homeomorphic_sym by blast - -lemma homotopy_eqv_homotopic_triviality_imp: - fixes S :: "'a::real_normed_vector set" - and T :: "'b::real_normed_vector set" - and U :: "'c::real_normed_vector set" - assumes "S homotopy_eqv T" - and f: "continuous_on U f" "f ` U \ T" - and g: "continuous_on U g" "g ` U \ T" - and homUS: "\f g. \continuous_on U f; f ` U \ S; - continuous_on U g; g ` U \ S\ - \ homotopic_with (\x. True) U S f g" - shows "homotopic_with (\x. True) U T f g" -proof - - obtain h k where h: "continuous_on S h" "h ` S \ T" - and k: "continuous_on T k" "k ` T \ S" - and hom: "homotopic_with (\x. True) S S (k \ h) id" - "homotopic_with (\x. True) T T (h \ k) id" - using assms by (auto simp: homotopy_eqv_def) - have "homotopic_with (\f. True) U S (k \ f) (k \ g)" - apply (rule homUS) - using f g k - apply (safe intro!: continuous_on_compose h k f elim!: continuous_on_subset) - apply (force simp: o_def)+ - done - then have "homotopic_with (\x. True) U T (h \ (k \ f)) (h \ (k \ g))" - apply (rule homotopic_with_compose_continuous_left) - apply (simp_all add: h) - done - moreover have "homotopic_with (\x. True) U T (h \ k \ f) (id \ f)" - apply (rule homotopic_with_compose_continuous_right [where X=T and Y=T]) - apply (auto simp: hom f) - done - moreover have "homotopic_with (\x. True) U T (h \ k \ g) (id \ g)" - apply (rule homotopic_with_compose_continuous_right [where X=T and Y=T]) - apply (auto simp: hom g) - done - ultimately show "homotopic_with (\x. True) U T f g" - apply (simp add: o_assoc) - using homotopic_with_trans homotopic_with_sym by blast -qed - -lemma homotopy_eqv_homotopic_triviality: - fixes S :: "'a::real_normed_vector set" - and T :: "'b::real_normed_vector set" - and U :: "'c::real_normed_vector set" - assumes "S homotopy_eqv T" - shows "(\f g. continuous_on U f \ f ` U \ S \ - continuous_on U g \ g ` U \ S - \ homotopic_with (\x. True) U S f g) \ - (\f g. continuous_on U f \ f ` U \ T \ - continuous_on U g \ g ` U \ T - \ homotopic_with (\x. True) U T f g)" -apply (rule iffI) -apply (metis assms homotopy_eqv_homotopic_triviality_imp) -by (metis (no_types) assms homotopy_eqv_homotopic_triviality_imp homotopy_eqv_sym) - -lemma homotopy_eqv_cohomotopic_triviality_null_imp: - fixes S :: "'a::real_normed_vector set" - and T :: "'b::real_normed_vector set" - and U :: "'c::real_normed_vector set" - assumes "S homotopy_eqv T" - and f: "continuous_on T f" "f ` T \ U" - and homSU: "\f. \continuous_on S f; f ` S \ U\ - \ \c. homotopic_with (\x. True) S U f (\x. c)" - obtains c where "homotopic_with (\x. True) T U f (\x. c)" -proof - - obtain h k where h: "continuous_on S h" "h ` S \ T" - and k: "continuous_on T k" "k ` T \ S" - and hom: "homotopic_with (\x. True) S S (k \ h) id" - "homotopic_with (\x. True) T T (h \ k) id" - using assms by (auto simp: homotopy_eqv_def) - obtain c where "homotopic_with (\x. True) S U (f \ h) (\x. c)" - apply (rule exE [OF homSU [of "f \ h"]]) - apply (intro continuous_on_compose h) - using h f apply (force elim!: continuous_on_subset)+ - done - then have "homotopic_with (\x. True) T U ((f \ h) \ k) ((\x. c) \ k)" - apply (rule homotopic_with_compose_continuous_right [where X=S]) - using k by auto - moreover have "homotopic_with (\x. True) T U (f \ id) (f \ (h \ k))" - apply (rule homotopic_with_compose_continuous_left [where Y=T]) - apply (simp add: hom homotopic_with_symD) - using f apply auto - done - ultimately show ?thesis - apply (rule_tac c=c in that) - apply (simp add: o_def) - using homotopic_with_trans by blast -qed - -lemma homotopy_eqv_cohomotopic_triviality_null: - fixes S :: "'a::real_normed_vector set" - and T :: "'b::real_normed_vector set" - and U :: "'c::real_normed_vector set" - assumes "S homotopy_eqv T" - shows "(\f. continuous_on S f \ f ` S \ U - \ (\c. homotopic_with (\x. True) S U f (\x. c))) \ - (\f. continuous_on T f \ f ` T \ U - \ (\c. homotopic_with (\x. True) T U f (\x. c)))" -apply (rule iffI) -apply (metis assms homotopy_eqv_cohomotopic_triviality_null_imp) -by (metis assms homotopy_eqv_cohomotopic_triviality_null_imp homotopy_eqv_sym) - -lemma homotopy_eqv_homotopic_triviality_null_imp: - fixes S :: "'a::real_normed_vector set" - and T :: "'b::real_normed_vector set" - and U :: "'c::real_normed_vector set" - assumes "S homotopy_eqv T" - and f: "continuous_on U f" "f ` U \ T" - and homSU: "\f. \continuous_on U f; f ` U \ S\ - \ \c. homotopic_with (\x. True) U S f (\x. c)" - shows "\c. homotopic_with (\x. True) U T f (\x. c)" -proof - - obtain h k where h: "continuous_on S h" "h ` S \ T" - and k: "continuous_on T k" "k ` T \ S" - and hom: "homotopic_with (\x. True) S S (k \ h) id" - "homotopic_with (\x. True) T T (h \ k) id" - using assms by (auto simp: homotopy_eqv_def) - obtain c::'a where "homotopic_with (\x. True) U S (k \ f) (\x. c)" - apply (rule exE [OF homSU [of "k \ f"]]) - apply (intro continuous_on_compose h) - using k f apply (force elim!: continuous_on_subset)+ - done - then have "homotopic_with (\x. True) U T (h \ (k \ f)) (h \ (\x. c))" - apply (rule homotopic_with_compose_continuous_left [where Y=S]) - using h by auto - moreover have "homotopic_with (\x. True) U T (id \ f) ((h \ k) \ f)" - apply (rule homotopic_with_compose_continuous_right [where X=T]) - apply (simp add: hom homotopic_with_symD) - using f apply auto - done - ultimately show ?thesis - using homotopic_with_trans by (fastforce simp add: o_def) -qed - -lemma homotopy_eqv_homotopic_triviality_null: - fixes S :: "'a::real_normed_vector set" - and T :: "'b::real_normed_vector set" - and U :: "'c::real_normed_vector set" - assumes "S homotopy_eqv T" - shows "(\f. continuous_on U f \ f ` U \ S - \ (\c. homotopic_with (\x. True) U S f (\x. c))) \ - (\f. continuous_on U f \ f ` U \ T - \ (\c. homotopic_with (\x. True) U T f (\x. c)))" -apply (rule iffI) -apply (metis assms homotopy_eqv_homotopic_triviality_null_imp) -by (metis assms homotopy_eqv_homotopic_triviality_null_imp homotopy_eqv_sym) - -lemma homotopy_eqv_contractible_sets: - fixes S :: "'a::real_normed_vector set" - and T :: "'b::real_normed_vector set" - assumes "contractible S" "contractible T" "S = {} \ T = {}" - shows "S homotopy_eqv T" -proof (cases "S = {}") - case True with assms show ?thesis - by (simp add: homeomorphic_imp_homotopy_eqv) -next - case False - with assms obtain a b where "a \ S" "b \ T" - by auto - then show ?thesis - unfolding homotopy_eqv_def - apply (rule_tac x="\x. b" in exI) - apply (rule_tac x="\x. a" in exI) - apply (intro assms conjI continuous_on_id' homotopic_into_contractible) - apply (auto simp: o_def continuous_on_const) - done -qed - -lemma homotopy_eqv_empty1 [simp]: - fixes S :: "'a::real_normed_vector set" - shows "S homotopy_eqv ({}::'b::real_normed_vector set) \ S = {}" -apply (rule iffI) -using homotopy_eqv_def apply fastforce -by (simp add: homotopy_eqv_contractible_sets) - -lemma homotopy_eqv_empty2 [simp]: - fixes S :: "'a::real_normed_vector set" - shows "({}::'b::real_normed_vector set) homotopy_eqv S \ S = {}" -by (metis homotopy_eqv_empty1 homotopy_eqv_sym) - -lemma homotopy_eqv_contractibility: - fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" - shows "S homotopy_eqv T \ (contractible S \ contractible T)" -unfolding homotopy_eqv_def -by (blast intro: homotopy_dominated_contractibility) - -lemma homotopy_eqv_sing: - fixes S :: "'a::real_normed_vector set" and a :: "'b::real_normed_vector" - shows "S homotopy_eqv {a} \ S \ {} \ contractible S" -proof (cases "S = {}") - case True then show ?thesis - by simp -next - case False then show ?thesis - by (metis contractible_sing empty_not_insert homotopy_eqv_contractibility homotopy_eqv_contractible_sets) -qed - -lemma homeomorphic_contractible_eq: - fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" - shows "S homeomorphic T \ (contractible S \ contractible T)" -by (simp add: homeomorphic_imp_homotopy_eqv homotopy_eqv_contractibility) - -lemma homeomorphic_contractible: - fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" - shows "\contractible S; S homeomorphic T\ \ contractible T" - by (metis homeomorphic_contractible_eq) - - -subsection%unimportant\Misc other results\ - -lemma bounded_connected_Compl_real: - fixes S :: "real set" - assumes "bounded S" and conn: "connected(- S)" - shows "S = {}" -proof - - obtain a b where "S \ box a b" - by (meson assms bounded_subset_box_symmetric) - then have "a \ S" "b \ S" - by auto - then have "\x. a \ x \ x \ b \ x \ - S" - by (meson Compl_iff conn connected_iff_interval) - then show ?thesis - using \S \ box a b\ by auto -qed - -lemma bounded_connected_Compl_1: - fixes S :: "'a::{euclidean_space} set" - assumes "bounded S" and conn: "connected(- S)" and 1: "DIM('a) = 1" - shows "S = {}" -proof - - have "DIM('a) = DIM(real)" - by (simp add: "1") - then obtain f::"'a \ real" and g - where "linear f" "\x. norm(f x) = norm x" "\x. g(f x) = x" "\y. f(g y) = y" - by (rule isomorphisms_UNIV_UNIV) blast - with \bounded S\ have "bounded (f ` S)" - using bounded_linear_image linear_linear by blast - have "connected (f ` (-S))" - using connected_linear_image assms \linear f\ by blast - moreover have "f ` (-S) = - (f ` S)" - apply (rule bij_image_Compl_eq) - apply (auto simp: bij_def) - apply (metis \\x. g (f x) = x\ injI) - by (metis UNIV_I \\y. f (g y) = y\ image_iff) - finally have "connected (- (f ` S))" - by simp - then have "f ` S = {}" - using \bounded (f ` S)\ bounded_connected_Compl_real by blast - then show ?thesis - by blast -qed - - -subsection%unimportant\Some Uncountable Sets\ - -lemma uncountable_closed_segment: - fixes a :: "'a::real_normed_vector" - assumes "a \ b" shows "uncountable (closed_segment a b)" -unfolding path_image_linepath [symmetric] path_image_def - using inj_on_linepath [OF assms] uncountable_closed_interval [of 0 1] - countable_image_inj_on by auto - -lemma uncountable_open_segment: - fixes a :: "'a::real_normed_vector" - assumes "a \ b" shows "uncountable (open_segment a b)" - by (simp add: assms open_segment_def uncountable_closed_segment uncountable_minus_countable) - -lemma uncountable_convex: - fixes a :: "'a::real_normed_vector" - assumes "convex S" "a \ S" "b \ S" "a \ b" - shows "uncountable S" -proof - - have "uncountable (closed_segment a b)" - by (simp add: uncountable_closed_segment assms) - then show ?thesis - by (meson assms convex_contains_segment countable_subset) -qed - -lemma uncountable_ball: - fixes a :: "'a::euclidean_space" - assumes "r > 0" - shows "uncountable (ball a r)" -proof - - have "uncountable (open_segment a (a + r *\<^sub>R (SOME i. i \ Basis)))" - by (metis Basis_zero SOME_Basis add_cancel_right_right assms less_le scale_eq_0_iff uncountable_open_segment) - moreover have "open_segment a (a + r *\<^sub>R (SOME i. i \ Basis)) \ ball a r" - using assms by (auto simp: in_segment algebra_simps dist_norm SOME_Basis) - ultimately show ?thesis - by (metis countable_subset) -qed - -lemma ball_minus_countable_nonempty: - assumes "countable (A :: 'a :: euclidean_space set)" "r > 0" - shows "ball z r - A \ {}" -proof - assume *: "ball z r - A = {}" - have "uncountable (ball z r - A)" - by (intro uncountable_minus_countable assms uncountable_ball) - thus False by (subst (asm) *) auto -qed - -lemma uncountable_cball: - fixes a :: "'a::euclidean_space" - assumes "r > 0" - shows "uncountable (cball a r)" - using assms countable_subset uncountable_ball by auto - -lemma pairwise_disjnt_countable: - fixes \ :: "nat set set" - assumes "pairwise disjnt \" - shows "countable \" -proof - - have "inj_on (\X. SOME n. n \ X) (\ - {{}})" - apply (clarsimp simp add: inj_on_def) - by (metis assms disjnt_insert2 insert_absorb pairwise_def subsetI subset_empty tfl_some) - then show ?thesis - by (metis countable_Diff_eq countable_def) -qed - -lemma pairwise_disjnt_countable_Union: - assumes "countable (\\)" and pwd: "pairwise disjnt \" - shows "countable \" -proof - - obtain f :: "_ \ nat" where f: "inj_on f (\\)" - using assms by blast - then have "pairwise disjnt (\ X \ \. {f ` X})" - using assms by (force simp: pairwise_def disjnt_inj_on_iff [OF f]) - then have "countable (\ X \ \. {f ` X})" - using pairwise_disjnt_countable by blast - then show ?thesis - by (meson pwd countable_image_inj_on disjoint_image f inj_on_image pairwise_disjnt_countable) -qed - -lemma connected_uncountable: - fixes S :: "'a::metric_space set" - assumes "connected S" "a \ S" "b \ S" "a \ b" shows "uncountable S" -proof - - have "continuous_on S (dist a)" - by (intro continuous_intros) - then have "connected (dist a ` S)" - by (metis connected_continuous_image \connected S\) - then have "closed_segment 0 (dist a b) \ (dist a ` S)" - by (simp add: assms closed_segment_subset is_interval_connected_1 is_interval_convex) - then have "uncountable (dist a ` S)" - by (metis \a \ b\ countable_subset dist_eq_0_iff uncountable_closed_segment) - then show ?thesis - by blast -qed - -lemma path_connected_uncountable: - fixes S :: "'a::metric_space set" - assumes "path_connected S" "a \ S" "b \ S" "a \ b" shows "uncountable S" - using path_connected_imp_connected assms connected_uncountable by metis - -lemma connected_finite_iff_sing: - fixes S :: "'a::metric_space set" - assumes "connected S" - shows "finite S \ S = {} \ (\a. S = {a})" (is "_ = ?rhs") -proof - - have "uncountable S" if "\ ?rhs" - using connected_uncountable assms that by blast - then show ?thesis - using uncountable_infinite by auto -qed - -lemma connected_card_eq_iff_nontrivial: - fixes S :: "'a::metric_space set" - shows "connected S \ uncountable S \ \(\a. S \ {a})" - apply (auto simp: countable_finite finite_subset) - by (metis connected_uncountable is_singletonI' is_singleton_the_elem subset_singleton_iff) - -lemma simple_path_image_uncountable: - fixes g :: "real \ 'a::metric_space" - assumes "simple_path g" - shows "uncountable (path_image g)" -proof - - have "g 0 \ path_image g" "g (1/2) \ path_image g" - by (simp_all add: path_defs) - moreover have "g 0 \ g (1/2)" - using assms by (fastforce simp add: simple_path_def) - ultimately show ?thesis - apply (simp add: assms connected_card_eq_iff_nontrivial connected_simple_path_image) - by blast -qed - -lemma arc_image_uncountable: - fixes g :: "real \ 'a::metric_space" - assumes "arc g" - shows "uncountable (path_image g)" - by (simp add: arc_imp_simple_path assms simple_path_image_uncountable) - - -subsection%unimportant\ Some simple positive connection theorems\ - -proposition path_connected_convex_diff_countable: - fixes U :: "'a::euclidean_space set" - assumes "convex U" "\ collinear U" "countable S" - shows "path_connected(U - S)" -proof (clarsimp simp add: path_connected_def) - fix a b - assume "a \ U" "a \ S" "b \ U" "b \ S" - let ?m = "midpoint a b" - show "\g. path g \ path_image g \ U - S \ pathstart g = a \ pathfinish g = b" - proof (cases "a = b") - case True - then show ?thesis - by (metis DiffI \a \ U\ \a \ S\ path_component_def path_component_refl) - next - case False - then have "a \ ?m" "b \ ?m" - using midpoint_eq_endpoint by fastforce+ - have "?m \ U" - using \a \ U\ \b \ U\ \convex U\ convex_contains_segment by force - obtain c where "c \ U" and nc_abc: "\ collinear {a,b,c}" - by (metis False \a \ U\ \b \ U\ \\ collinear U\ collinear_triples insert_absorb) - have ncoll_mca: "\ collinear {?m,c,a}" - by (metis (full_types) \a \ ?m\ collinear_3_trans collinear_midpoint insert_commute nc_abc) - have ncoll_mcb: "\ collinear {?m,c,b}" - by (metis (full_types) \b \ ?m\ collinear_3_trans collinear_midpoint insert_commute nc_abc) - have "c \ ?m" - by (metis collinear_midpoint insert_commute nc_abc) - then have "closed_segment ?m c \ U" - by (simp add: \c \ U\ \?m \ U\ \convex U\ closed_segment_subset) - then obtain z where z: "z \ closed_segment ?m c" - and disjS: "(closed_segment a z \ closed_segment z b) \ S = {}" - proof - - have False if "closed_segment ?m c \ {z. (closed_segment a z \ closed_segment z b) \ S \ {}}" - proof - - have closb: "closed_segment ?m c \ - {z \ closed_segment ?m c. closed_segment a z \ S \ {}} \ {z \ closed_segment ?m c. closed_segment z b \ S \ {}}" - using that by blast - have *: "countable {z \ closed_segment ?m c. closed_segment z u \ S \ {}}" - if "u \ U" "u \ S" and ncoll: "\ collinear {?m, c, u}" for u - proof - - have **: False if x1: "x1 \ closed_segment ?m c" and x2: "x2 \ closed_segment ?m c" - and "x1 \ x2" "x1 \ u" - and w: "w \ closed_segment x1 u" "w \ closed_segment x2 u" - and "w \ S" for x1 x2 w - proof - - have "x1 \ affine hull {?m,c}" "x2 \ affine hull {?m,c}" - using segment_as_ball x1 x2 by auto - then have coll_x1: "collinear {x1, ?m, c}" and coll_x2: "collinear {?m, c, x2}" - by (simp_all add: affine_hull_3_imp_collinear) (metis affine_hull_3_imp_collinear insert_commute) - have "\ collinear {x1, u, x2}" - proof - assume "collinear {x1, u, x2}" - then have "collinear {?m, c, u}" - by (metis (full_types) \c \ ?m\ coll_x1 coll_x2 collinear_3_trans insert_commute ncoll \x1 \ x2\) - with ncoll show False .. - qed - then have "closed_segment x1 u \ closed_segment u x2 = {u}" - by (blast intro!: Int_closed_segment) - then have "w = u" - using closed_segment_commute w by auto - show ?thesis - using \u \ S\ \w = u\ that(7) by auto - qed - then have disj: "disjoint ((\z\closed_segment ?m c. {closed_segment z u \ S}))" - by (fastforce simp: pairwise_def disjnt_def) - have cou: "countable ((\z \ closed_segment ?m c. {closed_segment z u \ S}) - {{}})" - apply (rule pairwise_disjnt_countable_Union [OF _ pairwise_subset [OF disj]]) - apply (rule countable_subset [OF _ \countable S\], auto) - done - define f where "f \ \X. (THE z. z \ closed_segment ?m c \ X = closed_segment z u \ S)" - show ?thesis - proof (rule countable_subset [OF _ countable_image [OF cou, where f=f]], clarify) - fix x - assume x: "x \ closed_segment ?m c" "closed_segment x u \ S \ {}" - show "x \ f ` ((\z\closed_segment ?m c. {closed_segment z u \ S}) - {{}})" - proof (rule_tac x="closed_segment x u \ S" in image_eqI) - show "x = f (closed_segment x u \ S)" - unfolding f_def - apply (rule the_equality [symmetric]) - using x apply (auto simp: dest: **) - done - qed (use x in auto) - qed - qed - have "uncountable (closed_segment ?m c)" - by (metis \c \ ?m\ uncountable_closed_segment) - then show False - using closb * [OF \a \ U\ \a \ S\ ncoll_mca] * [OF \b \ U\ \b \ S\ ncoll_mcb] - apply (simp add: closed_segment_commute) - by (simp add: countable_subset) - qed - then show ?thesis - by (force intro: that) - qed - show ?thesis - proof (intro exI conjI) - have "path_image (linepath a z +++ linepath z b) \ U" - by (metis \a \ U\ \b \ U\ \closed_segment ?m c \ U\ z \convex U\ closed_segment_subset contra_subsetD path_image_linepath subset_path_image_join) - with disjS show "path_image (linepath a z +++ linepath z b) \ U - S" - by (force simp: path_image_join) - qed auto - qed -qed - - -corollary connected_convex_diff_countable: - fixes U :: "'a::euclidean_space set" - assumes "convex U" "\ collinear U" "countable S" - shows "connected(U - S)" - by (simp add: assms path_connected_convex_diff_countable path_connected_imp_connected) - -lemma path_connected_punctured_convex: - assumes "convex S" and aff: "aff_dim S \ 1" - shows "path_connected(S - {a})" -proof - - consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S \ 2" - using assms aff_dim_geq [of S] by linarith - then show ?thesis - proof cases - assume "aff_dim S = -1" - then show ?thesis - by (metis aff_dim_empty empty_Diff path_connected_empty) - next - assume "aff_dim S = 0" - then show ?thesis - by (metis aff_dim_eq_0 Diff_cancel Diff_empty Diff_insert0 convex_empty convex_imp_path_connected path_connected_singleton singletonD) - next - assume ge2: "aff_dim S \ 2" - then have "\ collinear S" - proof (clarsimp simp add: collinear_affine_hull) - fix u v - assume "S \ affine hull {u, v}" - then have "aff_dim S \ aff_dim {u, v}" - by (metis (no_types) aff_dim_affine_hull aff_dim_subset) - with ge2 show False - by (metis (no_types) aff_dim_2 antisym aff not_numeral_le_zero one_le_numeral order_trans) - qed - then show ?thesis - apply (rule path_connected_convex_diff_countable [OF \convex S\]) - by simp - qed -qed - -lemma connected_punctured_convex: - shows "\convex S; aff_dim S \ 1\ \ connected(S - {a})" - using path_connected_imp_connected path_connected_punctured_convex by blast - -lemma path_connected_complement_countable: - fixes S :: "'a::euclidean_space set" - assumes "2 \ DIM('a)" "countable S" - shows "path_connected(- S)" -proof - - have "path_connected(UNIV - S)" - apply (rule path_connected_convex_diff_countable) - using assms by (auto simp: collinear_aff_dim [of "UNIV :: 'a set"]) - then show ?thesis - by (simp add: Compl_eq_Diff_UNIV) -qed - -proposition path_connected_openin_diff_countable: - fixes S :: "'a::euclidean_space set" - assumes "connected S" and ope: "openin (subtopology euclidean (affine hull S)) S" - and "\ collinear S" "countable T" - shows "path_connected(S - T)" -proof (clarsimp simp add: path_connected_component) - fix x y - assume xy: "x \ S" "x \ T" "y \ S" "y \ T" - show "path_component (S - T) x y" - proof (rule connected_equivalence_relation_gen [OF \connected S\, where P = "\x. x \ T"]) - show "\z. z \ U \ z \ T" if opeU: "openin (subtopology euclidean S) U" and "x \ U" for U x - proof - - have "openin (subtopology euclidean (affine hull S)) U" - using opeU ope openin_trans by blast - with \x \ U\ obtain r where Usub: "U \ affine hull S" and "r > 0" - and subU: "ball x r \ affine hull S \ U" - by (auto simp: openin_contains_ball) - with \x \ U\ have x: "x \ ball x r \ affine hull S" - by auto - have "\ S \ {x}" - using \\ collinear S\ collinear_subset by blast - then obtain x' where "x' \ x" "x' \ S" - by blast - obtain y where y: "y \ x" "y \ ball x r \ affine hull S" - proof - show "x + (r / 2 / norm(x' - x)) *\<^sub>R (x' - x) \ x" - using \x' \ x\ \r > 0\ by auto - show "x + (r / 2 / norm (x' - x)) *\<^sub>R (x' - x) \ ball x r \ affine hull S" - using \x' \ x\ \r > 0\ \x' \ S\ x - by (simp add: dist_norm mem_affine_3_minus hull_inc) - qed - have "convex (ball x r \ affine hull S)" - by (simp add: affine_imp_convex convex_Int) - with x y subU have "uncountable U" - by (meson countable_subset uncountable_convex) - then have "\ U \ T" - using \countable T\ countable_subset by blast - then show ?thesis by blast - qed - show "\U. openin (subtopology euclidean S) U \ x \ U \ - (\x\U. \y\U. x \ T \ y \ T \ path_component (S - T) x y)" - if "x \ S" for x - proof - - obtain r where Ssub: "S \ affine hull S" and "r > 0" - and subS: "ball x r \ affine hull S \ S" - using ope \x \ S\ by (auto simp: openin_contains_ball) - then have conv: "convex (ball x r \ affine hull S)" - by (simp add: affine_imp_convex convex_Int) - have "\ aff_dim (affine hull S) \ 1" - using \\ collinear S\ collinear_aff_dim by auto - then have "\ collinear (ball x r \ affine hull S)" - apply (simp add: collinear_aff_dim) - by (metis (no_types, hide_lams) aff_dim_convex_Int_open IntI open_ball \0 < r\ aff_dim_affine_hull affine_affine_hull affine_imp_convex centre_in_ball empty_iff hull_subset inf_commute subsetCE that) - then have *: "path_connected ((ball x r \ affine hull S) - T)" - by (rule path_connected_convex_diff_countable [OF conv _ \countable T\]) - have ST: "ball x r \ affine hull S - T \ S - T" - using subS by auto - show ?thesis - proof (intro exI conjI) - show "x \ ball x r \ affine hull S" - using \x \ S\ \r > 0\ by (simp add: hull_inc) - have "openin (subtopology euclidean (affine hull S)) (ball x r \ affine hull S)" - by (subst inf.commute) (simp add: openin_Int_open) - then show "openin (subtopology euclidean S) (ball x r \ affine hull S)" - by (rule openin_subset_trans [OF _ subS Ssub]) - qed (use * path_component_trans in \auto simp: path_connected_component path_component_of_subset [OF ST]\) - qed - qed (use xy path_component_trans in auto) -qed - -corollary connected_openin_diff_countable: - fixes S :: "'a::euclidean_space set" - assumes "connected S" and ope: "openin (subtopology euclidean (affine hull S)) S" - and "\ collinear S" "countable T" - shows "connected(S - T)" - by (metis path_connected_imp_connected path_connected_openin_diff_countable [OF assms]) - -corollary path_connected_open_diff_countable: - fixes S :: "'a::euclidean_space set" - assumes "2 \ DIM('a)" "open S" "connected S" "countable T" - shows "path_connected(S - T)" -proof (cases "S = {}") - case True - then show ?thesis - by (simp add: path_connected_empty) -next - case False - show ?thesis - proof (rule path_connected_openin_diff_countable) - show "openin (subtopology euclidean (affine hull S)) S" - by (simp add: assms hull_subset open_subset) - show "\ collinear S" - using assms False by (simp add: collinear_aff_dim aff_dim_open) - qed (simp_all add: assms) -qed - -corollary connected_open_diff_countable: - fixes S :: "'a::euclidean_space set" - assumes "2 \ DIM('a)" "open S" "connected S" "countable T" - shows "connected(S - T)" -by (simp add: assms path_connected_imp_connected path_connected_open_diff_countable) - - - -subsection%unimportant \Self-homeomorphisms shuffling points about\ - -subsubsection%unimportant\The theorem \homeomorphism_moving_points_exists\\ - -lemma homeomorphism_moving_point_1: - fixes a :: "'a::euclidean_space" - assumes "affine T" "a \ T" and u: "u \ ball a r \ T" - obtains f g where "homeomorphism (cball a r \ T) (cball a r \ T) f g" - "f a = u" "\x. x \ sphere a r \ f x = x" -proof - - have nou: "norm (u - a) < r" and "u \ T" - using u by (auto simp: dist_norm norm_minus_commute) - then have "0 < r" - by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) - define f where "f \ \x. (1 - norm(x - a) / r) *\<^sub>R (u - a) + x" - have *: "False" if eq: "x + (norm y / r) *\<^sub>R u = y + (norm x / r) *\<^sub>R u" - and nou: "norm u < r" and yx: "norm y < norm x" for x y and u::'a - proof - - have "x = y + (norm x / r - (norm y / r)) *\<^sub>R u" - using eq by (simp add: algebra_simps) - then have "norm x = norm (y + ((norm x - norm y) / r) *\<^sub>R u)" - by (metis diff_divide_distrib) - also have "\ \ norm y + norm(((norm x - norm y) / r) *\<^sub>R u)" - using norm_triangle_ineq by blast - also have "\ = norm y + (norm x - norm y) * (norm u / r)" - using yx \r > 0\ - by (simp add: divide_simps) - also have "\ < norm y + (norm x - norm y) * 1" - apply (subst add_less_cancel_left) - apply (rule mult_strict_left_mono) - using nou \0 < r\ yx - apply (simp_all add: field_simps) - done - also have "\ = norm x" - by simp - finally show False by simp - qed - have "inj f" - unfolding f_def - proof (clarsimp simp: inj_on_def) - fix x y - assume "(1 - norm (x - a) / r) *\<^sub>R (u - a) + x = - (1 - norm (y - a) / r) *\<^sub>R (u - a) + y" - then have eq: "(x - a) + (norm (y - a) / r) *\<^sub>R (u - a) = (y - a) + (norm (x - a) / r) *\<^sub>R (u - a)" - by (auto simp: algebra_simps) - show "x=y" - proof (cases "norm (x - a) = norm (y - a)") - case True - then show ?thesis - using eq by auto - next - case False - then consider "norm (x - a) < norm (y - a)" | "norm (x - a) > norm (y - a)" - by linarith - then have "False" - proof cases - case 1 show False - using * [OF _ nou 1] eq by simp - next - case 2 with * [OF eq nou] show False - by auto - qed - then show "x=y" .. - qed - qed - then have inj_onf: "inj_on f (cball a r \ T)" - using inj_on_Int by fastforce - have contf: "continuous_on (cball a r \ T) f" - unfolding f_def using \0 < r\ by (intro continuous_intros) blast - have fim: "f ` (cball a r \ T) = cball a r \ T" - proof - have *: "norm (y + (1 - norm y / r) *\<^sub>R u) \ r" if "norm y \ r" "norm u < r" for y u::'a - proof - - have "norm (y + (1 - norm y / r) *\<^sub>R u) \ norm y + norm((1 - norm y / r) *\<^sub>R u)" - using norm_triangle_ineq by blast - also have "\ = norm y + abs(1 - norm y / r) * norm u" - by simp - also have "\ \ r" - proof - - have "(r - norm u) * (r - norm y) \ 0" - using that by auto - then have "r * norm u + r * norm y \ r * r + norm u * norm y" - by (simp add: algebra_simps) - then show ?thesis - using that \0 < r\ by (simp add: abs_if field_simps) - qed - finally show ?thesis . - qed - have "f ` (cball a r) \ cball a r" - apply (clarsimp simp add: dist_norm norm_minus_commute f_def) - using * by (metis diff_add_eq diff_diff_add diff_diff_eq2 norm_minus_commute nou) - moreover have "f ` T \ T" - unfolding f_def using \affine T\ \a \ T\ \u \ T\ - by (force simp: add.commute mem_affine_3_minus) - ultimately show "f ` (cball a r \ T) \ cball a r \ T" - by blast - next - show "cball a r \ T \ f ` (cball a r \ T)" - proof (clarsimp simp add: dist_norm norm_minus_commute) - fix x - assume x: "norm (x - a) \ r" and "x \ T" - have "\v \ {0..1}. ((1 - v) * r - norm ((x - a) - v *\<^sub>R (u - a))) \ 1 = 0" - by (rule ivt_decreasing_component_on_1) (auto simp: x continuous_intros) - then obtain v where "0\v" "v\1" and v: "(1 - v) * r = norm ((x - a) - v *\<^sub>R (u - a))" - by auto - show "x \ f ` (cball a r \ T)" - proof (rule image_eqI) - show "x = f (x - v *\<^sub>R (u - a))" - using \r > 0\ v by (simp add: f_def field_simps) - have "x - v *\<^sub>R (u - a) \ cball a r" - using \r > 0\ v \0 \ v\ - apply (simp add: field_simps dist_norm norm_minus_commute) - by (metis le_add_same_cancel2 order.order_iff_strict zero_le_mult_iff) - moreover have "x - v *\<^sub>R (u - a) \ T" - by (simp add: f_def \affine T\ \u \ T\ \x \ T\ assms mem_affine_3_minus2) - ultimately show "x - v *\<^sub>R (u - a) \ cball a r \ T" - by blast - qed - qed - qed - have "\g. homeomorphism (cball a r \ T) (cball a r \ T) f g" - apply (rule homeomorphism_compact [OF _ contf fim inj_onf]) - apply (simp add: affine_closed compact_Int_closed \affine T\) - done - then show ?thesis - apply (rule exE) - apply (erule_tac f=f in that) - using \r > 0\ - apply (simp_all add: f_def dist_norm norm_minus_commute) - done -qed - -corollary%unimportant homeomorphism_moving_point_2: - fixes a :: "'a::euclidean_space" - assumes "affine T" "a \ T" and u: "u \ ball a r \ T" and v: "v \ ball a r \ T" - obtains f g where "homeomorphism (cball a r \ T) (cball a r \ T) f g" - "f u = v" "\x. \x \ sphere a r; x \ T\ \ f x = x" -proof - - have "0 < r" - by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) - obtain f1 g1 where hom1: "homeomorphism (cball a r \ T) (cball a r \ T) f1 g1" - and "f1 a = u" and f1: "\x. x \ sphere a r \ f1 x = x" - using homeomorphism_moving_point_1 [OF \affine T\ \a \ T\ u] by blast - obtain f2 g2 where hom2: "homeomorphism (cball a r \ T) (cball a r \ T) f2 g2" - and "f2 a = v" and f2: "\x. x \ sphere a r \ f2 x = x" - using homeomorphism_moving_point_1 [OF \affine T\ \a \ T\ v] by blast - show ?thesis - proof - show "homeomorphism (cball a r \ T) (cball a r \ T) (f2 \ g1) (f1 \ g2)" - by (metis homeomorphism_compose homeomorphism_symD hom1 hom2) - have "g1 u = a" - using \0 < r\ \f1 a = u\ assms hom1 homeomorphism_apply1 by fastforce - then show "(f2 \ g1) u = v" - by (simp add: \f2 a = v\) - show "\x. \x \ sphere a r; x \ T\ \ (f2 \ g1) x = x" - using f1 f2 hom1 homeomorphism_apply1 by fastforce - qed -qed - - -corollary%unimportant homeomorphism_moving_point_3: - fixes a :: "'a::euclidean_space" - assumes "affine T" "a \ T" and ST: "ball a r \ T \ S" "S \ T" - and u: "u \ ball a r \ T" and v: "v \ ball a r \ T" - obtains f g where "homeomorphism S S f g" - "f u = v" "{x. \ (f x = x \ g x = x)} \ ball a r \ T" -proof - - obtain f g where hom: "homeomorphism (cball a r \ T) (cball a r \ T) f g" - and "f u = v" and fid: "\x. \x \ sphere a r; x \ T\ \ f x = x" - using homeomorphism_moving_point_2 [OF \affine T\ \a \ T\ u v] by blast - have gid: "\x. \x \ sphere a r; x \ T\ \ g x = x" - using fid hom homeomorphism_apply1 by fastforce - define ff where "ff \ \x. if x \ ball a r \ T then f x else x" - define gg where "gg \ \x. if x \ ball a r \ T then g x else x" - show ?thesis - proof - show "homeomorphism S S ff gg" - proof (rule homeomorphismI) - have "continuous_on ((cball a r \ T) \ (T - ball a r)) ff" - apply (simp add: ff_def) - apply (rule continuous_on_cases) - using homeomorphism_cont1 [OF hom] - apply (auto simp: affine_closed \affine T\ continuous_on_id fid) - done - then show "continuous_on S ff" - apply (rule continuous_on_subset) - using ST by auto - have "continuous_on ((cball a r \ T) \ (T - ball a r)) gg" - apply (simp add: gg_def) - apply (rule continuous_on_cases) - using homeomorphism_cont2 [OF hom] - apply (auto simp: affine_closed \affine T\ continuous_on_id gid) - done - then show "continuous_on S gg" - apply (rule continuous_on_subset) - using ST by auto - show "ff ` S \ S" - proof (clarsimp simp add: ff_def) - fix x - assume "x \ S" and x: "dist a x < r" and "x \ T" - then have "f x \ cball a r \ T" - using homeomorphism_image1 [OF hom] by force - then show "f x \ S" - using ST(1) \x \ T\ gid hom homeomorphism_def x by fastforce - qed - show "gg ` S \ S" - proof (clarsimp simp add: gg_def) - fix x - assume "x \ S" and x: "dist a x < r" and "x \ T" - then have "g x \ cball a r \ T" - using homeomorphism_image2 [OF hom] by force - then have "g x \ ball a r" - using homeomorphism_apply2 [OF hom] - by (metis Diff_Diff_Int Diff_iff \x \ T\ cball_def fid le_less mem_Collect_eq mem_ball mem_sphere x) - then show "g x \ S" - using ST(1) \g x \ cball a r \ T\ by force - qed - show "\x. x \ S \ gg (ff x) = x" - apply (simp add: ff_def gg_def) - using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] - apply auto - apply (metis Int_iff homeomorphism_apply1 [OF hom] fid image_eqI less_eq_real_def mem_cball mem_sphere) - done - show "\x. x \ S \ ff (gg x) = x" - apply (simp add: ff_def gg_def) - using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] - apply auto - apply (metis Int_iff fid image_eqI less_eq_real_def mem_cball mem_sphere) - done - qed - show "ff u = v" - using u by (auto simp: ff_def \f u = v\) - show "{x. \ (ff x = x \ gg x = x)} \ ball a r \ T" - by (auto simp: ff_def gg_def) - qed -qed - - -proposition%unimportant homeomorphism_moving_point: - fixes a :: "'a::euclidean_space" - assumes ope: "openin (subtopology euclidean (affine hull S)) S" - and "S \ T" - and TS: "T \ affine hull S" - and S: "connected S" "a \ S" "b \ S" - obtains f g where "homeomorphism T T f g" "f a = b" - "{x. \ (f x = x \ g x = x)} \ S" - "bounded {x. \ (f x = x \ g x = x)}" -proof - - have 1: "\h k. homeomorphism T T h k \ h (f d) = d \ - {x. \ (h x = x \ k x = x)} \ S \ bounded {x. \ (h x = x \ k x = x)}" - if "d \ S" "f d \ S" and homfg: "homeomorphism T T f g" - and S: "{x. \ (f x = x \ g x = x)} \ S" - and bo: "bounded {x. \ (f x = x \ g x = x)}" for d f g - proof (intro exI conjI) - show homgf: "homeomorphism T T g f" - by (metis homeomorphism_symD homfg) - then show "g (f d) = d" - by (meson \S \ T\ homeomorphism_def subsetD \d \ S\) - show "{x. \ (g x = x \ f x = x)} \ S" - using S by blast - show "bounded {x. \ (g x = x \ f x = x)}" - using bo by (simp add: conj_commute) - qed - have 2: "\f g. homeomorphism T T f g \ f x = f2 (f1 x) \ - {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" - if "x \ S" "f1 x \ S" "f2 (f1 x) \ S" - and hom: "homeomorphism T T f1 g1" "homeomorphism T T f2 g2" - and sub: "{x. \ (f1 x = x \ g1 x = x)} \ S" "{x. \ (f2 x = x \ g2 x = x)} \ S" - and bo: "bounded {x. \ (f1 x = x \ g1 x = x)}" "bounded {x. \ (f2 x = x \ g2 x = x)}" - for x f1 f2 g1 g2 - proof (intro exI conjI) - show homgf: "homeomorphism T T (f2 \ f1) (g1 \ g2)" - by (metis homeomorphism_compose hom) - then show "(f2 \ f1) x = f2 (f1 x)" - by force - show "{x. \ ((f2 \ f1) x = x \ (g1 \ g2) x = x)} \ S" - using sub by force - have "bounded ({x. \(f1 x = x \ g1 x = x)} \ {x. \(f2 x = x \ g2 x = x)})" - using bo by simp - then show "bounded {x. \ ((f2 \ f1) x = x \ (g1 \ g2) x = x)}" - by (rule bounded_subset) auto - qed - have 3: "\U. openin (subtopology euclidean S) U \ - d \ U \ - (\x\U. - \f g. homeomorphism T T f g \ f d = x \ - {x. \ (f x = x \ g x = x)} \ S \ - bounded {x. \ (f x = x \ g x = x)})" - if "d \ S" for d - proof - - obtain r where "r > 0" and r: "ball d r \ affine hull S \ S" - by (metis \d \ S\ ope openin_contains_ball) - have *: "\f g. homeomorphism T T f g \ f d = e \ - {x. \ (f x = x \ g x = x)} \ S \ - bounded {x. \ (f x = x \ g x = x)}" if "e \ S" "e \ ball d r" for e - apply (rule homeomorphism_moving_point_3 [of "affine hull S" d r T d e]) - using r \S \ T\ TS that - apply (auto simp: \d \ S\ \0 < r\ hull_inc) - using bounded_subset by blast - show ?thesis - apply (rule_tac x="S \ ball d r" in exI) - apply (intro conjI) - apply (simp add: openin_open_Int) - apply (simp add: \0 < r\ that) - apply (blast intro: *) - done - qed - have "\f g. homeomorphism T T f g \ f a = b \ - {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" - apply (rule connected_equivalence_relation [OF S], safe) - apply (blast intro: 1 2 3)+ - done - then show ?thesis - using that by auto -qed - - -lemma homeomorphism_moving_points_exists_gen: - assumes K: "finite K" "\i. i \ K \ x i \ S \ y i \ S" - "pairwise (\i j. (x i \ x j) \ (y i \ y j)) K" - and "2 \ aff_dim S" - and ope: "openin (subtopology euclidean (affine hull S)) S" - and "S \ T" "T \ affine hull S" "connected S" - shows "\f g. homeomorphism T T f g \ (\i \ K. f(x i) = y i) \ - {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" - using assms -proof (induction K) - case empty - then show ?case - by (force simp: homeomorphism_ident) -next - case (insert i K) - then have xney: "\j. \j \ K; j \ i\ \ x i \ x j \ y i \ y j" - and pw: "pairwise (\i j. x i \ x j \ y i \ y j) K" - and "x i \ S" "y i \ S" - and xyS: "\i. i \ K \ x i \ S \ y i \ S" - by (simp_all add: pairwise_insert) - obtain f g where homfg: "homeomorphism T T f g" and feq: "\i. i \ K \ f(x i) = y i" - and fg_sub: "{x. \ (f x = x \ g x = x)} \ S" - and bo_fg: "bounded {x. \ (f x = x \ g x = x)}" - using insert.IH [OF xyS pw] insert.prems by (blast intro: that) - then have "\f g. homeomorphism T T f g \ (\i \ K. f(x i) = y i) \ - {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" - using insert by blast - have aff_eq: "affine hull (S - y ` K) = affine hull S" - apply (rule affine_hull_Diff) - apply (auto simp: insert) - using \y i \ S\ insert.hyps(2) xney xyS by fastforce - have f_in_S: "f x \ S" if "x \ S" for x - using homfg fg_sub homeomorphism_apply1 \S \ T\ - proof - - have "(f (f x) \ f x \ g (f x) \ f x) \ f x \ S" - by (metis \S \ T\ homfg subsetD homeomorphism_apply1 that) - then show ?thesis - using fg_sub by force - qed - obtain h k where homhk: "homeomorphism T T h k" and heq: "h (f (x i)) = y i" - and hk_sub: "{x. \ (h x = x \ k x = x)} \ S - y ` K" - and bo_hk: "bounded {x. \ (h x = x \ k x = x)}" - proof (rule homeomorphism_moving_point [of "S - y`K" T "f(x i)" "y i"]) - show "openin (subtopology euclidean (affine hull (S - y ` K))) (S - y ` K)" - by (simp add: aff_eq openin_diff finite_imp_closedin image_subset_iff hull_inc insert xyS) - show "S - y ` K \ T" - using \S \ T\ by auto - show "T \ affine hull (S - y ` K)" - using insert by (simp add: aff_eq) - show "connected (S - y ` K)" - proof (rule connected_openin_diff_countable [OF \connected S\ ope]) - show "\ collinear S" - using collinear_aff_dim \2 \ aff_dim S\ by force - show "countable (y ` K)" - using countable_finite insert.hyps(1) by blast - qed - show "f (x i) \ S - y ` K" - apply (auto simp: f_in_S \x i \ S\) - by (metis feq homfg \x i \ S\ homeomorphism_def \S \ T\ \i \ K\ subsetCE xney xyS) - show "y i \ S - y ` K" - using insert.hyps xney by (auto simp: \y i \ S\) - qed blast - show ?case - proof (intro exI conjI) - show "homeomorphism T T (h \ f) (g \ k)" - using homfg homhk homeomorphism_compose by blast - show "\i \ insert i K. (h \ f) (x i) = y i" - using feq hk_sub by (auto simp: heq) - show "{x. \ ((h \ f) x = x \ (g \ k) x = x)} \ S" - using fg_sub hk_sub by force - have "bounded ({x. \(f x = x \ g x = x)} \ {x. \(h x = x \ k x = x)})" - using bo_fg bo_hk bounded_Un by blast - then show "bounded {x. \ ((h \ f) x = x \ (g \ k) x = x)}" - by (rule bounded_subset) auto - qed -qed - -proposition%unimportant homeomorphism_moving_points_exists: - fixes S :: "'a::euclidean_space set" - assumes 2: "2 \ DIM('a)" "open S" "connected S" "S \ T" "finite K" - and KS: "\i. i \ K \ x i \ S \ y i \ S" - and pw: "pairwise (\i j. (x i \ x j) \ (y i \ y j)) K" - and S: "S \ T" "T \ affine hull S" "connected S" - obtains f g where "homeomorphism T T f g" "\i. i \ K \ f(x i) = y i" - "{x. \ (f x = x \ g x = x)} \ S" "bounded {x. (\ (f x = x \ g x = x))}" -proof (cases "S = {}") - case True - then show ?thesis - using KS homeomorphism_ident that by fastforce -next - case False - then have affS: "affine hull S = UNIV" - by (simp add: affine_hull_open \open S\) - then have ope: "openin (subtopology euclidean (affine hull S)) S" - using \open S\ open_openin by auto - have "2 \ DIM('a)" by (rule 2) - also have "\ = aff_dim (UNIV :: 'a set)" - by simp - also have "\ \ aff_dim S" - by (metis aff_dim_UNIV aff_dim_affine_hull aff_dim_le_DIM affS) - finally have "2 \ aff_dim S" - by linarith - then show ?thesis - using homeomorphism_moving_points_exists_gen [OF \finite K\ KS pw _ ope S] that by fastforce -qed - -subsubsection%unimportant\The theorem \homeomorphism_grouping_points_exists\\ - -lemma homeomorphism_grouping_point_1: - fixes a::real and c::real - assumes "a < b" "c < d" - obtains f g where "homeomorphism (cbox a b) (cbox c d) f g" "f a = c" "f b = d" -proof - - define f where "f \ \x. ((d - c) / (b - a)) * x + (c - a * ((d - c) / (b - a)))" - have "\g. homeomorphism (cbox a b) (cbox c d) f g" - proof (rule homeomorphism_compact) - show "continuous_on (cbox a b) f" - apply (simp add: f_def) - apply (intro continuous_intros) - using assms by auto - have "f ` {a..b} = {c..d}" - unfolding f_def image_affinity_atLeastAtMost - using assms sum_sqs_eq by (auto simp: divide_simps algebra_simps) - then show "f ` cbox a b = cbox c d" - by auto - show "inj_on f (cbox a b)" - unfolding f_def inj_on_def using assms by auto - qed auto - then obtain g where "homeomorphism (cbox a b) (cbox c d) f g" .. - then show ?thesis - proof - show "f a = c" - by (simp add: f_def) - show "f b = d" - using assms sum_sqs_eq [of a b] by (auto simp: f_def divide_simps algebra_simps) - qed -qed - -lemma homeomorphism_grouping_point_2: - fixes a::real and w::real - assumes hom_ab: "homeomorphism (cbox a b) (cbox u v) f1 g1" - and hom_bc: "homeomorphism (cbox b c) (cbox v w) f2 g2" - and "b \ cbox a c" "v \ cbox u w" - and eq: "f1 a = u" "f1 b = v" "f2 b = v" "f2 c = w" - obtains f g where "homeomorphism (cbox a c) (cbox u w) f g" "f a = u" "f c = w" - "\x. x \ cbox a b \ f x = f1 x" "\x. x \ cbox b c \ f x = f2 x" -proof - - have le: "a \ b" "b \ c" "u \ v" "v \ w" - using assms by simp_all - then have ac: "cbox a c = cbox a b \ cbox b c" and uw: "cbox u w = cbox u v \ cbox v w" - by auto - define f where "f \ \x. if x \ b then f1 x else f2 x" - have "\g. homeomorphism (cbox a c) (cbox u w) f g" - proof (rule homeomorphism_compact) - have cf1: "continuous_on (cbox a b) f1" - using hom_ab homeomorphism_cont1 by blast - have cf2: "continuous_on (cbox b c) f2" - using hom_bc homeomorphism_cont1 by blast - show "continuous_on (cbox a c) f" - apply (simp add: f_def) - apply (rule continuous_on_cases_le [OF continuous_on_subset [OF cf1] continuous_on_subset [OF cf2]]) - using le eq apply (force simp: continuous_on_id)+ - done - have "f ` cbox a b = f1 ` cbox a b" "f ` cbox b c = f2 ` cbox b c" - unfolding f_def using eq by force+ - then show "f ` cbox a c = cbox u w" - apply (simp only: ac uw image_Un) - by (metis hom_ab hom_bc homeomorphism_def) - have neq12: "f1 x \ f2 y" if x: "a \ x" "x \ b" and y: "b < y" "y \ c" for x y - proof - - have "f1 x \ cbox u v" - by (metis hom_ab homeomorphism_def image_eqI mem_box_real(2) x) - moreover have "f2 y \ cbox v w" - by (metis (full_types) hom_bc homeomorphism_def image_subset_iff mem_box_real(2) not_le not_less_iff_gr_or_eq order_refl y) - moreover have "f2 y \ f2 b" - by (metis cancel_comm_monoid_add_class.diff_cancel diff_gt_0_iff_gt hom_bc homeomorphism_def le(2) less_imp_le less_numeral_extra(3) mem_box_real(2) order_refl y) - ultimately show ?thesis - using le eq by simp - qed - have "inj_on f1 (cbox a b)" - by (metis (full_types) hom_ab homeomorphism_def inj_onI) - moreover have "inj_on f2 (cbox b c)" - by (metis (full_types) hom_bc homeomorphism_def inj_onI) - ultimately show "inj_on f (cbox a c)" - apply (simp (no_asm) add: inj_on_def) - apply (simp add: f_def inj_on_eq_iff) - using neq12 apply force - done - qed auto - then obtain g where "homeomorphism (cbox a c) (cbox u w) f g" .. - then show ?thesis - apply (rule that) - using eq le by (auto simp: f_def) -qed - -lemma homeomorphism_grouping_point_3: - fixes a::real - assumes cbox_sub: "cbox c d \ box a b" "cbox u v \ box a b" - and box_ne: "box c d \ {}" "box u v \ {}" - obtains f g where "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" - "\x. x \ cbox c d \ f x \ cbox u v" -proof - - have less: "a < c" "a < u" "d < b" "v < b" "c < d" "u < v" "cbox c d \ {}" - using assms - by (simp_all add: cbox_sub subset_eq) - obtain f1 g1 where 1: "homeomorphism (cbox a c) (cbox a u) f1 g1" - and f1_eq: "f1 a = a" "f1 c = u" - using homeomorphism_grouping_point_1 [OF \a < c\ \a < u\] . - obtain f2 g2 where 2: "homeomorphism (cbox c d) (cbox u v) f2 g2" - and f2_eq: "f2 c = u" "f2 d = v" - using homeomorphism_grouping_point_1 [OF \c < d\ \u < v\] . - obtain f3 g3 where 3: "homeomorphism (cbox d b) (cbox v b) f3 g3" - and f3_eq: "f3 d = v" "f3 b = b" - using homeomorphism_grouping_point_1 [OF \d < b\ \v < b\] . - obtain f4 g4 where 4: "homeomorphism (cbox a d) (cbox a v) f4 g4" and "f4 a = a" "f4 d = v" - and f4_eq: "\x. x \ cbox a c \ f4 x = f1 x" "\x. x \ cbox c d \ f4 x = f2 x" - using homeomorphism_grouping_point_2 [OF 1 2] less by (auto simp: f1_eq f2_eq) - obtain f g where fg: "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" - and f_eq: "\x. x \ cbox a d \ f x = f4 x" "\x. x \ cbox d b \ f x = f3 x" - using homeomorphism_grouping_point_2 [OF 4 3] less by (auto simp: f4_eq f3_eq f2_eq f1_eq) - show ?thesis - apply (rule that [OF fg]) - using f4_eq f_eq homeomorphism_image1 [OF 2] - apply simp - by (metis atLeastAtMost_iff box_real(1) box_real(2) cbox_sub(1) greaterThanLessThan_iff imageI less_eq_real_def subset_eq) -qed - - -lemma homeomorphism_grouping_point_4: - fixes T :: "real set" - assumes "open U" "open S" "connected S" "U \ {}" "finite K" "K \ S" "U \ S" "S \ T" - obtains f g where "homeomorphism T T f g" - "\x. x \ K \ f x \ U" "{x. (\ (f x = x \ g x = x))} \ S" - "bounded {x. (\ (f x = x \ g x = x))}" -proof - - obtain c d where "box c d \ {}" "cbox c d \ U" - proof - - obtain u where "u \ U" - using \U \ {}\ by blast - then obtain e where "e > 0" "cball u e \ U" - using \open U\ open_contains_cball by blast - then show ?thesis - by (rule_tac c=u and d="u+e" in that) (auto simp: dist_norm subset_iff) - qed - have "compact K" - by (simp add: \finite K\ finite_imp_compact) - obtain a b where "box a b \ {}" "K \ cbox a b" "cbox a b \ S" - proof (cases "K = {}") - case True then show ?thesis - using \box c d \ {}\ \cbox c d \ U\ \U \ S\ that by blast - next - case False - then obtain a b where "a \ K" "b \ K" - and a: "\x. x \ K \ a \ x" and b: "\x. x \ K \ x \ b" - using compact_attains_inf compact_attains_sup by (metis \compact K\)+ - obtain e where "e > 0" "cball b e \ S" - using \open S\ open_contains_cball - by (metis \b \ K\ \K \ S\ subsetD) - show ?thesis - proof - show "box a (b + e) \ {}" - using \0 < e\ \b \ K\ a by force - show "K \ cbox a (b + e)" - using \0 < e\ a b by fastforce - have "a \ S" - using \a \ K\ assms(6) by blast - have "b + e \ S" - using \0 < e\ \cball b e \ S\ by (force simp: dist_norm) - show "cbox a (b + e) \ S" - using \a \ S\ \b + e \ S\ \connected S\ connected_contains_Icc by auto - qed - qed - obtain w z where "cbox w z \ S" and sub_wz: "cbox a b \ cbox c d \ box w z" - proof - - have "a \ S" "b \ S" - using \box a b \ {}\ \cbox a b \ S\ by auto - moreover have "c \ S" "d \ S" - using \box c d \ {}\ \cbox c d \ U\ \U \ S\ by force+ - ultimately have "min a c \ S" "max b d \ S" - by linarith+ - then obtain e1 e2 where "e1 > 0" "cball (min a c) e1 \ S" "e2 > 0" "cball (max b d) e2 \ S" - using \open S\ open_contains_cball by metis - then have *: "min a c - e1 \ S" "max b d + e2 \ S" - by (auto simp: dist_norm) - show ?thesis - proof - show "cbox (min a c - e1) (max b d+ e2) \ S" - using * \connected S\ connected_contains_Icc by auto - show "cbox a b \ cbox c d \ box (min a c - e1) (max b d + e2)" - using \0 < e1\ \0 < e2\ by auto - qed - qed - then - obtain f g where hom: "homeomorphism (cbox w z) (cbox w z) f g" - and "f w = w" "f z = z" - and fin: "\x. x \ cbox a b \ f x \ cbox c d" - using homeomorphism_grouping_point_3 [of a b w z c d] - using \box a b \ {}\ \box c d \ {}\ by blast - have contfg: "continuous_on (cbox w z) f" "continuous_on (cbox w z) g" - using hom homeomorphism_def by blast+ - define f' where "f' \ \x. if x \ cbox w z then f x else x" - define g' where "g' \ \x. if x \ cbox w z then g x else x" - show ?thesis - proof - have T: "cbox w z \ (T - box w z) = T" - using \cbox w z \ S\ \S \ T\ by auto - show "homeomorphism T T f' g'" - proof - have clo: "closedin (subtopology euclidean (cbox w z \ (T - box w z))) (T - box w z)" - by (metis Diff_Diff_Int Diff_subset T closedin_def open_box openin_open_Int topspace_euclidean_subtopology) - have "continuous_on (cbox w z \ (T - box w z)) f'" "continuous_on (cbox w z \ (T - box w z)) g'" - unfolding f'_def g'_def - apply (safe intro!: continuous_on_cases_local contfg continuous_on_id clo) - apply (simp_all add: closed_subset) - using \f w = w\ \f z = z\ apply force - by (metis \f w = w\ \f z = z\ hom homeomorphism_def less_eq_real_def mem_box_real(2)) - then show "continuous_on T f'" "continuous_on T g'" - by (simp_all only: T) - show "f' ` T \ T" - unfolding f'_def - by clarsimp (metis \cbox w z \ S\ \S \ T\ subsetD hom homeomorphism_def imageI mem_box_real(2)) - show "g' ` T \ T" - unfolding g'_def - by clarsimp (metis \cbox w z \ S\ \S \ T\ subsetD hom homeomorphism_def imageI mem_box_real(2)) - show "\x. x \ T \ g' (f' x) = x" - unfolding f'_def g'_def - using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] by fastforce - show "\y. y \ T \ f' (g' y) = y" - unfolding f'_def g'_def - using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] by fastforce - qed - show "\x. x \ K \ f' x \ U" - using fin sub_wz \K \ cbox a b\ \cbox c d \ U\ by (force simp: f'_def) - show "{x. \ (f' x = x \ g' x = x)} \ S" - using \cbox w z \ S\ by (auto simp: f'_def g'_def) - show "bounded {x. \ (f' x = x \ g' x = x)}" - apply (rule bounded_subset [of "cbox w z"]) - using bounded_cbox apply blast - apply (auto simp: f'_def g'_def) - done - qed -qed - -proposition%unimportant homeomorphism_grouping_points_exists: - fixes S :: "'a::euclidean_space set" - assumes "open U" "open S" "connected S" "U \ {}" "finite K" "K \ S" "U \ S" "S \ T" - obtains f g where "homeomorphism T T f g" "{x. (\ (f x = x \ g x = x))} \ S" - "bounded {x. (\ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" -proof (cases "2 \ DIM('a)") - case True - have TS: "T \ affine hull S" - using affine_hull_open assms by blast - have "infinite U" - using \open U\ \U \ {}\ finite_imp_not_open by blast - then obtain P where "P \ U" "finite P" "card K = card P" - using infinite_arbitrarily_large by metis - then obtain \ where \: "bij_betw \ K P" - using \finite K\ finite_same_card_bij by blast - obtain f g where "homeomorphism T T f g" "\i. i \ K \ f (id i) = \ i" "{x. \ (f x = x \ g x = x)} \ S" "bounded {x. \ (f x = x \ g x = x)}" - proof (rule homeomorphism_moving_points_exists [OF True \open S\ \connected S\ \S \ T\ \finite K\]) - show "\i. i \ K \ id i \ S \ \ i \ S" - using \P \ U\ \bij_betw \ K P\ \K \ S\ \U \ S\ bij_betwE by blast - show "pairwise (\i j. id i \ id j \ \ i \ \ j) K" - using \ by (auto simp: pairwise_def bij_betw_def inj_on_def) - qed (use affine_hull_open assms that in auto) - then show ?thesis - using \ \P \ U\ bij_betwE by (fastforce simp add: intro!: that) -next - case False - with DIM_positive have "DIM('a) = 1" - by (simp add: dual_order.antisym) - then obtain h::"'a \real" and j - where "linear h" "linear j" - and noh: "\x. norm(h x) = norm x" and noj: "\y. norm(j y) = norm y" - and hj: "\x. j(h x) = x" "\y. h(j y) = y" - and ranh: "surj h" - using isomorphisms_UNIV_UNIV - by (metis (mono_tags, hide_lams) DIM_real UNIV_eq_I range_eqI) - obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" - and f: "\x. x \ h ` K \ f x \ h ` U" - and sub: "{x. \ (f x = x \ g x = x)} \ h ` S" - and bou: "bounded {x. \ (f x = x \ g x = x)}" - apply (rule homeomorphism_grouping_point_4 [of "h ` U" "h ` S" "h ` K" "h ` T"]) - by (simp_all add: assms image_mono \linear h\ open_surjective_linear_image connected_linear_image ranh) - have jf: "j (f (h x)) = x \ f (h x) = h x" for x - by (metis hj) - have jg: "j (g (h x)) = x \ g (h x) = h x" for x - by (metis hj) - have cont_hj: "continuous_on X h" "continuous_on Y j" for X Y - by (simp_all add: \linear h\ \linear j\ linear_linear linear_continuous_on) - show ?thesis - proof - show "homeomorphism T T (j \ f \ h) (j \ g \ h)" - proof - show "continuous_on T (j \ f \ h)" "continuous_on T (j \ g \ h)" - using hom homeomorphism_def - by (blast intro: continuous_on_compose cont_hj)+ - show "(j \ f \ h) ` T \ T" "(j \ g \ h) ` T \ T" - by auto (metis (mono_tags, hide_lams) hj(1) hom homeomorphism_def imageE imageI)+ - show "\x. x \ T \ (j \ g \ h) ((j \ f \ h) x) = x" - using hj hom homeomorphism_apply1 by fastforce - show "\y. y \ T \ (j \ f \ h) ((j \ g \ h) y) = y" - using hj hom homeomorphism_apply2 by fastforce - qed - show "{x. \ ((j \ f \ h) x = x \ (j \ g \ h) x = x)} \ S" - apply (clarsimp simp: jf jg hj) - using sub hj - apply (drule_tac c="h x" in subsetD, force) - by (metis imageE) - have "bounded (j ` {x. (\ (f x = x \ g x = x))})" - by (rule bounded_linear_image [OF bou]) (use \linear j\ linear_conv_bounded_linear in auto) - moreover - have *: "{x. \((j \ f \ h) x = x \ (j \ g \ h) x = x)} = j ` {x. (\ (f x = x \ g x = x))}" - using hj by (auto simp: jf jg image_iff, metis+) - ultimately show "bounded {x. \ ((j \ f \ h) x = x \ (j \ g \ h) x = x)}" - by metis - show "\x. x \ K \ (j \ f \ h) x \ U" - using f hj by fastforce - qed -qed - - -proposition%unimportant homeomorphism_grouping_points_exists_gen: - fixes S :: "'a::euclidean_space set" - assumes opeU: "openin (subtopology euclidean S) U" - and opeS: "openin (subtopology euclidean (affine hull S)) S" - and "U \ {}" "finite K" "K \ S" and S: "S \ T" "T \ affine hull S" "connected S" - obtains f g where "homeomorphism T T f g" "{x. (\ (f x = x \ g x = x))} \ S" - "bounded {x. (\ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" -proof (cases "2 \ aff_dim S") - case True - have opeU': "openin (subtopology euclidean (affine hull S)) U" - using opeS opeU openin_trans by blast - obtain u where "u \ U" "u \ S" - using \U \ {}\ opeU openin_imp_subset by fastforce+ - have "infinite U" - apply (rule infinite_openin [OF opeU \u \ U\]) - apply (rule connected_imp_perfect_aff_dim [OF \connected S\ _ \u \ S\]) - using True apply simp - done - then obtain P where "P \ U" "finite P" "card K = card P" - using infinite_arbitrarily_large by metis - then obtain \ where \: "bij_betw \ K P" - using \finite K\ finite_same_card_bij by blast - have "\f g. homeomorphism T T f g \ (\i \ K. f(id i) = \ i) \ - {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" - proof (rule homeomorphism_moving_points_exists_gen [OF \finite K\ _ _ True opeS S]) - show "\i. i \ K \ id i \ S \ \ i \ S" - by (metis id_apply opeU openin_contains_cball subsetCE \P \ U\ \bij_betw \ K P\ \K \ S\ bij_betwE) - show "pairwise (\i j. id i \ id j \ \ i \ \ j) K" - using \ by (auto simp: pairwise_def bij_betw_def inj_on_def) - qed - then show ?thesis - using \ \P \ U\ bij_betwE by (fastforce simp add: intro!: that) -next - case False - with aff_dim_geq [of S] consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S = 1" by linarith - then show ?thesis - proof cases - assume "aff_dim S = -1" - then have "S = {}" - using aff_dim_empty by blast - then have "False" - using \U \ {}\ \K \ S\ openin_imp_subset [OF opeU] by blast - then show ?thesis .. - next - assume "aff_dim S = 0" - then obtain a where "S = {a}" - using aff_dim_eq_0 by blast - then have "K \ U" - using \U \ {}\ \K \ S\ openin_imp_subset [OF opeU] by blast - show ?thesis - apply (rule that [of id id]) - using \K \ U\ by (auto simp: continuous_on_id intro: homeomorphismI) - next - assume "aff_dim S = 1" - then have "affine hull S homeomorphic (UNIV :: real set)" - by (auto simp: homeomorphic_affine_sets) - then obtain h::"'a\real" and j where homhj: "homeomorphism (affine hull S) UNIV h j" - using homeomorphic_def by blast - then have h: "\x. x \ affine hull S \ j(h(x)) = x" and j: "\y. j y \ affine hull S \ h(j y) = y" - by (auto simp: homeomorphism_def) - have connh: "connected (h ` S)" - by (meson Topological_Spaces.connected_continuous_image \connected S\ homeomorphism_cont1 homeomorphism_of_subsets homhj hull_subset top_greatest) - have hUS: "h ` U \ h ` S" - by (meson homeomorphism_imp_open_map homeomorphism_of_subsets homhj hull_subset opeS opeU open_UNIV openin_open_eq) - have opn: "openin (subtopology euclidean (affine hull S)) U \ open (h ` U)" for U - using homeomorphism_imp_open_map [OF homhj] by simp - have "open (h ` U)" "open (h ` S)" - by (auto intro: opeS opeU openin_trans opn) - then obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" - and f: "\x. x \ h ` K \ f x \ h ` U" - and sub: "{x. \ (f x = x \ g x = x)} \ h ` S" - and bou: "bounded {x. \ (f x = x \ g x = x)}" - apply (rule homeomorphism_grouping_points_exists [of "h ` U" "h ` S" "h ` K" "h ` T"]) - using assms by (auto simp: connh hUS) - have jf: "\x. x \ affine hull S \ j (f (h x)) = x \ f (h x) = h x" - by (metis h j) - have jg: "\x. x \ affine hull S \ j (g (h x)) = x \ g (h x) = h x" - by (metis h j) - have cont_hj: "continuous_on T h" "continuous_on Y j" for Y - apply (rule continuous_on_subset [OF _ \T \ affine hull S\]) - using homeomorphism_def homhj apply blast - by (meson continuous_on_subset homeomorphism_def homhj top_greatest) - define f' where "f' \ \x. if x \ affine hull S then (j \ f \ h) x else x" - define g' where "g' \ \x. if x \ affine hull S then (j \ g \ h) x else x" - show ?thesis - proof - show "homeomorphism T T f' g'" - proof - have "continuous_on T (j \ f \ h)" - apply (intro continuous_on_compose cont_hj) - using hom homeomorphism_def by blast - then show "continuous_on T f'" - apply (rule continuous_on_eq) - using \T \ affine hull S\ f'_def by auto - have "continuous_on T (j \ g \ h)" - apply (intro continuous_on_compose cont_hj) - using hom homeomorphism_def by blast - then show "continuous_on T g'" - apply (rule continuous_on_eq) - using \T \ affine hull S\ g'_def by auto - show "f' ` T \ T" - proof (clarsimp simp: f'_def) - fix x assume "x \ T" - then have "f (h x) \ h ` T" - by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) - then show "j (f (h x)) \ T" - using \T \ affine hull S\ h by auto - qed - show "g' ` T \ T" - proof (clarsimp simp: g'_def) - fix x assume "x \ T" - then have "g (h x) \ h ` T" - by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) - then show "j (g (h x)) \ T" - using \T \ affine hull S\ h by auto - qed - show "\x. x \ T \ g' (f' x) = x" - using h j hom homeomorphism_apply1 by (fastforce simp add: f'_def g'_def) - show "\y. y \ T \ f' (g' y) = y" - using h j hom homeomorphism_apply2 by (fastforce simp add: f'_def g'_def) - qed - next - show "{x. \ (f' x = x \ g' x = x)} \ S" - apply (clarsimp simp: f'_def g'_def jf jg) - apply (rule imageE [OF subsetD [OF sub]], force) - by (metis h hull_inc) - next - have "compact (j ` closure {x. \ (f x = x \ g x = x)})" - using bou by (auto simp: compact_continuous_image cont_hj) - then have "bounded (j ` {x. \ (f x = x \ g x = x)})" - by (rule bounded_closure_image [OF compact_imp_bounded]) - moreover - have *: "{x \ affine hull S. j (f (h x)) \ x \ j (g (h x)) \ x} = j ` {x. (\ (f x = x \ g x = x))}" - using h j by (auto simp: image_iff; metis) - ultimately have "bounded {x \ affine hull S. j (f (h x)) \ x \ j (g (h x)) \ x}" - by metis - then show "bounded {x. \ (f' x = x \ g' x = x)}" - by (simp add: f'_def g'_def Collect_mono bounded_subset) - next - show "f' x \ U" if "x \ K" for x - proof - - have "U \ S" - using opeU openin_imp_subset by blast - then have "j (f (h x)) \ U" - using f h hull_subset that by fastforce - then show "f' x \ U" - using \K \ S\ S f'_def that by auto - qed - qed - qed -qed - - -subsection\Nullhomotopic mappings\ - -text\ A mapping out of a sphere is nullhomotopic iff it extends to the ball. -This even works out in the degenerate cases when the radius is \\\ 0, and -we also don't need to explicitly assume continuity since it's already implicit -in both sides of the equivalence.\ - -lemma nullhomotopic_from_lemma: - assumes contg: "continuous_on (cball a r - {a}) g" - and fa: "\e. 0 < e - \ \d. 0 < d \ (\x. x \ a \ norm(x - a) < d \ norm(g x - f a) < e)" - and r: "\x. x \ cball a r \ x \ a \ f x = g x" - shows "continuous_on (cball a r) f" -proof (clarsimp simp: continuous_on_eq_continuous_within Ball_def) - fix x - assume x: "dist a x \ r" - show "continuous (at x within cball a r) f" - proof (cases "x=a") - case True - then show ?thesis - by (metis continuous_within_eps_delta fa dist_norm dist_self r) - next - case False - show ?thesis - proof (rule continuous_transform_within [where f=g and d = "norm(x-a)"]) - have "\d>0. \x'\cball a r. - dist x' x < d \ dist (g x') (g x) < e" if "e>0" for e - proof - - obtain d where "d > 0" - and d: "\x'. \dist x' a \ r; x' \ a; dist x' x < d\ \ - dist (g x') (g x) < e" - using contg False x \e>0\ - unfolding continuous_on_iff by (fastforce simp add: dist_commute intro: that) - show ?thesis - using \d > 0\ \x \ a\ - by (rule_tac x="min d (norm(x - a))" in exI) - (auto simp: dist_commute dist_norm [symmetric] intro!: d) - qed - then show "continuous (at x within cball a r) g" - using contg False by (auto simp: continuous_within_eps_delta) - show "0 < norm (x - a)" - using False by force - show "x \ cball a r" - by (simp add: x) - show "\x'. \x' \ cball a r; dist x' x < norm (x - a)\ - \ g x' = f x'" - by (metis dist_commute dist_norm less_le r) - qed - qed -qed - -proposition nullhomotopic_from_sphere_extension: - fixes f :: "'M::euclidean_space \ 'a::real_normed_vector" - shows "(\c. homotopic_with (\x. True) (sphere a r) S f (\x. c)) \ - (\g. continuous_on (cball a r) g \ g ` (cball a r) \ S \ - (\x \ sphere a r. g x = f x))" - (is "?lhs = ?rhs") -proof (cases r "0::real" rule: linorder_cases) - case equal - then show ?thesis - apply (auto simp: homotopic_with) - apply (rule_tac x="\x. h (0, a)" in exI) - apply (fastforce simp add:) - using continuous_on_const by blast -next - case greater - let ?P = "continuous_on {x. norm(x - a) = r} f \ f ` {x. norm(x - a) = r} \ S" - have ?P if ?lhs using that - proof - fix c - assume c: "homotopic_with (\x. True) (sphere a r) S f (\x. c)" - then have contf: "continuous_on (sphere a r) f" and fim: "f ` sphere a r \ S" - by (auto simp: homotopic_with_imp_subset1 homotopic_with_imp_continuous) - show ?P - using contf fim by (auto simp: sphere_def dist_norm norm_minus_commute) - qed - moreover have ?P if ?rhs using that - proof - fix g - assume g: "continuous_on (cball a r) g \ g ` cball a r \ S \ (\xa\sphere a r. g xa = f xa)" - then - show ?P - apply (safe elim!: continuous_on_eq [OF continuous_on_subset]) - apply (auto simp: dist_norm norm_minus_commute) - by (metis dist_norm image_subset_iff mem_sphere norm_minus_commute sphere_cball subsetCE) - qed - moreover have ?thesis if ?P - proof - assume ?lhs - then obtain c where "homotopic_with (\x. True) (sphere a r) S (\x. c) f" - using homotopic_with_sym by blast - then obtain h where conth: "continuous_on ({0..1::real} \ sphere a r) h" - and him: "h ` ({0..1} \ sphere a r) \ S" - and h: "\x. h(0, x) = c" "\x. h(1, x) = f x" - by (auto simp: homotopic_with_def) - obtain b1::'M where "b1 \ Basis" - using SOME_Basis by auto - have "c \ S" - apply (rule him [THEN subsetD]) - apply (rule_tac x = "(0, a + r *\<^sub>R b1)" in image_eqI) - using h greater \b1 \ Basis\ - apply (auto simp: dist_norm) - done - have uconth: "uniformly_continuous_on ({0..1::real} \ (sphere a r)) h" - by (force intro: compact_Times conth compact_uniformly_continuous) - let ?g = "\x. h (norm (x - a)/r, - a + (if x = a then r *\<^sub>R b1 else (r / norm(x - a)) *\<^sub>R (x - a)))" - let ?g' = "\x. h (norm (x - a)/r, a + (r / norm(x - a)) *\<^sub>R (x - a))" - show ?rhs - proof (intro exI conjI) - have "continuous_on (cball a r - {a}) ?g'" - apply (rule continuous_on_compose2 [OF conth]) - apply (intro continuous_intros) - using greater apply (auto simp: dist_norm norm_minus_commute) - done - then show "continuous_on (cball a r) ?g" - proof (rule nullhomotopic_from_lemma) - show "\d>0. \x. x \ a \ norm (x - a) < d \ norm (?g' x - ?g a) < e" if "0 < e" for e - proof - - obtain d where "0 < d" - and d: "\x x'. \x \ {0..1} \ sphere a r; x' \ {0..1} \ sphere a r; dist x' x < d\ - \ dist (h x') (h x) < e" - using uniformly_continuous_onE [OF uconth \0 < e\] by auto - have *: "norm (h (norm (x - a) / r, - a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + r *\<^sub>R b1)) < e" - if "x \ a" "norm (x - a) < r" "norm (x - a) < d * r" for x - proof - - have "norm (h (norm (x - a) / r, a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + r *\<^sub>R b1)) = - norm (h (norm (x - a) / r, a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + (r / norm (x - a)) *\<^sub>R (x - a)))" - by (simp add: h) - also have "\ < e" - apply (rule d [unfolded dist_norm]) - using greater \0 < d\ \b1 \ Basis\ that - by (auto simp: dist_norm divide_simps) - finally show ?thesis . - qed - show ?thesis - apply (rule_tac x = "min r (d * r)" in exI) - using greater \0 < d\ by (auto simp: *) - qed - show "\x. x \ cball a r \ x \ a \ ?g x = ?g' x" - by auto - qed - next - show "?g ` cball a r \ S" - using greater him \c \ S\ - by (force simp: h dist_norm norm_minus_commute) - next - show "\x\sphere a r. ?g x = f x" - using greater by (auto simp: h dist_norm norm_minus_commute) - qed - next - assume ?rhs - then obtain g where contg: "continuous_on (cball a r) g" - and gim: "g ` cball a r \ S" - and gf: "\x \ sphere a r. g x = f x" - by auto - let ?h = "\y. g (a + (fst y) *\<^sub>R (snd y - a))" - have "continuous_on ({0..1} \ sphere a r) ?h" - apply (rule continuous_on_compose2 [OF contg]) - apply (intro continuous_intros) - apply (auto simp: dist_norm norm_minus_commute mult_left_le_one_le) - done - moreover - have "?h ` ({0..1} \ sphere a r) \ S" - by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gim [THEN subsetD]) - moreover - have "\x\sphere a r. ?h (0, x) = g a" "\x\sphere a r. ?h (1, x) = f x" - by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gf) - ultimately - show ?lhs - apply (subst homotopic_with_sym) - apply (rule_tac x="g a" in exI) - apply (auto simp: homotopic_with) - done - qed - ultimately - show ?thesis by meson -qed simp - -end