diff --git a/src/HOL/Analysis/Homotopy.thy b/src/HOL/Analysis/Homotopy.thy --- a/src/HOL/Analysis/Homotopy.thy +++ b/src/HOL/Analysis/Homotopy.thy @@ -1,5677 +1,5709 @@ (* 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 Product_Topology begin definition\<^marker>\tag important\ homotopic_with where "homotopic_with P X Y f g \ (\h. continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h \ (\x. h(0, x) = f x) \ (\x. h(1, x) = g 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\.\ abbreviation homotopic_with_canon :: "[('a::topological_space \ 'b::topological_space) \ bool, 'a set, 'b set, 'a \ 'b, 'a \ 'b] \ bool" where "homotopic_with_canon P S T p q \ homotopic_with P (top_of_set S) (top_of_set T) p q" 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 lemma image_Pair_const: "(\x. (x, c)) ` A = A \ {c}" by auto lemma fst_o_paired [simp]: "fst \ (\(x,y). (f x y, g x y)) = (\(x,y). f x y)" by auto lemma snd_o_paired [simp]: "snd \ (\(x,y). (f x y, g x y)) = (\(x,y). g x y)" by auto 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 continuous_map_o_Pair: assumes h: "continuous_map (prod_topology X Y) Z h" and t: "t \ topspace X" shows "continuous_map Y Z (h \ Pair t)" apply (intro continuous_map_compose [OF _ h] continuous_map_id [unfolded id_def] continuous_intros) apply (simp add: t) done subsection\<^marker>\tag unimportant\\Trivial properties\ text \We often want to just localize the ending function equality or whatever.\ text\<^marker>\tag important\ \%whitespace\ proposition homotopic_with: assumes "\h k. (\x. x \ topspace X \ h x = k x) \ (P h \ P k)" shows "homotopic_with P X Y p q \ (\h. continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h \ (\x \ topspace X. h(0,x) = p x) \ (\x \ topspace 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 \ topspace X then h(u,v) else if u = 0 then p v else q v" in exI) apply auto using continuous_map_eq apply fastforce apply (drule_tac x=t in bspec, force) apply (subst assms; simp) done lemma homotopic_with_mono: assumes hom: "homotopic_with P X Y f g" and Q: "\h. \continuous_map X Y h; 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: o_def dest: continuous_map_o_Pair intro: Q) done lemma homotopic_with_imp_continuous_maps: assumes "homotopic_with P X Y f g" shows "continuous_map X Y f \ continuous_map X Y g" proof - obtain h where conth: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y 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_map X Y (h \ (\x. (t,x)))" for t by (rule continuous_map_compose [OF _ conth]) (simp add: o_def continuous_map_pairwise) show ?thesis using h *[of 0] *[of 1] by (simp add: continuous_map_eq) qed lemma homotopic_with_imp_continuous: assumes "homotopic_with_canon P X Y f g" shows "continuous_on X f \ continuous_on X g" by (meson assms continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) lemma homotopic_with_imp_property: assumes "homotopic_with P X Y f g" shows "P f \ P g" proof obtain h where h: "\x. h(0, x) = f x" "\x. h(1, x) = g x" and P: "\t. t \ {0..1::real} \ P(\x. h(t,x))" using assms by (force simp: homotopic_with_def) show "P f" "P g" using P [of 0] P [of 1] by (force simp: h)+ qed lemma homotopic_with_equal: assumes "P f" "P g" and contf: "continuous_map X Y f" and fg: "\x. x \ topspace X \ f x = g x" shows "homotopic_with P X Y f g" unfolding homotopic_with_def proof (intro exI conjI allI ballI) let ?h = "\(t::real,x). if t = 1 then g x else f x" show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y ?h" proof (rule continuous_map_eq) show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y (f \ snd)" by (simp add: contf continuous_map_of_snd) qed (auto simp: fg) show "P (\x. ?h (t, x))" if "t \ {0..1}" for t by (cases "t = 1") (simp_all add: assms) qed auto lemma homotopic_with_imp_subset1: "homotopic_with_canon 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) lemma homotopic_with_imp_subset2: "homotopic_with_canon 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) lemma homotopic_with_subset_left: "\homotopic_with_canon P X Y f g; Z \ X\ \ homotopic_with_canon P Z Y f g" apply (simp add: homotopic_with_def) apply (fast elim!: continuous_on_subset ex_forward) done lemma homotopic_with_subset_right: "\homotopic_with_canon P X Y f g; Y \ Z\ \ homotopic_with_canon P X Z f g" apply (simp add: homotopic_with_def) apply (fast elim!: continuous_on_subset ex_forward) done subsection\Homotopy with P is an equivalence relation\ text \(on continuous functions mapping X into Y that satisfy P, though this only affects reflexivity)\ lemma homotopic_with_refl [simp]: "homotopic_with P X Y f f \ continuous_map X Y f \ P f" by (auto simp: homotopic_with_imp_continuous_maps intro: homotopic_with_equal dest: homotopic_with_imp_property) lemma homotopic_with_symD: assumes "homotopic_with P X Y f g" shows "homotopic_with P X Y g f" proof - let ?I01 = "subtopology euclideanreal {0..1}" let ?j = "\y. (1 - fst y, snd y)" have 1: "continuous_map (prod_topology ?I01 X) (prod_topology euclideanreal X) ?j" apply (intro continuous_intros) apply (simp_all add: prod_topology_subtopology continuous_map_from_subtopology [OF continuous_map_fst]) done have *: "continuous_map (prod_topology ?I01 X) (prod_topology ?I01 X) ?j" proof - have "continuous_map (prod_topology ?I01 X) (subtopology (prod_topology euclideanreal X) ({0..1} \ topspace X)) ?j" by (simp add: continuous_map_into_subtopology [OF 1] image_subset_iff) then show ?thesis by (simp add: prod_topology_subtopology(1)) qed show ?thesis 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 (simp add: continuous_map_compose [OF *]) done qed lemma homotopic_with_sym: "homotopic_with P X Y f g \ homotopic_with P X Y g f" by (metis homotopic_with_symD) proposition homotopic_with_trans: assumes "homotopic_with P X Y f g" "homotopic_with P X Y g h" shows "homotopic_with P X Y f h" proof - let ?X01 = "prod_topology (subtopology euclideanreal {0..1}) X" obtain k1 k2 where contk1: "continuous_map ?X01 Y k1" and contk2: "continuous_map ?X01 Y k2" and k12: "\x. k1 (1, x) = g x" "\x. k2 (0, x) = g x" "\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))" using assms by (auto simp: homotopic_with_def) 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" have keq: "k1 (2 * u, v) = k2 (2 * u -1, v)" if "u = 1/2" for u v by (simp add: k12 that) show ?thesis unfolding homotopic_with_def proof (intro exI conjI) show "continuous_map ?X01 Y k" unfolding k_def proof (rule continuous_map_cases_le) show fst: "continuous_map ?X01 euclideanreal fst" using continuous_map_fst continuous_map_in_subtopology by blast show "continuous_map ?X01 euclideanreal (\x. 1/2)" by simp show "continuous_map (subtopology ?X01 {y \ topspace ?X01. fst y \ 1/2}) Y (k1 \ (\x. (2 *\<^sub>R fst x, snd x)))" apply (rule fst continuous_map_compose [OF _ contk1] continuous_intros continuous_map_into_subtopology | simp)+ apply (intro continuous_intros fst continuous_map_from_subtopology) apply (force simp: prod_topology_subtopology) using continuous_map_snd continuous_map_from_subtopology by blast show "continuous_map (subtopology ?X01 {y \ topspace ?X01. 1/2 \ fst y}) Y (k2 \ (\x. (2 *\<^sub>R fst x -1, snd x)))" apply (rule fst continuous_map_compose [OF _ contk2] continuous_intros continuous_map_into_subtopology | simp)+ apply (rule continuous_intros fst continuous_map_from_subtopology | simp)+ apply (force simp: prod_topology_subtopology) using continuous_map_snd continuous_map_from_subtopology by blast show "(k1 \ (\x. (2 *\<^sub>R fst x, snd x))) y = (k2 \ (\x. (2 *\<^sub>R fst x -1, snd x))) y" if "y \ topspace ?X01" and "fst y = 1/2" for y using that by (simp add: keq) qed show "\x. k (0, x) = f x" by (simp add: k12 k_def) show "\x. k (1, x) = h x" by (simp add: k12 k_def) show "\t\{0..1}. P (\x. k (t, x))" using P apply (clarsimp simp add: k_def) apply (case_tac "t \ 1/2", auto) done qed qed lemma homotopic_with_id2: "(\x. x \ topspace X \ g (f x) = x) \ homotopic_with (\x. True) X X (g \ f) id" by (metis comp_apply continuous_map_id eq_id_iff homotopic_with_equal homotopic_with_symD) subsection\Continuity lemmas\ lemma homotopic_with_compose_continuous_map_left: "\homotopic_with p X1 X2 f g; continuous_map X2 X3 h; \j. p j \ q(h \ j)\ \ homotopic_with q X1 X3 (h \ f) (h \ g)" unfolding homotopic_with_def apply clarify apply (rename_tac k) apply (rule_tac x="h \ k" in exI) apply (rule conjI continuous_map_compose | simp add: o_def)+ done lemma homotopic_compose_continuous_map_left: "\homotopic_with (\k. True) X1 X2 f g; continuous_map X2 X3 h\ \ homotopic_with (\k. True) X1 X3 (h \ f) (h \ g)" by (simp add: homotopic_with_compose_continuous_map_left) lemma homotopic_with_compose_continuous_map_right: assumes hom: "homotopic_with p X2 X3 f g" and conth: "continuous_map X1 X2 h" and q: "\j. p j \ q(j \ h)" shows "homotopic_with q X1 X3 (f \ h) (g \ h)" proof - obtain k where contk: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X2) X3 k" and k: "\x. k (0, x) = f x" "\x. k (1, x) = g x" and p: "\t. t\{0..1} \ p (\x. k (t, x))" using hom unfolding homotopic_with_def by blast have hsnd: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) X2 (h \ snd)" by (rule continuous_map_compose [OF continuous_map_snd conth]) let ?h = "k \ (\(t,x). (t,h x))" show ?thesis unfolding homotopic_with_def proof (intro exI conjI allI ballI) have "continuous_map (prod_topology (top_of_set {0..1}) X1) (prod_topology (top_of_set {0..1::real}) X2) (\(t, x). (t, h x))" by (metis (mono_tags, lifting) case_prod_beta' comp_def continuous_map_eq continuous_map_fst continuous_map_pairedI hsnd) then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) X3 ?h" by (intro conjI continuous_map_compose [OF _ contk]) show "q (\x. ?h (t, x))" if "t \ {0..1}" for t using q [OF p [OF that]] by (simp add: o_def) qed (auto simp: k) qed lemma homotopic_compose_continuous_map_right: "\homotopic_with (\k. True) X2 X3 f g; continuous_map X1 X2 h\ \ homotopic_with (\k. True) X1 X3 (f \ h) (g \ h)" by (meson homotopic_with_compose_continuous_map_right) corollary homotopic_compose: shows "\homotopic_with (\x. True) X Y f f'; homotopic_with (\x. True) Y Z g g'\ \ homotopic_with (\x. True) X Z (g \ f) (g' \ f')" apply (rule homotopic_with_trans [where g = "g \ f'"]) apply (simp add: homotopic_compose_continuous_map_left homotopic_with_imp_continuous_maps) by (simp add: homotopic_compose_continuous_map_right homotopic_with_imp_continuous_maps) proposition homotopic_with_compose_continuous_right: "\homotopic_with_canon (\f. p (f \ h)) X Y f g; continuous_on W h; h ` W \ X\ \ homotopic_with_canon 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_canon (\f. True) X Y f g; continuous_on W h; h ` W \ X\ \ homotopic_with_canon (\f. True) W Y (f \ h) (g \ h)" using homotopic_with_compose_continuous_right by fastforce proposition homotopic_with_compose_continuous_left: "\homotopic_with_canon (\f. p (h \ f)) X Y f g; continuous_on Y h; h ` Y \ Z\ \ homotopic_with_canon 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_canon (\_. True) X Y f g; continuous_on Y h; h ` Y \ Z\ \ homotopic_with_canon (\f. True) X Z (h \ f) (h \ g)" using homotopic_with_compose_continuous_left by fastforce lemma homotopic_from_subtopology: "homotopic_with P X X' f g \ homotopic_with P (subtopology X s) X' f g" unfolding homotopic_with_def apply (erule ex_forward) by (simp add: continuous_map_from_subtopology prod_topology_subtopology(2)) lemma homotopic_on_emptyI: assumes "topspace X = {}" "P f" "P g" shows "homotopic_with P X X' f g" unfolding homotopic_with_def proof (intro exI conjI ballI) show "P (\x. (\(t,x). if t = 0 then f x else g x) (t, x))" if "t \ {0..1}" for t::real by (cases "t = 0", auto simp: assms) qed (auto simp: continuous_map_atin assms) lemma homotopic_on_empty: "topspace X = {} \ (homotopic_with P X X' f g \ P f \ P g)" using homotopic_on_emptyI homotopic_with_imp_property by metis lemma homotopic_with_canon_on_empty [simp]: "homotopic_with_canon (\x. True) {} t f g" by (auto intro: homotopic_with_equal) lemma homotopic_constant_maps: "homotopic_with (\x. True) X X' (\x. a) (\x. b) \ topspace X = {} \ path_component_of X' a b" (is "?lhs = ?rhs") proof (cases "topspace X = {}") case False then obtain c where c: "c \ topspace X" by blast have "\g. continuous_map (top_of_set {0..1::real}) X' g \ g 0 = a \ g 1 = b" if "x \ topspace X" and hom: "homotopic_with (\x. True) X X' (\x. a) (\x. b)" for x proof - obtain h :: "real \ 'a \ 'b" where conth: "continuous_map (prod_topology (top_of_set {0..1}) X) X' h" and h: "\x. h (0, x) = a" "\x. h (1, x) = b" using hom by (auto simp: homotopic_with_def) have cont: "continuous_map (top_of_set {0..1}) X' (h \ (\t. (t, c)))" apply (rule continuous_map_compose [OF _ conth]) apply (rule continuous_intros c | simp)+ done then show ?thesis by (force simp: h) qed moreover have "homotopic_with (\x. True) X X' (\x. g 0) (\x. g 1)" if "x \ topspace X" "a = g 0" "b = g 1" "continuous_map (top_of_set {0..1}) X' g" for x and g :: "real \ 'b" unfolding homotopic_with_def by (force intro!: continuous_map_compose continuous_intros c that) ultimately show ?thesis using False by (auto simp: path_component_of_def pathin_def) qed (simp add: homotopic_on_empty) proposition homotopic_with_eq: assumes h: "homotopic_with P X Y f g" and f': "\x. x \ topspace X \ f' x = f x" and g': "\x. x \ topspace X \ g' x = g x" and P: "(\h k. (\x. x \ topspace 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 \ topspace 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_map_eq) apply (subst P; fastforce) done lemma homotopic_with_prod_topology: assumes "homotopic_with p X1 Y1 f f'" and "homotopic_with q X2 Y2 g g'" and r: "\i j. \p i; q j\ \ r(\(x,y). (i x, j y))" shows "homotopic_with r (prod_topology X1 X2) (prod_topology Y1 Y2) (\z. (f(fst z),g(snd z))) (\z. (f'(fst z), g'(snd z)))" proof - obtain h where h: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) Y1 h" and h0: "\x. h (0, x) = f x" and h1: "\x. h (1, x) = f' x" and p: "\t. \0 \ t; t \ 1\ \ p (\x. h (t,x))" using assms unfolding homotopic_with_def by auto obtain k where k: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X2) Y2 k" and k0: "\x. k (0, x) = g x" and k1: "\x. k (1, x) = g' x" and q: "\t. \0 \ t; t \ 1\ \ q (\x. k (t,x))" using assms unfolding homotopic_with_def by auto let ?hk = "\(t,x,y). (h(t,x), k(t,y))" show ?thesis unfolding homotopic_with_def proof (intro conjI allI exI) show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (prod_topology X1 X2)) (prod_topology Y1 Y2) ?hk" unfolding continuous_map_pairwise case_prod_unfold by (rule conjI continuous_map_pairedI continuous_intros continuous_map_id [unfolded id_def] continuous_map_fst_of [unfolded o_def] continuous_map_snd_of [unfolded o_def] continuous_map_compose [OF _ h, unfolded o_def] continuous_map_compose [OF _ k, unfolded o_def])+ next fix x show "?hk (0, x) = (f (fst x), g (snd x))" "?hk (1, x) = (f' (fst x), g' (snd x))" by (auto simp: case_prod_beta h0 k0 h1 k1) qed (auto simp: p q r) qed lemma homotopic_with_product_topology: assumes ht: "\i. i \ I \ homotopic_with (p i) (X i) (Y i) (f i) (g i)" and pq: "\h. (\i. i \ I \ p i (h i)) \ q(\x. (\i\I. h i (x i)))" shows "homotopic_with q (product_topology X I) (product_topology Y I) (\z. (\i\I. (f i) (z i))) (\z. (\i\I. (g i) (z i)))" proof - obtain h where h: "\i. i \ I \ continuous_map (prod_topology (subtopology euclideanreal {0..1}) (X i)) (Y i) (h i)" and h0: "\i x. i \ I \ h i (0, x) = f i x" and h1: "\i x. i \ I \ h i (1, x) = g i x" and p: "\i t. \i \ I; t \ {0..1}\ \ p i (\x. h i (t,x))" using ht unfolding homotopic_with_def by metis show ?thesis unfolding homotopic_with_def proof (intro conjI allI exI) let ?h = "\(t,z). \i\I. h i (t,z i)" have "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I)) (Y i) (\x. h i (fst x, snd x i))" if "i \ I" for i unfolding continuous_map_pairwise case_prod_unfold apply (intro conjI that continuous_intros continuous_map_compose [OF _ h, unfolded o_def]) using continuous_map_componentwise continuous_map_snd that apply fastforce done then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I)) (product_topology Y I) ?h" by (auto simp: continuous_map_componentwise case_prod_beta) show "?h (0, x) = (\i\I. f i (x i))" "?h (1, x) = (\i\I. g i (x i))" for x by (auto simp: case_prod_beta h0 h1) show "\t\{0..1}. q (\x. ?h (t, x))" by (force intro: p pq) qed qed text\Homotopic triviality implicitly incorporates path-connectedness.\ lemma homotopic_triviality: shows "(\f g. continuous_on S f \ f ` S \ T \ continuous_on S g \ g ` S \ T \ homotopic_with_canon (\x. True) S T f g) \ (S = {} \ path_connected T) \ (\f. continuous_on S f \ f ` S \ T \ (\c. homotopic_with_canon (\x. True) S T f (\x. c)))" (is "?lhs = ?rhs") proof (cases "S = {} \ T = {}") case True then show ?thesis by (auto simp: homotopic_on_emptyI) 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_canon (\x. True) S T (\x. a) (\x. b)" by (simp add: LHS image_subset_iff that) then show ?thesis using False homotopic_constant_maps [of "top_of_set S" "top_of_set T" a b] by auto qed moreover have "\c. homotopic_with_canon (\x. True) S T f (\x. c)" if "continuous_on S f" "f ` S \ T" for f using False LHS continuous_on_const that by blast 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_canon (\x. True) S T f (\x. c)" and d: "homotopic_with_canon (\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_continuous_maps by fastforce+ with T have "path_component T c d" using path_connected_component by blast then have "homotopic_with_canon (\x. True) S T (\x. c) (\x. d)" by (simp add: homotopic_constant_maps) with c d show "homotopic_with_canon (\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\<^marker>\tag important\ homotopic_paths :: "['a set, real \ 'a, real \ 'a::topological_space] \ bool" where "homotopic_paths s p q \ homotopic_with_canon (\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_maps path_def continuous_map_subtopology_eu by blast lemma homotopic_paths_imp_subset: "homotopic_paths s p q \ path_image p \ s \ path_image q \ s" by (metis (mono_tags) continuous_map_subtopology_eu homotopic_paths_def homotopic_with_imp_continuous_maps 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 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]: assumes "homotopic_paths s p q" "homotopic_paths s q r" shows "homotopic_paths s p r" proof - have "pathstart q = pathstart p" "pathfinish q = pathfinish p" using assms by (simp_all add: homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish) then have "homotopic_with_canon (\f. pathstart f = pathstart p \ pathfinish f = pathfinish p) {0..1} s q r" using \homotopic_paths s q r\ homotopic_paths_def by force then show ?thesis using assms homotopic_paths_def homotopic_with_trans by blast qed 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 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" unfolding homotopic_paths by fast 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 by (simp add: homotopic_with_compose_continuous_map_left pathfinish_compose pathstart_compose) subsection\Group properties for homotopy of paths\ text\<^marker>\tag 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\<^marker>\tag important\ homotopic_loops :: "'a::topological_space set \ (real \ 'a) \ (real \ 'a) \ bool" where "homotopic_loops s p q \ homotopic_with_canon (\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_maps continuous_map_subtopology_eu 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 (meson continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) 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 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 (fastforce simp add: homotopic_loops) 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 by (simp add: homotopic_with_compose_continuous_map_left 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_with_def homotopic_paths_def homotopic_loops_def) 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\<^marker>\tag 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_canon (\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\<^marker>\tag important\\defined as "all loops are homotopic (as loops)\ definition\<^marker>\tag 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\<^marker>\tag important\ contractible where "contractible S \ \a. homotopic_with_canon (\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_canon (\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_in_topspace topspace_euclidean_subtopology) 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_canon (\h. True) S U (g \ f) (\x. c)" proof - obtain b where b: "homotopic_with_canon (\x. True) T T id (\x. b)" using assms by (force simp: contractible_def) have "homotopic_with_canon (\f. True) T U (g \ id) (g \ (\x. b))" by (metis Abstract_Topology.continuous_map_subtopology_eu b g homotopic_compose_continuous_map_left) then have "homotopic_with_canon (\f. True) S U (g \ id \ f) (g \ (\x. b) \ f)" by (simp add: f homotopic_with_compose_continuous_map_right) 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_canon (\h. True) S T f (\x. c)" apply (rule nullhomotopic_through_contractible [OF f, of id T]) using assms apply (auto) done lemma nullhomotopic_from_contractible: assumes f: "continuous_on S f" "f ` S \ T" and S: "contractible S" obtains c where "homotopic_with_canon (\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_canon (\h. True) S U (g1 \ f1) (g2 \ f2)" proof - obtain c1 where c1: "homotopic_with_canon (\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_canon (\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_continuous_maps by fastforce+ 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_canon (\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) 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_canon (\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) +subsection\Starlike sets\ + +definition\<^marker>\tag important\ "starlike S \ (\a\S. \x\S. closed_segment a x \ S)" + +lemma starlike_UNIV [simp]: "starlike UNIV" + by (simp add: starlike_def) + +lemma convex_imp_starlike: + "convex S \ S \ {} \ starlike S" + unfolding convex_contains_segment starlike_def by auto + +lemma starlike_convex_tweak_boundary_points: + fixes S :: "'a::euclidean_space set" + assumes "convex S" "S \ {}" and ST: "rel_interior S \ T" and TS: "T \ closure S" + shows "starlike T" +proof - + have "rel_interior S \ {}" + by (simp add: assms rel_interior_eq_empty) + then obtain a where a: "a \ rel_interior S" by blast + with ST have "a \ T" by blast + have *: "\x. x \ T \ open_segment a x \ rel_interior S" + apply (rule rel_interior_closure_convex_segment [OF \convex S\ a]) + using assms by blast + show ?thesis + unfolding starlike_def + apply (rule bexI [OF _ \a \ T\]) + apply (simp add: closed_segment_eq_open) + apply (intro conjI ballI a \a \ T\ rel_interior_closure_convex_segment [OF \convex S\ a]) + apply (simp add: order_trans [OF * ST]) + done +qed + 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_canon 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_on_emptyI) 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 subsection\Local versions of topological properties in general\ definition\<^marker>\tag important\ locally :: "('a::topological_space set \ bool) \ 'a set \ bool" where "locally P S \ \w x. openin (top_of_set S) w \ x \ w \ (\u v. openin (top_of_set S) u \ P v \ x \ u \ u \ v \ v \ w)" lemma locallyI: assumes "\w x. \openin (top_of_set S) w; x \ w\ \ \u v. openin (top_of_set 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 (top_of_set S) w" "x \ w" obtains u v where "openin (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set (S \ T)) W" and xy: "(x, y) \ W" then obtain U V where "openin (top_of_set S) U" "x \ U" "openin (top_of_set T) V" "y \ V" "U \ V \ W" using Times_in_interior_subtopology by metis then obtain U1 U2 V1 V2 where opeS: "openin (top_of_set S) U1 \ P U2 \ x \ U1 \ U1 \ U2 \ U2 \ U" and opeT: "openin (top_of_set T) V1 \ Q V2 \ y \ V1 \ V1 \ V2 \ V2 \ V" by (meson PS QT locallyE) with \U \ V \ W\ show "\u v. openin (top_of_set (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 openin_Times_eq) 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 (top_of_set 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 (top_of_set S) (g ` W)" proof - have "continuous_on S f" using f(3) by blast then show "openin (top_of_set S) (g ` W)" by (simp add: gw Collect_conj_eq \openin (top_of_set t) W\ continuous_on_open f(2)) qed then obtain u v where osu: "openin (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set S) t \ openin (top_of_set (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 (top_of_set (f ` S)) W" and "y \ W" then have "W \ f ` S" by (simp add: openin_euclidean_subtopology_iff) have oivf: "openin (top_of_set 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 (top_of_set 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 (top_of_set (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 (top_of_set S) T; a \ T\ \ \z. z \ T \ P z" and opI: "\a. a \ S \ \T. openin (top_of_set 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 (top_of_set S) {b. \T. openin (top_of_set 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 (top_of_set S) {b. \T. openin (top_of_set 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 (top_of_set S) T; a \ T\ \ \z. z \ T \ P z" and opI: "\a. a \ S \ \T. openin (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set s) (u x) \ compact (v x)" by (metis locally_compactE) have *: "\u v. k \ u \ u \ v \ v \ s \ openin (top_of_set s) u \ compact v" if "k \ s" "compact k" for k proof - have "\C. (\c\C. openin (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set (S \ T)) S" and opT: "openin (top_of_set (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 (top_of_set (S \ T)) S" and clT: "closedin (top_of_set (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 (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set S) T \ closedin (top_of_set 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 (top_of_set (\?\)) K1" and K2: "closedin (top_of_set (\?\)) 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 (top_of_set S) D" and cloD: "closedin (top_of_set 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 (top_of_set S) (\\)" using openin_Inter \finite \\ False \ by blast then show "closedin (top_of_set 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 (top_of_set D) (D \ closure V1)" and cloV2: "closedin (top_of_set 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 (top_of_set D) (D \ V1)" and cloDV2: "closedin (top_of_set 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 (top_of_set S) K" "compact K" "C \ K" "K \ U" proof (rule ccontr) assume "\ thesis" with that have neg: "\K. openin (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set K) T \ closedin (top_of_set 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 (top_of_set K) (\\)" using False \ \finite \\ by blast then have opeVF: "openin (top_of_set V) (\\)" using W \K \ S\ \V \ K\ opeKF \\\ \ K\ FUW openin_subset_trans by fastforce then have "openin (top_of_set 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 (top_of_set S) U" and "C \ U" obtains K where "openin (top_of_set 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 (top_of_set S) K}" (is "C = ?rhs") proof show "?rhs \ C" proof (clarsimp, rule ccontr) fix x assume *: "\X. C \ X \ compact X \ openin (top_of_set 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 (top_of_set S) v; x \ v\ \ \u. openin (top_of_set 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 (top_of_set S) t" "x \ t" shows "openin (top_of_set S) (connected_component_set t x)" proof - { fix y :: 'a let ?SS = "top_of_set 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 (top_of_set S) t; x \ t\ \ openin (top_of_set S) (connected_component_set t x)" "openin (top_of_set S) v" "x \ v" shows "\u. openin (top_of_set S) u \ connected u \ x \ u \ u \ v" using assms connected_component_subset by fastforce lemma locally_connected: "locally connected S \ (\v x. openin (top_of_set S) v \ x \ v \ (\u. openin (top_of_set 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 (top_of_set S) t \ x \ t \ openin (top_of_set 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 (top_of_set S) v; x \ v\ \ \u. openin (top_of_set 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 (top_of_set S) t" "x \ t" shows "openin (top_of_set S) (path_component_set t x)" proof - { fix y :: 'a let ?SS = "top_of_set 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 rev_subsetD 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 (top_of_set S) t; x \ t\ \ openin (top_of_set S) (path_component_set t x)" "openin (top_of_set S) v" "x \ v" shows "\u. openin (top_of_set 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 (top_of_set S) v \ x \ v \ (\u. openin (top_of_set 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 (top_of_set S) t \ x \ t \ openin (top_of_set 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 (top_of_set S) t \ c \ components t \ openin (top_of_set S) c)" by (metis components_iff locally_connected_open_connected_component) proposition locally_connected_im_kleinen: "locally connected S \ (\v x. openin (top_of_set S) v \ x \ v \ (\u. openin (top_of_set 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 (top_of_set S) T \ x \ T \ T \ c" if "openin (top_of_set 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 (top_of_set 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 (top_of_set S) v \ x \ v \ (\u. openin (top_of_set 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 (top_of_set S) T \ x \ T \ T \ path_component_set u z" if "openin (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set S) c" using locally_connected_open_component openin_subtopology_self by blast lemma openin_path_component_locally_path_connected: "locally path_connected S \ openin (top_of_set 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 (top_of_set 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 (top_of_set (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 (top_of_set (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 (top_of_set S) (S \ f -` T) \ openin (top_of_set (f ` S)) T" shows "locally connected (f ` S)" proof (clarsimp simp: locally_connected_open_component) fix U C assume opefSU: "openin (top_of_set (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 (top_of_set (f ` S)) C \ openin (top_of_set S) (S \ f -` C)" by (auto simp: oo) moreover have "openin (top_of_set S) (S \ f -` C)" proof (subst openin_subopen, clarify) fix x assume "x \ S" "f x \ C" show "\T. openin (top_of_set S) T \ x \ T \ T \ (S \ f -` C)" proof (intro conjI exI) show "openin (top_of_set S) (connected_component_set (S \ f -` U) x)" proof (rule ccontr) assume **: "\ openin (top_of_set 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 (top_of_set (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 (top_of_set S) (S \ f -` T) \ openin (top_of_set (f ` S)) T" shows "locally path_connected (f ` S)" proof (clarsimp simp: locally_path_connected_open_path_component) fix U y assume opefSU: "openin (top_of_set (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 (top_of_set (f ` S)) (path_component_set U y) \ openin (top_of_set 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 (top_of_set 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 (top_of_set S) T \ x \ T \ T \ (S \ f -` path_component_set U y)" proof (intro conjI exI) show "openin (top_of_set S) (path_component_set (S \ f -` U) x)" proof (rule ccontr) assume **: "\ openin (top_of_set 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 (top_of_set (f ` S)) (path_component_set U y)" by metis qed subsection\<^marker>\tag unimportant\\Components, continuity, openin, closedin\ lemma continuous_on_components_gen: fixes f :: "'a::topological_space \ 'b::topological_space" assumes "\c. c \ components S \ openin (top_of_set 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 (top_of_set 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 (top_of_set u) S" and cuS: "c \ components(u - S)" shows "closedin (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set 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 (top_of_set u) S" and c: " c \ components(u - S)" shows "closedin (top_of_set u) (S \ c)" proof - have "closedin (top_of_set 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 (fact homeomorphic_translation) also have "\ homeomorphic ((+) (- b) ` T)" by (rule homeomorphic_subspaces [OF ss dd]) also have "\ homeomorphic T" using homeomorphic_translation [of T "- b"] by (simp add: homeomorphic_sym [of T]) finally show ?thesis . qed subsection\Retracts, in a general sense, preserve (co)homotopic triviality)\ locale\<^marker>\tag 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_canon 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_canon 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_canon P u s (k \ f) (k \ g)" by (rule hom) then have "homotopic_with_canon 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) using feq geq apply fastforce+ using Qeq topspace_euclidean_subtopology by blast 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_canon 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_canon 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_canon P u s (k \ f) (\x. c)" by (metis hom) then have "homotopic_with_canon 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) using feq apply auto[1] apply simp using Qeq topspace_euclidean_subtopology by blast 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_canon 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_canon 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_canon P s u (f \ h) (g \ h)" by (rule hom) then have "homotopic_with_canon 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) using feq apply auto[1] using geq apply auto[1] using Qeq topspace_euclidean_subtopology by blast 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_canon 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_canon 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_canon P s u (f \ h) (\x. c)" by (metis hom) then have "homotopic_with_canon 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) using feq apply auto[1] apply simp using Qeq topspace_euclidean_subtopology by blast 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\ subsection\Homotopy equivalence of topological spaces.\ definition\<^marker>\tag important\ homotopy_equivalent_space (infix "homotopy'_equivalent'_space" 50) where "X homotopy_equivalent_space Y \ (\f g. continuous_map X Y f \ continuous_map Y X g \ homotopic_with (\x. True) X X (g \ f) id \ homotopic_with (\x. True) Y Y (f \ g) id)" lemma homeomorphic_imp_homotopy_equivalent_space: "X homeomorphic_space Y \ X homotopy_equivalent_space Y" unfolding homeomorphic_space_def homotopy_equivalent_space_def apply (erule ex_forward)+ by (simp add: homotopic_with_equal homotopic_with_sym homeomorphic_maps_def) lemma homotopy_equivalent_space_refl: "X homotopy_equivalent_space X" by (simp add: homeomorphic_imp_homotopy_equivalent_space homeomorphic_space_refl) lemma homotopy_equivalent_space_sym: "X homotopy_equivalent_space Y \ Y homotopy_equivalent_space X" by (meson homotopy_equivalent_space_def) lemma homotopy_eqv_trans [trans]: assumes 1: "X homotopy_equivalent_space Y" and 2: "Y homotopy_equivalent_space U" shows "X homotopy_equivalent_space U" proof - obtain f1 g1 where f1: "continuous_map X Y f1" and g1: "continuous_map Y X g1" and hom1: "homotopic_with (\x. True) X X (g1 \ f1) id" "homotopic_with (\x. True) Y Y (f1 \ g1) id" using 1 by (auto simp: homotopy_equivalent_space_def) obtain f2 g2 where f2: "continuous_map Y U f2" and g2: "continuous_map U Y g2" and hom2: "homotopic_with (\x. True) Y Y (g2 \ f2) id" "homotopic_with (\x. True) U U (f2 \ g2) id" using 2 by (auto simp: homotopy_equivalent_space_def) have "homotopic_with (\f. True) X Y (g2 \ f2 \ f1) (id \ f1)" using f1 hom2(1) homotopic_compose_continuous_map_right by blast then have "homotopic_with (\f. True) X Y (g2 \ (f2 \ f1)) (id \ f1)" by (simp add: o_assoc) then have "homotopic_with (\x. True) X X (g1 \ (g2 \ (f2 \ f1))) (g1 \ (id \ f1))" by (simp add: g1 homotopic_compose_continuous_map_left) moreover have "homotopic_with (\x. True) X X (g1 \ id \ f1) id" using hom1 by simp ultimately have SS: "homotopic_with (\x. True) X X (g1 \ g2 \ (f2 \ f1)) id" apply (simp add: o_assoc) apply (blast intro: homotopic_with_trans) done have "homotopic_with (\f. True) U Y (f1 \ g1 \ g2) (id \ g2)" using g2 hom1(2) homotopic_with_compose_continuous_map_right by fastforce then have "homotopic_with (\f. True) U Y (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_map_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_equivalent_space_def by (blast intro: f1 f2 g1 g2 continuous_map_compose SS UU) qed lemma deformation_retraction_imp_homotopy_equivalent_space: "\homotopic_with (\x. True) X X (s \ r) id; retraction_maps X Y r s\ \ X homotopy_equivalent_space Y" unfolding homotopy_equivalent_space_def retraction_maps_def apply (rule_tac x=r in exI) apply (rule_tac x=s in exI) apply (auto simp: homotopic_with_equal continuous_map_compose) done lemma deformation_retract_imp_homotopy_equivalent_space: "\homotopic_with (\x. True) X X r id; retraction_maps X Y r id\ \ X homotopy_equivalent_space Y" using deformation_retraction_imp_homotopy_equivalent_space by force lemma deformation_retract_of_space: "S \ topspace X \ (\r. homotopic_with (\x. True) X X id r \ retraction_maps X (subtopology X S) r id) \ S retract_of_space X \ (\f. homotopic_with (\x. True) X X id f \ f ` (topspace X) \ S)" proof (cases "S \ topspace X") case True moreover have "(\r. homotopic_with (\x. True) X X id r \ retraction_maps X (subtopology X S) r id) \ (S retract_of_space X \ (\f. homotopic_with (\x. True) X X id f \ f ` topspace X \ S))" unfolding retract_of_space_def proof safe fix f r assume f: "homotopic_with (\x. True) X X id f" and fS: "f ` topspace X \ S" and r: "continuous_map X (subtopology X S) r" and req: "\x\S. r x = x" show "\r. homotopic_with (\x. True) X X id r \ retraction_maps X (subtopology X S) r id" proof (intro exI conjI) have "homotopic_with (\x. True) X X f r" proof (rule homotopic_with_eq) show "homotopic_with (\x. True) X X (r \ f) (r \ id)" using homotopic_with_symD continuous_map_into_fulltopology f homotopic_compose_continuous_map_left r by blast show "f x = (r \ f) x" if "x \ topspace X" for x using that fS req by auto qed auto then show "homotopic_with (\x. True) X X id r" by (rule homotopic_with_trans [OF f]) next show "retraction_maps X (subtopology X S) r id" by (simp add: r req retraction_maps_def) qed qed (use True in \auto simp: retraction_maps_def topspace_subtopology_subset continuous_map_in_subtopology\) ultimately show ?thesis by simp qed (auto simp: retract_of_space_def retraction_maps_def) subsection\Contractible spaces\ text\The definition (which agrees with "contractible" on subsets of Euclidean space) is a little cryptic because we don't in fact assume that the constant "a" is in the space. This forces the convention that the empty space / set is contractible, avoiding some special cases. \ definition contractible_space where "contractible_space X \ \a. homotopic_with (\x. True) X X id (\x. a)" lemma contractible_space_top_of_set [simp]:"contractible_space (top_of_set S) \ contractible S" by (auto simp: contractible_space_def contractible_def) lemma contractible_space_empty: "topspace X = {} \ contractible_space X" apply (simp add: contractible_space_def homotopic_with_def) apply (rule_tac x=undefined in exI) apply (rule_tac x="\(t,x). if t = 0 then x else undefined" in exI) apply (auto simp: continuous_map_on_empty) done lemma contractible_space_singleton: "topspace X = {a} \ contractible_space X" apply (simp add: contractible_space_def homotopic_with_def) apply (rule_tac x=a in exI) apply (rule_tac x="\(t,x). if t = 0 then x else a" in exI) apply (auto intro: continuous_map_eq [where f = "\z. a"]) done lemma contractible_space_subset_singleton: "topspace X \ {a} \ contractible_space X" by (meson contractible_space_empty contractible_space_singleton subset_singletonD) lemma contractible_space_subtopology_singleton: "contractible_space(subtopology X {a})" by (meson contractible_space_subset_singleton insert_subset path_connectedin_singleton path_connectedin_subtopology subsetI) lemma contractible_space: "contractible_space X \ topspace X = {} \ (\a \ topspace X. homotopic_with (\x. True) X X id (\x. a))" proof (cases "topspace X = {}") case False then show ?thesis apply (auto simp: contractible_space_def) using homotopic_with_imp_continuous_maps by fastforce qed (simp add: contractible_space_empty) lemma contractible_imp_path_connected_space: assumes "contractible_space X" shows "path_connected_space X" proof (cases "topspace X = {}") case False have *: "path_connected_space X" if "a \ topspace X" and conth: "continuous_map (prod_topology (top_of_set {0..1}) X) X h" and h: "\x. h (0, x) = x" "\x. h (1, x) = a" for a and h :: "real \ 'a \ 'a" proof - have "path_component_of X b a" if "b \ topspace X" for b using that unfolding path_component_of_def apply (rule_tac x="h \ (\x. (x,b))" in exI) apply (simp add: h pathin_def) apply (rule continuous_map_compose [OF _ conth]) apply (auto simp: continuous_map_pairwise intro!: continuous_intros continuous_map_compose continuous_map_id [unfolded id_def]) done then show ?thesis by (metis path_component_of_equiv path_connected_space_iff_path_component) qed show ?thesis using assms False by (auto simp: contractible_space homotopic_with_def *) qed (simp add: path_connected_space_topspace_empty) lemma contractible_imp_connected_space: "contractible_space X \ connected_space X" by (simp add: contractible_imp_path_connected_space path_connected_imp_connected_space) lemma contractible_space_alt: "contractible_space X \ (\a \ topspace X. homotopic_with (\x. True) X X id (\x. a))" (is "?lhs = ?rhs") proof assume X: ?lhs then obtain a where a: "homotopic_with (\x. True) X X id (\x. a)" by (auto simp: contractible_space_def) show ?rhs proof show "homotopic_with (\x. True) X X id (\x. b)" if "b \ topspace X" for b apply (rule homotopic_with_trans [OF a]) using homotopic_constant_maps path_connected_space_imp_path_component_of by (metis (full_types) X a continuous_map_const contractible_imp_path_connected_space homotopic_with_imp_continuous_maps that) qed next assume R: ?rhs then show ?lhs apply (simp add: contractible_space_def) by (metis equals0I homotopic_on_emptyI) qed lemma compose_const [simp]: "f \ (\x. a) = (\x. f a)" "(\x. a) \ g = (\x. a)" by (simp_all add: o_def) lemma nullhomotopic_through_contractible_space: assumes f: "continuous_map X Y f" and g: "continuous_map Y Z g" and Y: "contractible_space Y" obtains c where "homotopic_with (\h. True) X Z (g \ f) (\x. c)" proof - obtain b where b: "homotopic_with (\x. True) Y Y id (\x. b)" using Y by (auto simp: contractible_space_def) show thesis using homotopic_compose_continuous_map_right [OF homotopic_compose_continuous_map_left [OF b g] f] by (simp add: that) qed lemma nullhomotopic_into_contractible_space: assumes f: "continuous_map X Y f" and Y: "contractible_space Y" obtains c where "homotopic_with (\h. True) X Y f (\x. c)" using nullhomotopic_through_contractible_space [OF f _ Y] by (metis continuous_map_id id_comp) lemma nullhomotopic_from_contractible_space: assumes f: "continuous_map X Y f" and X: "contractible_space X" obtains c where "homotopic_with (\h. True) X Y f (\x. c)" using nullhomotopic_through_contractible_space [OF _ f X] by (metis comp_id continuous_map_id) lemma homotopy_dominated_contractibility: assumes f: "continuous_map X Y f" and g: "continuous_map Y X g" and hom: "homotopic_with (\x. True) Y Y (f \ g) id" and X: "contractible_space X" shows "contractible_space Y" proof - obtain c where c: "homotopic_with (\h. True) X Y f (\x. c)" using nullhomotopic_from_contractible_space [OF f X] . have "homotopic_with (\x. True) Y Y (f \ g) (\x. c)" using homotopic_compose_continuous_map_right [OF c g] by fastforce then have "homotopic_with (\x. True) Y Y id (\x. c)" using homotopic_with_trans [OF _ hom] homotopic_with_symD by blast then show ?thesis unfolding contractible_space_def .. qed lemma homotopy_equivalent_space_contractibility: "X homotopy_equivalent_space Y \ (contractible_space X \ contractible_space Y)" unfolding homotopy_equivalent_space_def by (blast intro: homotopy_dominated_contractibility) lemma homeomorphic_space_contractibility: "X homeomorphic_space Y \ (contractible_space X \ contractible_space Y)" by (simp add: homeomorphic_imp_homotopy_equivalent_space homotopy_equivalent_space_contractibility) lemma contractible_eq_homotopy_equivalent_singleton_subtopology: "contractible_space X \ topspace X = {} \ (\a \ topspace X. X homotopy_equivalent_space (subtopology X {a}))"(is "?lhs = ?rhs") proof (cases "topspace X = {}") case False show ?thesis proof assume ?lhs then obtain a where a: "homotopic_with (\x. True) X X id (\x. a)" by (auto simp: contractible_space_def) then have "a \ topspace X" by (metis False continuous_map_const homotopic_with_imp_continuous_maps) then have "X homotopy_equivalent_space subtopology X {a}" unfolding homotopy_equivalent_space_def apply (rule_tac x="\x. a" in exI) apply (rule_tac x=id in exI) apply (auto simp: homotopic_with_sym topspace_subtopology_subset a) using connectedin_absolute connectedin_sing contractible_space_alt contractible_space_subtopology_singleton by fastforce with \a \ topspace X\ show ?rhs by blast next assume ?rhs then show ?lhs by (meson False contractible_space_subtopology_singleton homotopy_equivalent_space_contractibility) qed qed (simp add: contractible_space_empty) lemma contractible_space_retraction_map_image: assumes "retraction_map X Y f" and X: "contractible_space X" shows "contractible_space Y" proof - obtain g where f: "continuous_map X Y f" and g: "continuous_map Y X g" and fg: "\y \ topspace Y. f(g y) = y" using assms by (auto simp: retraction_map_def retraction_maps_def) obtain a where a: "homotopic_with (\x. True) X X id (\x. a)" using X by (auto simp: contractible_space_def) have "homotopic_with (\x. True) Y Y id (\x. f a)" proof (rule homotopic_with_eq) show "homotopic_with (\x. True) Y Y (f \ id \ g) (f \ (\x. a) \ g)" using f g a homotopic_compose_continuous_map_left homotopic_compose_continuous_map_right by metis qed (use fg in auto) then show ?thesis unfolding contractible_space_def by blast qed lemma contractible_space_prod_topology: "contractible_space(prod_topology X Y) \ topspace X = {} \ topspace Y = {} \ contractible_space X \ contractible_space Y" proof (cases "topspace X = {} \ topspace Y = {}") case True then have "topspace (prod_topology X Y) = {}" by simp then show ?thesis by (auto simp: contractible_space_empty) next case False have "contractible_space(prod_topology X Y) \ contractible_space X \ contractible_space Y" proof safe assume XY: "contractible_space (prod_topology X Y)" with False have "retraction_map (prod_topology X Y) X fst" by (auto simp: contractible_space False retraction_map_fst) then show "contractible_space X" by (rule contractible_space_retraction_map_image [OF _ XY]) have "retraction_map (prod_topology X Y) Y snd" using False XY by (auto simp: contractible_space False retraction_map_snd) then show "contractible_space Y" by (rule contractible_space_retraction_map_image [OF _ XY]) next assume "contractible_space X" and "contractible_space Y" with False obtain a b where "a \ topspace X" and a: "homotopic_with (\x. True) X X id (\x. a)" and "b \ topspace Y" and b: "homotopic_with (\x. True) Y Y id (\x. b)" by (auto simp: contractible_space) with False show "contractible_space (prod_topology X Y)" apply (simp add: contractible_space) apply (rule_tac x=a in bexI) apply (rule_tac x=b in bexI) using homotopic_with_prod_topology [OF a b] apply (metis (no_types, lifting) case_prod_Pair case_prod_beta' eq_id_iff) apply auto done qed with False show ?thesis by auto qed lemma contractible_space_product_topology: "contractible_space(product_topology X I) \ topspace (product_topology X I) = {} \ (\i \ I. contractible_space(X i))" proof (cases "topspace (product_topology X I) = {}") case False have 1: "contractible_space (X i)" if XI: "contractible_space (product_topology X I)" and "i \ I" for i proof (rule contractible_space_retraction_map_image [OF _ XI]) show "retraction_map (product_topology X I) (X i) (\x. x i)" using False by (simp add: retraction_map_product_projection \i \ I\) qed have 2: "contractible_space (product_topology X I)" if "x \ topspace (product_topology X I)" and cs: "\i\I. contractible_space (X i)" for x :: "'a \ 'b" proof - obtain f where f: "\i. i\I \ homotopic_with (\x. True) (X i) (X i) id (\x. f i)" using cs unfolding contractible_space_def by metis have "homotopic_with (\x. True) (product_topology X I) (product_topology X I) id (\x. restrict f I)" by (rule homotopic_with_eq [OF homotopic_with_product_topology [OF f]]) (auto) then show ?thesis by (auto simp: contractible_space_def) qed show ?thesis using False 1 2 by blast qed (simp add: contractible_space_empty) lemma contractible_space_subtopology_euclideanreal [simp]: "contractible_space(subtopology euclideanreal S) \ is_interval S" (is "?lhs = ?rhs") proof assume ?lhs then have "path_connectedin (subtopology euclideanreal S) S" using contractible_imp_path_connected_space path_connectedin_topspace path_connectedin_absolute by (simp add: contractible_imp_path_connected) then show ?rhs by (simp add: is_interval_path_connected_1) next assume ?rhs then have "convex S" by (simp add: is_interval_convex_1) show ?lhs proof (cases "S = {}") case False then obtain z where "z \ S" by blast show ?thesis unfolding contractible_space_def homotopic_with_def proof (intro exI conjI allI) show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set S)) (top_of_set S) (\(t,x). (1 - t) * x + t * z)" apply (auto simp: case_prod_unfold) apply (intro continuous_intros) using \z \ S\ apply (auto intro: convexD [OF \convex S\, simplified]) done qed auto qed (simp add: contractible_space_empty) qed corollary contractible_space_euclideanreal: "contractible_space euclideanreal" proof - have "contractible_space (subtopology euclideanreal UNIV)" using contractible_space_subtopology_euclideanreal by blast then show ?thesis by simp qed abbreviation\<^marker>\tag important\ homotopy_eqv :: "'a::topological_space set \ 'b::topological_space set \ bool" (infix "homotopy'_eqv" 50) where "S homotopy_eqv T \ top_of_set S homotopy_equivalent_space top_of_set T" lemma homeomorphic_imp_homotopy_eqv: "S homeomorphic T \ S homotopy_eqv T" unfolding homeomorphic_def homeomorphism_def homotopy_equivalent_space_def apply (erule ex_forward)+ apply auto apply (metis homotopic_with_id2 topspace_euclidean_subtopology) by (metis (full_types) homotopic_with_id2 imageE topspace_euclidean_subtopology) 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_canon (\x. True) U S f g" shows "homotopic_with_canon (\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_canon (\x. True) S S (k \ h) id" "homotopic_with_canon (\x. True) T T (h \ k) id" using assms by (auto simp: homotopy_equivalent_space_def) have "homotopic_with_canon (\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_canon (\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_canon (\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_canon (\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_canon (\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_canon (\x. True) U S f g) \ (\f g. continuous_on U f \ f ` U \ T \ continuous_on U g \ g ` U \ T \ homotopic_with_canon (\x. True) U T f g)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis assms homotopy_eqv_homotopic_triviality_imp) next assume ?rhs moreover have "T homotopy_eqv S" using assms homotopy_equivalent_space_sym by blast ultimately show ?lhs by (blast intro: homotopy_eqv_homotopic_triviality_imp) qed 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_canon (\x. True) S U f (\x. c)" obtains c where "homotopic_with_canon (\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_canon (\x. True) S S (k \ h) id" "homotopic_with_canon (\x. True) T T (h \ k) id" using assms by (auto simp: homotopy_equivalent_space_def) obtain c where "homotopic_with_canon (\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_canon (\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_canon (\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_canon (\x. True) S U f (\x. c))) \ (\f. continuous_on T f \ f ` T \ U \ (\c. homotopic_with_canon (\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_equivalent_space_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_canon (\x. True) U S f (\x. c)" shows "\c. homotopic_with_canon (\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_canon (\x. True) S S (k \ h) id" "homotopic_with_canon (\x. True) T T (h \ k) id" using assms by (auto simp: homotopy_equivalent_space_def) obtain c::'a where "homotopic_with_canon (\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_canon (\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_canon (\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_canon (\x. True) U S f (\x. c))) \ (\f. continuous_on U f \ f ` U \ T \ (\c. homotopic_with_canon (\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_equivalent_space_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_equivalent_space_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) 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) apply (metis Abstract_Topology.continuous_map_subtopology_eu emptyE equals0I homotopy_equivalent_space_def image_subset_iff) 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 = {}" using homotopy_equivalent_space_sym homotopy_eqv_empty1 by blast 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)" by (meson contractible_space_top_of_set homotopy_equivalent_space_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\<^marker>\tag 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 corollary bounded_path_connected_Compl_real: fixes S :: "real set" assumes "bounded S" "path_connected(- S)" shows "S = {}" by (simp add: assms bounded_connected_Compl_real path_connected_imp_connected) 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\<^marker>\tag 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\<^marker>\tag 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 (top_of_set (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 (top_of_set S) U" and "x \ U" for U x proof - have "openin (top_of_set (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 (top_of_set 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 (top_of_set (affine hull S)) (ball x r \ affine hull S)" by (subst inf.commute) (simp add: openin_Int_open) then show "openin (top_of_set 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 (top_of_set (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) next case False show ?thesis proof (rule path_connected_openin_diff_countable) show "openin (top_of_set (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\<^marker>\tag unimportant\ \Self-homeomorphisms shuffling points about\ subsubsection\<^marker>\tag 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: field_split_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 then have n: "norm (a - (x - v *\<^sub>R (u - a))) = r - r * v" by (simp add: field_simps norm_minus_commute) 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) (simp add: field_simps) have "x - v *\<^sub>R (u - a) \ cball a r" using \r > 0\\0 \ v\ by (simp add: dist_norm n) moreover have "x - v *\<^sub>R (u - a) \ T" by (simp add: f_def \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\<^marker>\tag 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\<^marker>\tag 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\ 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\ 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\<^marker>\tag unimportant\ homeomorphism_moving_point: fixes a :: "'a::euclidean_space" assumes ope: "openin (top_of_set (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 (top_of_set 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 (top_of_set (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 (top_of_set (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\<^marker>\tag 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 (top_of_set (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\<^marker>\tag 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: field_split_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 field_split_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)+ 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 (top_of_set (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\<^marker>\tag 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\<^marker>\tag unimportant\ homeomorphism_grouping_points_exists_gen: fixes S :: "'a::euclidean_space set" assumes opeU: "openin (top_of_set S) U" and opeS: "openin (top_of_set (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 (top_of_set (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 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 (top_of_set (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_canon (\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 less then show ?thesis by (simp add: homotopic_on_emptyI) next 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_canon (\x. True) (sphere a r) S f (\x. c)" then have contf: "continuous_on (sphere a r) f" by (metis homotopic_with_imp_continuous) moreover have fim: "f ` sphere a r \ S" by (meson continuous_map_subtopology_eu c homotopic_with_imp_continuous_maps) 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_canon (\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 (simp_all add: dist_norm) (simp add: field_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" proof (rule continuous_on_compose2 [OF contg]) show "continuous_on ({0..1} \ sphere a r) (\x. a + fst x *\<^sub>R (snd x - a))" by (intro continuous_intros) qed (auto simp: dist_norm norm_minus_commute mult_left_le_one_le) 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 have "homotopic_with_canon (\x. True) (sphere a r) S (\x. g a) f" by (auto simp: homotopic_with) then show ?lhs using homotopic_with_symD by blast qed ultimately show ?thesis by meson qed end \ No newline at end of file diff --git a/src/HOL/Analysis/Starlike.thy b/src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy +++ b/src/HOL/Analysis/Starlike.thy @@ -1,6664 +1,6631 @@ (* Title: HOL/Analysis/Starlike.thy Author: L C Paulson, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Bogdan Grechuk, University of Edinburgh Author: Armin Heller, TU Muenchen Author: Johannes Hoelzl, TU Muenchen *) chapter \Unsorted\ theory Starlike imports Convex_Euclidean_Space Abstract_Limits Line_Segment begin -subsection\Starlike sets\ - -definition\<^marker>\tag important\ "starlike S \ (\a\S. \x\S. closed_segment a x \ S)" - -lemma starlike_UNIV [simp]: "starlike UNIV" - by (simp add: starlike_def) - -lemma convex_imp_starlike: - "convex S \ S \ {} \ starlike S" - unfolding convex_contains_segment starlike_def by auto - - lemma affine_hull_closed_segment [simp]: "affine hull (closed_segment a b) = affine hull {a,b}" by (simp add: segment_convex_hull) lemma affine_hull_open_segment [simp]: fixes a :: "'a::euclidean_space" shows "affine hull (open_segment a b) = (if a = b then {} else affine hull {a,b})" by (metis affine_hull_convex_hull affine_hull_empty closure_open_segment closure_same_affine_hull segment_convex_hull) lemma rel_interior_closure_convex_segment: fixes S :: "_::euclidean_space set" assumes "convex S" "a \ rel_interior S" "b \ closure S" shows "open_segment a b \ rel_interior S" proof fix x have [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" for u by (simp add: algebra_simps) assume "x \ open_segment a b" then show "x \ rel_interior S" unfolding closed_segment_def open_segment_def using assms by (auto intro: rel_interior_closure_convex_shrink) qed lemma convex_hull_insert_segments: "convex hull (insert a S) = (if S = {} then {a} else \x \ convex hull S. closed_segment a x)" by (force simp add: convex_hull_insert_alt in_segment) lemma Int_convex_hull_insert_rel_exterior: fixes z :: "'a::euclidean_space" assumes "convex C" "T \ C" and z: "z \ rel_interior C" and dis: "disjnt S (rel_interior C)" shows "S \ (convex hull (insert z T)) = S \ (convex hull T)" (is "?lhs = ?rhs") proof have "T = {} \ z \ S" using dis z by (auto simp add: disjnt_def) then show "?lhs \ ?rhs" proof (clarsimp simp add: convex_hull_insert_segments) fix x y assume "x \ S" and y: "y \ convex hull T" and "x \ closed_segment z y" have "y \ closure C" by (metis y \convex C\ \T \ C\ closure_subset contra_subsetD convex_hull_eq hull_mono) moreover have "x \ rel_interior C" by (meson \x \ S\ dis disjnt_iff) moreover have "x \ open_segment z y \ {z, y}" using \x \ closed_segment z y\ closed_segment_eq_open by blast ultimately show "x \ convex hull T" using rel_interior_closure_convex_segment [OF \convex C\ z] using y z by blast qed show "?rhs \ ?lhs" by (meson hull_mono inf_mono subset_insertI subset_refl) qed subsection\<^marker>\tag unimportant\ \Shrinking towards the interior of a convex set\ lemma mem_interior_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ interior S" and "x \ S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ interior S" proof - obtain d where "d > 0" and d: "ball c d \ S" using assms(2) unfolding mem_interior by auto show ?thesis unfolding mem_interior proof (intro exI subsetI conjI) fix y assume "y \ ball (x - e *\<^sub>R (x - c)) (e*d)" then have as: "dist (x - e *\<^sub>R (x - c)) y < e * d" by simp have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using \e > 0\ by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \1/e\ * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" unfolding dist_norm unfolding norm_scaleR[symmetric] apply (rule arg_cong[where f=norm]) using \e > 0\ by (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) also have "\ = \1/e\ * norm (x - e *\<^sub>R (x - c) - y)" by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and \e > 0\ by (auto simp add:pos_divide_less_eq[OF \e > 0\] mult.commute) finally show "y \ S" apply (subst *) apply (rule assms(1)[unfolded convex_alt,rule_format]) apply (rule d[unfolded subset_eq,rule_format]) unfolding mem_ball using assms(3-5) apply auto done qed (insert \e>0\ \d>0\, auto) qed lemma mem_interior_closure_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ interior S" and "x \ closure S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ interior S" proof - obtain d where "d > 0" and d: "ball c d \ S" using assms(2) unfolding mem_interior by auto have "\y\S. norm (y - x) * (1 - e) < e * d" proof (cases "x \ S") case True then show ?thesis using \e > 0\ \d > 0\ apply (rule_tac bexI[where x=x]) apply (auto) done next case False then have x: "x islimpt S" using assms(3)[unfolded closure_def] by auto show ?thesis proof (cases "e = 1") case True obtain y where "y \ S" "y \ x" "dist y x < 1" using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto then show ?thesis apply (rule_tac x=y in bexI) unfolding True using \d > 0\ apply auto done next case False then have "0 < e * d / (1 - e)" and *: "1 - e > 0" using \e \ 1\ \e > 0\ \d > 0\ by auto then obtain y where "y \ S" "y \ x" "dist y x < e * d / (1 - e)" using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto then show ?thesis apply (rule_tac x=y in bexI) unfolding dist_norm using pos_less_divide_eq[OF *] apply auto done qed qed then obtain y where "y \ S" and y: "norm (y - x) * (1 - e) < e * d" by auto define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using \e > 0\ by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) have "z \ interior S" apply (rule interior_mono[OF d,unfolded subset_eq,rule_format]) unfolding interior_open[OF open_ball] mem_ball z_def dist_norm using y and assms(4,5) by simp (simp add: field_simps norm_minus_commute) then show ?thesis unfolding * using mem_interior_convex_shrink \y \ S\ assms by blast qed lemma in_interior_closure_convex_segment: fixes S :: "'a::euclidean_space set" assumes "convex S" and a: "a \ interior S" and b: "b \ closure S" shows "open_segment a b \ interior S" proof (clarsimp simp: in_segment) fix u::real assume u: "0 < u" "u < 1" have "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" by (simp add: algebra_simps) also have "... \ interior S" using mem_interior_closure_convex_shrink [OF assms] u by simp finally show "(1 - u) *\<^sub>R a + u *\<^sub>R b \ interior S" . qed lemma convex_closure_interior: fixes S :: "'a::euclidean_space set" assumes "convex S" and int: "interior S \ {}" shows "closure(interior S) = closure S" proof - obtain a where a: "a \ interior S" using int by auto have "closure S \ closure(interior S)" proof fix x assume x: "x \ closure S" show "x \ closure (interior S)" proof (cases "x=a") case True then show ?thesis using \a \ interior S\ closure_subset by blast next case False show ?thesis proof (clarsimp simp add: closure_def islimpt_approachable) fix e::real assume xnotS: "x \ interior S" and "0 < e" show "\x'\interior S. x' \ x \ dist x' x < e" proof (intro bexI conjI) show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \ x" using False \0 < e\ by (auto simp: algebra_simps min_def) show "dist (x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a)) x < e" using \0 < e\ by (auto simp: dist_norm min_def) show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \ interior S" apply (clarsimp simp add: min_def a) apply (rule mem_interior_closure_convex_shrink [OF \convex S\ a x]) using \0 < e\ False apply (auto simp: field_split_simps) done qed qed qed qed then show ?thesis by (simp add: closure_mono interior_subset subset_antisym) qed lemma closure_convex_Int_superset: fixes S :: "'a::euclidean_space set" assumes "convex S" "interior S \ {}" "interior S \ closure T" shows "closure(S \ T) = closure S" proof - have "closure S \ closure(interior S)" by (simp add: convex_closure_interior assms) also have "... \ closure (S \ T)" using interior_subset [of S] assms by (metis (no_types, lifting) Int_assoc Int_lower2 closure_mono closure_open_Int_superset inf.orderE open_interior) finally show ?thesis by (simp add: closure_mono dual_order.antisym) qed subsection\<^marker>\tag unimportant\ \Some obvious but surprisingly hard simplex lemmas\ lemma simplex: assumes "finite S" and "0 \ S" shows "convex hull (insert 0 S) = {y. \u. (\x\S. 0 \ u x) \ sum u S \ 1 \ sum (\x. u x *\<^sub>R x) S = y}" proof (simp add: convex_hull_finite set_eq_iff assms, safe) fix x and u :: "'a \ real" assume "0 \ u 0" "\x\S. 0 \ u x" "u 0 + sum u S = 1" then show "\v. (\x\S. 0 \ v x) \ sum v S \ 1 \ (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" by force next fix x and u :: "'a \ real" assume "\x\S. 0 \ u x" "sum u S \ 1" then show "\v. 0 \ v 0 \ (\x\S. 0 \ v x) \ v 0 + sum v S = 1 \ (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" by (rule_tac x="\x. if x = 0 then 1 - sum u S else u x" in exI) (auto simp: sum_delta_notmem assms if_smult) qed lemma substd_simplex: assumes d: "d \ Basis" shows "convex hull (insert 0 d) = {x. (\i\Basis. 0 \ x\i) \ (\i\d. x\i) \ 1 \ (\i\Basis. i \ d \ x\i = 0)}" (is "convex hull (insert 0 ?p) = ?s") proof - let ?D = d have "0 \ ?p" using assms by (auto simp: image_def) from d have "finite d" by (blast intro: finite_subset finite_Basis) show ?thesis unfolding simplex[OF \finite d\ \0 \ ?p\] proof (intro set_eqI; safe) fix u :: "'a \ real" assume as: "\x\?D. 0 \ u x" "sum u ?D \ 1" let ?x = "(\x\?D. u x *\<^sub>R x)" have ind: "\i\Basis. i \ d \ u i = ?x \ i" and notind: "(\i\Basis. i \ d \ ?x \ i = 0)" using substdbasis_expansion_unique[OF assms] by blast+ then have **: "sum u ?D = sum ((\) ?x) ?D" using assms by (auto intro!: sum.cong) show "0 \ ?x \ i" if "i \ Basis" for i using as(1) ind notind that by fastforce show "sum ((\) ?x) ?D \ 1" using "**" as(2) by linarith show "?x \ i = 0" if "i \ Basis" "i \ d" for i using notind that by blast next fix x assume "\i\Basis. 0 \ x \ i" "sum ((\) x) ?D \ 1" "(\i\Basis. i \ d \ x \ i = 0)" with d show "\u. (\x\?D. 0 \ u x) \ sum u ?D \ 1 \ (\x\?D. u x *\<^sub>R x) = x" unfolding substdbasis_expansion_unique[OF assms] by (rule_tac x="inner x" in exI) auto qed qed lemma std_simplex: "convex hull (insert 0 Basis) = {x::'a::euclidean_space. (\i\Basis. 0 \ x\i) \ sum (\i. x\i) Basis \ 1}" using substd_simplex[of Basis] by auto lemma interior_std_simplex: "interior (convex hull (insert 0 Basis)) = {x::'a::euclidean_space. (\i\Basis. 0 < x\i) \ sum (\i. x\i) Basis < 1}" unfolding set_eq_iff mem_interior std_simplex proof (intro allI iffI CollectI; clarify) fix x :: 'a fix e assume "e > 0" and as: "ball x e \ {x. (\i\Basis. 0 \ x \ i) \ sum ((\) x) Basis \ 1}" show "(\i\Basis. 0 < x \ i) \ sum ((\) x) Basis < 1" proof safe fix i :: 'a assume i: "i \ Basis" then show "0 < x \ i" using as[THEN subsetD[where c="x - (e / 2) *\<^sub>R i"]] and \e > 0\ by (force simp add: inner_simps) next have **: "dist x (x + (e / 2) *\<^sub>R (SOME i. i\Basis)) < e" using \e > 0\ unfolding dist_norm by (auto intro!: mult_strict_left_mono simp: SOME_Basis) have "\i. i \ Basis \ (x + (e / 2) *\<^sub>R (SOME i. i\Basis)) \ i = x\i + (if i = (SOME i. i\Basis) then e/2 else 0)" by (auto simp: SOME_Basis inner_Basis inner_simps) then have *: "sum ((\) (x + (e / 2) *\<^sub>R (SOME i. i\Basis))) Basis = sum (\i. x\i + (if (SOME i. i\Basis) = i then e/2 else 0)) Basis" by (auto simp: intro!: sum.cong) have "sum ((\) x) Basis < sum ((\) (x + (e / 2) *\<^sub>R (SOME i. i\Basis))) Basis" using \e > 0\ DIM_positive by (auto simp: SOME_Basis sum.distrib *) also have "\ \ 1" using ** as by force finally show "sum ((\) x) Basis < 1" by auto qed next fix x :: 'a assume as: "\i\Basis. 0 < x \ i" "sum ((\) x) Basis < 1" obtain a :: 'b where "a \ UNIV" using UNIV_witness .. let ?d = "(1 - sum ((\) x) Basis) / real (DIM('a))" show "\e>0. ball x e \ {x. (\i\Basis. 0 \ x \ i) \ sum ((\) x) Basis \ 1}" proof (rule_tac x="min (Min (((\) x) ` Basis)) D" for D in exI, intro conjI subsetI CollectI) fix y assume y: "y \ ball x (min (Min ((\) x ` Basis)) ?d)" have "sum ((\) y) Basis \ sum (\i. x\i + ?d) Basis" proof (rule sum_mono) fix i :: 'a assume i: "i \ Basis" have "\y\i - x\i\ \ norm (y - x)" by (metis Basis_le_norm i inner_commute inner_diff_right) also have "... < ?d" using y by (simp add: dist_norm norm_minus_commute) finally have "\y\i - x\i\ < ?d" . then show "y \ i \ x \ i + ?d" by auto qed also have "\ \ 1" unfolding sum.distrib sum_constant by (auto simp add: Suc_le_eq) finally show "sum ((\) y) Basis \ 1" . show "(\i\Basis. 0 \ y \ i)" proof safe fix i :: 'a assume i: "i \ Basis" have "norm (x - y) < Min (((\) x) ` Basis)" using y by (auto simp: dist_norm less_eq_real_def) also have "... \ x\i" using i by auto finally have "norm (x - y) < x\i" . then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format, OF i] by (auto simp: inner_simps) qed next have "Min (((\) x) ` Basis) > 0" using as by simp moreover have "?d > 0" using as by (auto simp: Suc_le_eq) ultimately show "0 < min (Min ((\) x ` Basis)) ((1 - sum ((\) x) Basis) / real DIM('a))" by linarith qed qed lemma interior_std_simplex_nonempty: obtains a :: "'a::euclidean_space" where "a \ interior(convex hull (insert 0 Basis))" proof - let ?D = "Basis :: 'a set" let ?a = "sum (\b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis" { fix i :: 'a assume i: "i \ Basis" have "?a \ i = inverse (2 * real DIM('a))" by (rule trans[of _ "sum (\j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"]) (simp_all add: sum.If_cases i) } note ** = this show ?thesis apply (rule that[of ?a]) unfolding interior_std_simplex mem_Collect_eq proof safe fix i :: 'a assume i: "i \ Basis" show "0 < ?a \ i" unfolding **[OF i] by (auto simp add: Suc_le_eq) next have "sum ((\) ?a) ?D = sum (\i. inverse (2 * real DIM('a))) ?D" apply (rule sum.cong) apply rule apply auto done also have "\ < 1" unfolding sum_constant divide_inverse[symmetric] by (auto simp add: field_simps) finally show "sum ((\) ?a) ?D < 1" by auto qed qed lemma rel_interior_substd_simplex: assumes D: "D \ Basis" shows "rel_interior (convex hull (insert 0 D)) = {x::'a::euclidean_space. (\i\D. 0 < x\i) \ (\i\D. x\i) < 1 \ (\i\Basis. i \ D \ x\i = 0)}" (is "rel_interior (convex hull (insert 0 ?p)) = ?s") proof - have "finite D" using D finite_Basis finite_subset by blast show ?thesis proof (cases "D = {}") case True then show ?thesis using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto next case False have h0: "affine hull (convex hull (insert 0 ?p)) = {x::'a::euclidean_space. (\i\Basis. i \ D \ x\i = 0)}" using affine_hull_convex_hull affine_hull_substd_basis assms by auto have aux: "\x::'a. \i\Basis. (\i\D. 0 \ x\i) \ (\i\Basis. i \ D \ x\i = 0) \ 0 \ x\i" by auto { fix x :: "'a::euclidean_space" assume x: "x \ rel_interior (convex hull (insert 0 ?p))" then obtain e where "e > 0" and "ball x e \ {xa. (\i\Basis. i \ D \ xa\i = 0)} \ convex hull (insert 0 ?p)" using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto then have as [rule_format]: "\y. dist x y < e \ (\i\Basis. i \ D \ y\i = 0) \ (\i\D. 0 \ y \ i) \ sum ((\) y) D \ 1" unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto have x0: "(\i\Basis. i \ D \ x\i = 0)" using x rel_interior_subset substd_simplex[OF assms] by auto have "(\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x\i = 0)" proof (intro conjI ballI) fix i :: 'a assume "i \ D" then have "\j\D. 0 \ (x - (e / 2) *\<^sub>R i) \ j" apply - apply (rule as[THEN conjunct1]) using D \e > 0\ x0 apply (auto simp: dist_norm inner_simps inner_Basis) done then show "0 < x \ i" using \e > 0\ \i \ D\ D by (force simp: inner_simps inner_Basis) next obtain a where a: "a \ D" using \D \ {}\ by auto then have **: "dist x (x + (e / 2) *\<^sub>R a) < e" using \e > 0\ norm_Basis[of a] D unfolding dist_norm by auto have "\i. i \ Basis \ (x + (e / 2) *\<^sub>R a) \ i = x\i + (if i = a then e/2 else 0)" using a D by (auto simp: inner_simps inner_Basis) then have *: "sum ((\) (x + (e / 2) *\<^sub>R a)) D = sum (\i. x\i + (if a = i then e/2 else 0)) D" using D by (intro sum.cong) auto have "a \ Basis" using \a \ D\ D by auto then have h1: "(\i\Basis. i \ D \ (x + (e / 2) *\<^sub>R a) \ i = 0)" using x0 D \a\D\ by (auto simp add: inner_add_left inner_Basis) have "sum ((\) x) D < sum ((\) (x + (e / 2) *\<^sub>R a)) D" using \e > 0\ \a \ D\ \finite D\ by (auto simp add: * sum.distrib) also have "\ \ 1" using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"] by auto finally show "sum ((\) x) D < 1" "\i. i\Basis \ i \ D \ x\i = 0" using x0 by auto qed } moreover { fix x :: "'a::euclidean_space" assume as: "x \ ?s" have "\i. 0 < x\i \ 0 = x\i \ 0 \ x\i" by auto moreover have "\i. i \ D \ i \ D" by auto ultimately have "\i. (\i\D. 0 < x\i) \ (\i. i \ D \ x\i = 0) \ 0 \ x\i" by metis then have h2: "x \ convex hull (insert 0 ?p)" using as assms unfolding substd_simplex[OF assms] by fastforce obtain a where a: "a \ D" using \D \ {}\ by auto let ?d = "(1 - sum ((\) x) D) / real (card D)" have "0 < card D" using \D \ {}\ \finite D\ by (simp add: card_gt_0_iff) have "Min (((\) x) ` D) > 0" using as \D \ {}\ \finite D\ by (simp) moreover have "?d > 0" using as using \0 < card D\ by auto ultimately have h3: "min (Min (((\) x) ` D)) ?d > 0" by auto have "x \ rel_interior (convex hull (insert 0 ?p))" unfolding rel_interior_ball mem_Collect_eq h0 apply (rule,rule h2) unfolding substd_simplex[OF assms] apply (rule_tac x="min (Min (((\) x) ` D)) ?d" in exI) apply (rule, rule h3) apply safe unfolding mem_ball proof - fix y :: 'a assume y: "dist x y < min (Min ((\) x ` D)) ?d" assume y2: "\i\Basis. i \ D \ y\i = 0" have "sum ((\) y) D \ sum (\i. x\i + ?d) D" proof (rule sum_mono) fix i assume "i \ D" with D have i: "i \ Basis" by auto have "\y\i - x\i\ \ norm (y - x)" by (metis i inner_commute inner_diff_right norm_bound_Basis_le order_refl) also have "... < ?d" by (metis dist_norm min_less_iff_conj norm_minus_commute y) finally have "\y\i - x\i\ < ?d" . then show "y \ i \ x \ i + ?d" by auto qed also have "\ \ 1" unfolding sum.distrib sum_constant using \0 < card D\ by auto finally show "sum ((\) y) D \ 1" . fix i :: 'a assume i: "i \ Basis" then show "0 \ y\i" proof (cases "i\D") case True have "norm (x - y) < x\i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1] using Min_gr_iff[of "(\) x ` D" "norm (x - y)"] \0 < card D\ \i \ D\ by (simp add: card_gt_0_iff) then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format] by (auto simp: inner_simps) qed (insert y2, auto) qed } ultimately have "\x. x \ rel_interior (convex hull insert 0 D) \ x \ {x. (\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x \ i = 0)}" by blast then show ?thesis by (rule set_eqI) qed qed lemma rel_interior_substd_simplex_nonempty: assumes "D \ {}" and "D \ Basis" obtains a :: "'a::euclidean_space" where "a \ rel_interior (convex hull (insert 0 D))" proof - let ?D = D let ?a = "sum (\b::'a::euclidean_space. inverse (2 * real (card D)) *\<^sub>R b) ?D" have "finite D" apply (rule finite_subset) using assms(2) apply auto done then have d1: "0 < real (card D)" using \D \ {}\ by auto { fix i assume "i \ D" have "?a \ i = inverse (2 * real (card D))" apply (rule trans[of _ "sum (\j. if i = j then inverse (2 * real (card D)) else 0) ?D"]) unfolding inner_sum_left apply (rule sum.cong) using \i \ D\ \finite D\ sum.delta'[of D i "(\k. inverse (2 * real (card D)))"] d1 assms(2) by (auto simp: inner_Basis rev_subsetD[OF _ assms(2)]) } note ** = this show ?thesis apply (rule that[of ?a]) unfolding rel_interior_substd_simplex[OF assms(2)] mem_Collect_eq proof safe fix i assume "i \ D" have "0 < inverse (2 * real (card D))" using d1 by auto also have "\ = ?a \ i" using **[of i] \i \ D\ by auto finally show "0 < ?a \ i" by auto next have "sum ((\) ?a) ?D = sum (\i. inverse (2 * real (card D))) ?D" by (rule sum.cong) (rule refl, rule **) also have "\ < 1" unfolding sum_constant divide_real_def[symmetric] by (auto simp add: field_simps) finally show "sum ((\) ?a) ?D < 1" by auto next fix i assume "i \ Basis" and "i \ D" have "?a \ span D" proof (rule span_sum[of D "(\b. b /\<^sub>R (2 * real (card D)))" D]) { fix x :: "'a::euclidean_space" assume "x \ D" then have "x \ span D" using span_base[of _ "D"] by auto then have "x /\<^sub>R (2 * real (card D)) \ span D" using span_mul[of x "D" "(inverse (real (card D)) / 2)"] by auto } then show "\x. x\D \ x /\<^sub>R (2 * real (card D)) \ span D" by auto qed then show "?a \ i = 0 " using \i \ D\ unfolding span_substd_basis[OF assms(2)] using \i \ Basis\ by auto qed qed subsection\<^marker>\tag unimportant\ \Relative interior of convex set\ lemma rel_interior_convex_nonempty_aux: fixes S :: "'n::euclidean_space set" assumes "convex S" and "0 \ S" shows "rel_interior S \ {}" proof (cases "S = {0}") case True then show ?thesis using rel_interior_sing by auto next case False obtain B where B: "independent B \ B \ S \ S \ span B \ card B = dim S" using basis_exists[of S] by metis then have "B \ {}" using B assms \S \ {0}\ span_empty by auto have "insert 0 B \ span B" using subspace_span[of B] subspace_0[of "span B"] span_superset by auto then have "span (insert 0 B) \ span B" using span_span[of B] span_mono[of "insert 0 B" "span B"] by blast then have "convex hull insert 0 B \ span B" using convex_hull_subset_span[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) \ span B" using span_span[of B] span_mono[of "convex hull insert 0 B" "span B"] by blast then have *: "span (convex hull insert 0 B) = span B" using span_mono[of B "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) = span S" using B span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto moreover have "0 \ affine hull (convex hull insert 0 B)" using hull_subset[of "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S" using affine_hull_span_0[of "convex hull insert 0 B"] affine_hull_span_0[of "S"] assms hull_subset[of S] by auto obtain d and f :: "'n \ 'n" where fd: "card d = card B" "linear f" "f ` B = d" "f ` span B = {x. \i\Basis. i \ d \ x \ i = (0::real)} \ inj_on f (span B)" and d: "d \ Basis" using basis_to_substdbasis_subspace_isomorphism[of B,OF _ ] B by auto then have "bounded_linear f" using linear_conv_bounded_linear by auto have "d \ {}" using fd B \B \ {}\ by auto have "insert 0 d = f ` (insert 0 B)" using fd linear_0 by auto then have "(convex hull (insert 0 d)) = f ` (convex hull (insert 0 B))" using convex_hull_linear_image[of f "(insert 0 d)"] convex_hull_linear_image[of f "(insert 0 B)"] \linear f\ by auto moreover have "rel_interior (f ` (convex hull insert 0 B)) = f ` rel_interior (convex hull insert 0 B)" apply (rule rel_interior_injective_on_span_linear_image[of f "(convex hull insert 0 B)"]) using \bounded_linear f\ fd * apply auto done ultimately have "rel_interior (convex hull insert 0 B) \ {}" using rel_interior_substd_simplex_nonempty[OF \d \ {}\ d] apply auto apply blast done moreover have "convex hull (insert 0 B) \ S" using B assms hull_mono[of "insert 0 B" "S" "convex"] convex_hull_eq by auto ultimately show ?thesis using subset_rel_interior[of "convex hull insert 0 B" S] ** by auto qed lemma rel_interior_eq_empty: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior S = {} \ S = {}" proof - { assume "S \ {}" then obtain a where "a \ S" by auto then have "0 \ (+) (-a) ` S" using assms exI[of "(\x. x \ S \ - a + x = 0)" a] by auto then have "rel_interior ((+) (-a) ` S) \ {}" using rel_interior_convex_nonempty_aux[of "(+) (-a) ` S"] convex_translation[of S "-a"] assms by auto then have "rel_interior S \ {}" using rel_interior_translation [of "- a"] by simp } then show ?thesis by auto qed lemma interior_simplex_nonempty: fixes S :: "'N :: euclidean_space set" assumes "independent S" "finite S" "card S = DIM('N)" obtains a where "a \ interior (convex hull (insert 0 S))" proof - have "affine hull (insert 0 S) = UNIV" by (simp add: hull_inc affine_hull_span_0 dim_eq_full[symmetric] assms(1) assms(3) dim_eq_card_independent) moreover have "rel_interior (convex hull insert 0 S) \ {}" using rel_interior_eq_empty [of "convex hull (insert 0 S)"] by auto ultimately have "interior (convex hull insert 0 S) \ {}" by (simp add: rel_interior_interior) with that show ?thesis by auto qed lemma convex_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "convex (rel_interior S)" proof - { fix x y and u :: real assume assm: "x \ rel_interior S" "y \ rel_interior S" "0 \ u" "u \ 1" then have "x \ S" using rel_interior_subset by auto have "x - u *\<^sub>R (x-y) \ rel_interior S" proof (cases "0 = u") case False then have "0 < u" using assm by auto then show ?thesis using assm rel_interior_convex_shrink[of S y x u] assms \x \ S\ by auto next case True then show ?thesis using assm by auto qed then have "(1 - u) *\<^sub>R x + u *\<^sub>R y \ rel_interior S" by (simp add: algebra_simps) } then show ?thesis unfolding convex_alt by auto qed lemma convex_closure_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "closure (rel_interior S) = closure S" proof - have h1: "closure (rel_interior S) \ closure S" using closure_mono[of "rel_interior S" S] rel_interior_subset[of S] by auto show ?thesis proof (cases "S = {}") case False then obtain a where a: "a \ rel_interior S" using rel_interior_eq_empty assms by auto { fix x assume x: "x \ closure S" { assume "x = a" then have "x \ closure (rel_interior S)" using a unfolding closure_def by auto } moreover { assume "x \ a" { fix e :: real assume "e > 0" define e1 where "e1 = min 1 (e/norm (x - a))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (x - a) \ e" using \x \ a\ \e > 0\ le_divide_eq[of e1 e "norm (x - a)"] by simp_all then have *: "x - e1 *\<^sub>R (x - a) \ rel_interior S" using rel_interior_closure_convex_shrink[of S a x e1] assms x a e1_def by auto have "\y. y \ rel_interior S \ y \ x \ dist y x \ e" apply (rule_tac x="x - e1 *\<^sub>R (x - a)" in exI) using * e1 dist_norm[of "x - e1 *\<^sub>R (x - a)" x] \x \ a\ apply simp done } then have "x islimpt rel_interior S" unfolding islimpt_approachable_le by auto then have "x \ closure(rel_interior S)" unfolding closure_def by auto } ultimately have "x \ closure(rel_interior S)" by auto } then show ?thesis using h1 by auto next case True then have "rel_interior S = {}" by auto then have "closure (rel_interior S) = {}" using closure_empty by auto with True show ?thesis by auto qed qed lemma rel_interior_same_affine_hull: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "affine hull (rel_interior S) = affine hull S" by (metis assms closure_same_affine_hull convex_closure_rel_interior) lemma rel_interior_aff_dim: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "aff_dim (rel_interior S) = aff_dim S" by (metis aff_dim_affine_hull2 assms rel_interior_same_affine_hull) lemma rel_interior_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior (rel_interior S) = rel_interior S" proof - have "openin (top_of_set (affine hull (rel_interior S))) (rel_interior S)" using openin_rel_interior[of S] rel_interior_same_affine_hull[of S] assms by auto then show ?thesis using rel_interior_def by auto qed lemma rel_interior_rel_open: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_open (rel_interior S)" unfolding rel_open_def using rel_interior_rel_interior assms by auto lemma convex_rel_interior_closure_aux: fixes x y z :: "'n::euclidean_space" assumes "0 < a" "0 < b" "(a + b) *\<^sub>R z = a *\<^sub>R x + b *\<^sub>R y" obtains e where "0 < e" "e \ 1" "z = y - e *\<^sub>R (y - x)" proof - define e where "e = a / (a + b)" have "z = (1 / (a + b)) *\<^sub>R ((a + b) *\<^sub>R z)" using assms by (simp add: eq_vector_fraction_iff) also have "\ = (1 / (a + b)) *\<^sub>R (a *\<^sub>R x + b *\<^sub>R y)" using assms scaleR_cancel_left[of "1/(a+b)" "(a + b) *\<^sub>R z" "a *\<^sub>R x + b *\<^sub>R y"] by auto also have "\ = y - e *\<^sub>R (y-x)" using e_def apply (simp add: algebra_simps) using scaleR_left_distrib[of "a/(a+b)" "b/(a+b)" y] assms add_divide_distrib[of a b "a+b"] apply auto done finally have "z = y - e *\<^sub>R (y-x)" by auto moreover have "e > 0" using e_def assms by auto moreover have "e \ 1" using e_def assms by auto ultimately show ?thesis using that[of e] by auto qed lemma convex_rel_interior_closure: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior (closure S) = rel_interior S" proof (cases "S = {}") case True then show ?thesis using assms rel_interior_eq_empty by auto next case False have "rel_interior (closure S) \ rel_interior S" using subset_rel_interior[of S "closure S"] closure_same_affine_hull closure_subset by auto moreover { fix z assume z: "z \ rel_interior (closure S)" obtain x where x: "x \ rel_interior S" using \S \ {}\ assms rel_interior_eq_empty by auto have "z \ rel_interior S" proof (cases "x = z") case True then show ?thesis using x by auto next case False obtain e where e: "e > 0" "cball z e \ affine hull closure S \ closure S" using z rel_interior_cball[of "closure S"] by auto hence *: "0 < e/norm(z-x)" using e False by auto define y where "y = z + (e/norm(z-x)) *\<^sub>R (z-x)" have yball: "y \ cball z e" using y_def dist_norm[of z y] e by auto have "x \ affine hull closure S" using x rel_interior_subset_closure hull_inc[of x "closure S"] by blast moreover have "z \ affine hull closure S" using z rel_interior_subset hull_subset[of "closure S"] by blast ultimately have "y \ affine hull closure S" using y_def affine_affine_hull[of "closure S"] mem_affine_3_minus [of "affine hull closure S" z z x "e/norm(z-x)"] by auto then have "y \ closure S" using e yball by auto have "(1 + (e/norm(z-x))) *\<^sub>R z = (e/norm(z-x)) *\<^sub>R x + y" using y_def by (simp add: algebra_simps) then obtain e1 where "0 < e1" "e1 \ 1" "z = y - e1 *\<^sub>R (y - x)" using * convex_rel_interior_closure_aux[of "e / norm (z - x)" 1 z x y] by (auto simp add: algebra_simps) then show ?thesis using rel_interior_closure_convex_shrink assms x \y \ closure S\ by auto qed } ultimately show ?thesis by auto qed lemma convex_interior_closure: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "interior (closure S) = interior S" using closure_aff_dim[of S] interior_rel_interior_gen[of S] interior_rel_interior_gen[of "closure S"] convex_rel_interior_closure[of S] assms by auto lemma closure_eq_rel_interior_eq: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" shows "closure S1 = closure S2 \ rel_interior S1 = rel_interior S2" by (metis convex_rel_interior_closure convex_closure_rel_interior assms) lemma closure_eq_between: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" shows "closure S1 = closure S2 \ rel_interior S1 \ S2 \ S2 \ closure S1" (is "?A \ ?B") proof assume ?A then show ?B by (metis assms closure_subset convex_rel_interior_closure rel_interior_subset) next assume ?B then have "closure S1 \ closure S2" by (metis assms(1) convex_closure_rel_interior closure_mono) moreover from \?B\ have "closure S1 \ closure S2" by (metis closed_closure closure_minimal) ultimately show ?A .. qed lemma open_inter_closure_rel_interior: fixes S A :: "'n::euclidean_space set" assumes "convex S" and "open A" shows "A \ closure S = {} \ A \ rel_interior S = {}" by (metis assms convex_closure_rel_interior open_Int_closure_eq_empty) lemma rel_interior_open_segment: fixes a :: "'a :: euclidean_space" shows "rel_interior(open_segment a b) = open_segment a b" proof (cases "a = b") case True then show ?thesis by auto next case False then show ?thesis apply (simp add: rel_interior_eq openin_open) apply (rule_tac x="ball (inverse 2 *\<^sub>R (a + b)) (norm(b - a) / 2)" in exI) apply (simp add: open_segment_as_ball) done qed lemma rel_interior_closed_segment: fixes a :: "'a :: euclidean_space" shows "rel_interior(closed_segment a b) = (if a = b then {a} else open_segment a b)" proof (cases "a = b") case True then show ?thesis by auto next case False then show ?thesis by simp (metis closure_open_segment convex_open_segment convex_rel_interior_closure rel_interior_open_segment) qed lemmas rel_interior_segment = rel_interior_closed_segment rel_interior_open_segment -lemma starlike_convex_tweak_boundary_points: - fixes S :: "'a::euclidean_space set" - assumes "convex S" "S \ {}" and ST: "rel_interior S \ T" and TS: "T \ closure S" - shows "starlike T" -proof - - have "rel_interior S \ {}" - by (simp add: assms rel_interior_eq_empty) - then obtain a where a: "a \ rel_interior S" by blast - with ST have "a \ T" by blast - have *: "\x. x \ T \ open_segment a x \ rel_interior S" - apply (rule rel_interior_closure_convex_segment [OF \convex S\ a]) - using assms by blast - show ?thesis - unfolding starlike_def - apply (rule bexI [OF _ \a \ T\]) - apply (simp add: closed_segment_eq_open) - apply (intro conjI ballI a \a \ T\ rel_interior_closure_convex_segment [OF \convex S\ a]) - apply (simp add: order_trans [OF * ST]) - done -qed - subsection\The relative frontier of a set\ definition\<^marker>\tag important\ "rel_frontier S = closure S - rel_interior S" lemma rel_frontier_empty [simp]: "rel_frontier {} = {}" by (simp add: rel_frontier_def) lemma rel_frontier_eq_empty: fixes S :: "'n::euclidean_space set" shows "rel_frontier S = {} \ affine S" unfolding rel_frontier_def using rel_interior_subset_closure by (auto simp add: rel_interior_eq_closure [symmetric]) lemma rel_frontier_sing [simp]: fixes a :: "'n::euclidean_space" shows "rel_frontier {a} = {}" by (simp add: rel_frontier_def) lemma rel_frontier_affine_hull: fixes S :: "'a::euclidean_space set" shows "rel_frontier S \ affine hull S" using closure_affine_hull rel_frontier_def by fastforce lemma rel_frontier_cball [simp]: fixes a :: "'n::euclidean_space" shows "rel_frontier(cball a r) = (if r = 0 then {} else sphere a r)" proof (cases rule: linorder_cases [of r 0]) case less then show ?thesis by (force simp: sphere_def) next case equal then show ?thesis by simp next case greater then show ?thesis apply simp by (metis centre_in_ball empty_iff frontier_cball frontier_def interior_cball interior_rel_interior_gen rel_frontier_def) qed lemma rel_frontier_translation: fixes a :: "'a::euclidean_space" shows "rel_frontier((\x. a + x) ` S) = (\x. a + x) ` (rel_frontier S)" by (simp add: rel_frontier_def translation_diff rel_interior_translation closure_translation) lemma rel_frontier_nonempty_interior: fixes S :: "'n::euclidean_space set" shows "interior S \ {} \ rel_frontier S = frontier S" by (metis frontier_def interior_rel_interior_gen rel_frontier_def) lemma rel_frontier_frontier: fixes S :: "'n::euclidean_space set" shows "affine hull S = UNIV \ rel_frontier S = frontier S" by (simp add: frontier_def rel_frontier_def rel_interior_interior) lemma closest_point_in_rel_frontier: "\closed S; S \ {}; x \ affine hull S - rel_interior S\ \ closest_point S x \ rel_frontier S" by (simp add: closest_point_in_rel_interior closest_point_in_set rel_frontier_def) lemma closed_rel_frontier [iff]: fixes S :: "'n::euclidean_space set" shows "closed (rel_frontier S)" proof - have *: "closedin (top_of_set (affine hull S)) (closure S - rel_interior S)" by (simp add: closed_subset closedin_diff closure_affine_hull openin_rel_interior) show ?thesis apply (rule closedin_closed_trans[of "affine hull S" "rel_frontier S"]) unfolding rel_frontier_def using * closed_affine_hull apply auto done qed lemma closed_rel_boundary: fixes S :: "'n::euclidean_space set" shows "closed S \ closed(S - rel_interior S)" by (metis closed_rel_frontier closure_closed rel_frontier_def) lemma compact_rel_boundary: fixes S :: "'n::euclidean_space set" shows "compact S \ compact(S - rel_interior S)" by (metis bounded_diff closed_rel_boundary closure_eq compact_closure compact_imp_closed) lemma bounded_rel_frontier: fixes S :: "'n::euclidean_space set" shows "bounded S \ bounded(rel_frontier S)" by (simp add: bounded_closure bounded_diff rel_frontier_def) lemma compact_rel_frontier_bounded: fixes S :: "'n::euclidean_space set" shows "bounded S \ compact(rel_frontier S)" using bounded_rel_frontier closed_rel_frontier compact_eq_bounded_closed by blast lemma compact_rel_frontier: fixes S :: "'n::euclidean_space set" shows "compact S \ compact(rel_frontier S)" by (meson compact_eq_bounded_closed compact_rel_frontier_bounded) lemma convex_same_rel_interior_closure: fixes S :: "'n::euclidean_space set" shows "\convex S; convex T\ \ rel_interior S = rel_interior T \ closure S = closure T" by (simp add: closure_eq_rel_interior_eq) lemma convex_same_rel_interior_closure_straddle: fixes S :: "'n::euclidean_space set" shows "\convex S; convex T\ \ rel_interior S = rel_interior T \ rel_interior S \ T \ T \ closure S" by (simp add: closure_eq_between convex_same_rel_interior_closure) lemma convex_rel_frontier_aff_dim: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" and "S2 \ {}" and "S1 \ rel_frontier S2" shows "aff_dim S1 < aff_dim S2" proof - have "S1 \ closure S2" using assms unfolding rel_frontier_def by auto then have *: "affine hull S1 \ affine hull S2" using hull_mono[of "S1" "closure S2"] closure_same_affine_hull[of S2] by blast then have "aff_dim S1 \ aff_dim S2" using * aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] aff_dim_subset[of "affine hull S1" "affine hull S2"] by auto moreover { assume eq: "aff_dim S1 = aff_dim S2" then have "S1 \ {}" using aff_dim_empty[of S1] aff_dim_empty[of S2] \S2 \ {}\ by auto have **: "affine hull S1 = affine hull S2" apply (rule affine_dim_equal) using * affine_affine_hull apply auto using \S1 \ {}\ hull_subset[of S1] apply auto using eq aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] apply auto done obtain a where a: "a \ rel_interior S1" using \S1 \ {}\ rel_interior_eq_empty assms by auto obtain T where T: "open T" "a \ T \ S1" "T \ affine hull S1 \ S1" using mem_rel_interior[of a S1] a by auto then have "a \ T \ closure S2" using a assms unfolding rel_frontier_def by auto then obtain b where b: "b \ T \ rel_interior S2" using open_inter_closure_rel_interior[of S2 T] assms T by auto then have "b \ affine hull S1" using rel_interior_subset hull_subset[of S2] ** by auto then have "b \ S1" using T b by auto then have False using b assms unfolding rel_frontier_def by auto } ultimately show ?thesis using less_le by auto qed lemma convex_rel_interior_if: fixes S :: "'n::euclidean_space set" assumes "convex S" and "z \ rel_interior S" shows "\x\affine hull S. \m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" proof - obtain e1 where e1: "e1 > 0 \ cball z e1 \ affine hull S \ S" using mem_rel_interior_cball[of z S] assms by auto { fix x assume x: "x \ affine hull S" { assume "x \ z" define m where "m = 1 + e1/norm(x-z)" hence "m > 1" using e1 \x \ z\ by auto { fix e assume e: "e > 1 \ e \ m" have "z \ affine hull S" using assms rel_interior_subset hull_subset[of S] by auto then have *: "(1 - e)*\<^sub>R x + e *\<^sub>R z \ affine hull S" using mem_affine[of "affine hull S" x z "(1-e)" e] affine_affine_hull[of S] x by auto have "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) = norm ((e - 1) *\<^sub>R (x - z))" by (simp add: algebra_simps) also have "\ = (e - 1) * norm (x-z)" using norm_scaleR e by auto also have "\ \ (m - 1) * norm (x - z)" using e mult_right_mono[of _ _ "norm(x-z)"] by auto also have "\ = (e1 / norm (x - z)) * norm (x - z)" using m_def by auto also have "\ = e1" using \x \ z\ e1 by simp finally have **: "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) \ e1" by auto have "(1 - e)*\<^sub>R x+ e *\<^sub>R z \ cball z e1" using m_def ** unfolding cball_def dist_norm by (auto simp add: algebra_simps) then have "(1 - e) *\<^sub>R x+ e *\<^sub>R z \ S" using e * e1 by auto } then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S )" using \m> 1 \ by auto } moreover { assume "x = z" define m where "m = 1 + e1" then have "m > 1" using e1 by auto { fix e assume e: "e > 1 \ e \ m" then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using e1 x \x = z\ by (auto simp add: algebra_simps) then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using e by auto } then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using \m > 1\ by auto } ultimately have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S )" by blast } then show ?thesis by auto qed lemma convex_rel_interior_if2: fixes S :: "'n::euclidean_space set" assumes "convex S" assumes "z \ rel_interior S" shows "\x\affine hull S. \e. e > 1 \ (1 - e)*\<^sub>R x + e *\<^sub>R z \ S" using convex_rel_interior_if[of S z] assms by auto lemma convex_rel_interior_only_if: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" assumes "\x\S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" shows "z \ rel_interior S" proof - obtain x where x: "x \ rel_interior S" using rel_interior_eq_empty assms by auto then have "x \ S" using rel_interior_subset by auto then obtain e where e: "e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using assms by auto define y where [abs_def]: "y = (1 - e) *\<^sub>R x + e *\<^sub>R z" then have "y \ S" using e by auto define e1 where "e1 = 1/e" then have "0 < e1 \ e1 < 1" using e by auto then have "z =y - (1 - e1) *\<^sub>R (y - x)" using e1_def y_def by (auto simp add: algebra_simps) then show ?thesis using rel_interior_convex_shrink[of S x y "1-e1"] \0 < e1 \ e1 < 1\ \y \ S\ x assms by auto qed lemma convex_rel_interior_iff: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" shows "z \ rel_interior S \ (\x\S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using assms hull_subset[of S "affine"] convex_rel_interior_if[of S z] convex_rel_interior_only_if[of S z] by auto lemma convex_rel_interior_iff2: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" shows "z \ rel_interior S \ (\x\affine hull S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using assms hull_subset[of S] convex_rel_interior_if2[of S z] convex_rel_interior_only_if[of S z] by auto lemma convex_interior_iff: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "z \ interior S \ (\x. \e. e > 0 \ z + e *\<^sub>R x \ S)" proof (cases "aff_dim S = int DIM('n)") case False { assume "z \ interior S" then have False using False interior_rel_interior_gen[of S] by auto } moreover { assume r: "\x. \e. e > 0 \ z + e *\<^sub>R x \ S" { fix x obtain e1 where e1: "e1 > 0 \ z + e1 *\<^sub>R (x - z) \ S" using r by auto obtain e2 where e2: "e2 > 0 \ z + e2 *\<^sub>R (z - x) \ S" using r by auto define x1 where [abs_def]: "x1 = z + e1 *\<^sub>R (x - z)" then have x1: "x1 \ affine hull S" using e1 hull_subset[of S] by auto define x2 where [abs_def]: "x2 = z + e2 *\<^sub>R (z - x)" then have x2: "x2 \ affine hull S" using e2 hull_subset[of S] by auto have *: "e1/(e1+e2) + e2/(e1+e2) = 1" using add_divide_distrib[of e1 e2 "e1+e2"] e1 e2 by simp then have "z = (e2/(e1+e2)) *\<^sub>R x1 + (e1/(e1+e2)) *\<^sub>R x2" using x1_def x2_def apply (auto simp add: algebra_simps) using scaleR_left_distrib[of "e1/(e1+e2)" "e2/(e1+e2)" z] apply auto done then have z: "z \ affine hull S" using mem_affine[of "affine hull S" x1 x2 "e2/(e1+e2)" "e1/(e1+e2)"] x1 x2 affine_affine_hull[of S] * by auto have "x1 - x2 = (e1 + e2) *\<^sub>R (x - z)" using x1_def x2_def by (auto simp add: algebra_simps) then have "x = z+(1/(e1+e2)) *\<^sub>R (x1-x2)" using e1 e2 by simp then have "x \ affine hull S" using mem_affine_3_minus[of "affine hull S" z x1 x2 "1/(e1+e2)"] x1 x2 z affine_affine_hull[of S] by auto } then have "affine hull S = UNIV" by auto then have "aff_dim S = int DIM('n)" using aff_dim_affine_hull[of S] by (simp) then have False using False by auto } ultimately show ?thesis by auto next case True then have "S \ {}" using aff_dim_empty[of S] by auto have *: "affine hull S = UNIV" using True affine_hull_UNIV by auto { assume "z \ interior S" then have "z \ rel_interior S" using True interior_rel_interior_gen[of S] by auto then have **: "\x. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using convex_rel_interior_iff2[of S z] assms \S \ {}\ * by auto fix x obtain e1 where e1: "e1 > 1" "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z \ S" using **[rule_format, of "z-x"] by auto define e where [abs_def]: "e = e1 - 1" then have "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z = z + e *\<^sub>R x" by (simp add: algebra_simps) then have "e > 0" "z + e *\<^sub>R x \ S" using e1 e_def by auto then have "\e. e > 0 \ z + e *\<^sub>R x \ S" by auto } moreover { assume r: "\x. \e. e > 0 \ z + e *\<^sub>R x \ S" { fix x obtain e1 where e1: "e1 > 0" "z + e1 *\<^sub>R (z - x) \ S" using r[rule_format, of "z-x"] by auto define e where "e = e1 + 1" then have "z + e1 *\<^sub>R (z - x) = (1 - e) *\<^sub>R x + e *\<^sub>R z" by (simp add: algebra_simps) then have "e > 1" "(1 - e)*\<^sub>R x + e *\<^sub>R z \ S" using e1 e_def by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" by auto } then have "z \ rel_interior S" using convex_rel_interior_iff2[of S z] assms \S \ {}\ by auto then have "z \ interior S" using True interior_rel_interior_gen[of S] by auto } ultimately show ?thesis by auto qed subsubsection\<^marker>\tag unimportant\ \Relative interior and closure under common operations\ lemma rel_interior_inter_aux: "\{rel_interior S |S. S \ I} \ \I" proof - { fix y assume "y \ \{rel_interior S |S. S \ I}" then have y: "\S \ I. y \ rel_interior S" by auto { fix S assume "S \ I" then have "y \ S" using rel_interior_subset y by auto } then have "y \ \I" by auto } then show ?thesis by auto qed lemma convex_closure_rel_interior_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" proof - obtain x where x: "\S\I. x \ rel_interior S" using assms by auto { fix y assume "y \ \{closure S |S. S \ I}" then have y: "\S \ I. y \ closure S" by auto { assume "y = x" then have "y \ closure (\{rel_interior S |S. S \ I})" using x closure_subset[of "\{rel_interior S |S. S \ I}"] by auto } moreover { assume "y \ x" { fix e :: real assume e: "e > 0" define e1 where "e1 = min 1 (e/norm (y - x))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (y - x) \ e" using \y \ x\ \e > 0\ le_divide_eq[of e1 e "norm (y - x)"] by simp_all define z where "z = y - e1 *\<^sub>R (y - x)" { fix S assume "S \ I" then have "z \ rel_interior S" using rel_interior_closure_convex_shrink[of S x y e1] assms x y e1 z_def by auto } then have *: "z \ \{rel_interior S |S. S \ I}" by auto have "\z. z \ \{rel_interior S |S. S \ I} \ z \ y \ dist z y \ e" apply (rule_tac x="z" in exI) using \y \ x\ z_def * e1 e dist_norm[of z y] apply simp done } then have "y islimpt \{rel_interior S |S. S \ I}" unfolding islimpt_approachable_le by blast then have "y \ closure (\{rel_interior S |S. S \ I})" unfolding closure_def by auto } ultimately have "y \ closure (\{rel_interior S |S. S \ I})" by auto } then show ?thesis by auto qed lemma convex_closure_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "closure (\I) = \{closure S |S. S \ I}" proof - have "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" using convex_closure_rel_interior_inter assms by auto moreover have "closure (\{rel_interior S |S. S \ I}) \ closure (\I)" using rel_interior_inter_aux closure_mono[of "\{rel_interior S |S. S \ I}" "\I"] by auto ultimately show ?thesis using closure_Int[of I] by auto qed lemma convex_inter_rel_interior_same_closure: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "closure (\{rel_interior S |S. S \ I}) = closure (\I)" proof - have "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" using convex_closure_rel_interior_inter assms by auto moreover have "closure (\{rel_interior S |S. S \ I}) \ closure (\I)" using rel_interior_inter_aux closure_mono[of "\{rel_interior S |S. S \ I}" "\I"] by auto ultimately show ?thesis using closure_Int[of I] by auto qed lemma convex_rel_interior_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "rel_interior (\I) \ \{rel_interior S |S. S \ I}" proof - have "convex (\I)" using assms convex_Inter by auto moreover have "convex (\{rel_interior S |S. S \ I})" apply (rule convex_Inter) using assms convex_rel_interior apply auto done ultimately have "rel_interior (\{rel_interior S |S. S \ I}) = rel_interior (\I)" using convex_inter_rel_interior_same_closure assms closure_eq_rel_interior_eq[of "\{rel_interior S |S. S \ I}" "\I"] by blast then show ?thesis using rel_interior_subset[of "\{rel_interior S |S. S \ I}"] by auto qed lemma convex_rel_interior_finite_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" and "finite I" shows "rel_interior (\I) = \{rel_interior S |S. S \ I}" proof - have "\I \ {}" using assms rel_interior_inter_aux[of I] by auto have "convex (\I)" using convex_Inter assms by auto show ?thesis proof (cases "I = {}") case True then show ?thesis using Inter_empty rel_interior_UNIV by auto next case False { fix z assume z: "z \ \{rel_interior S |S. S \ I}" { fix x assume x: "x \ \I" { fix S assume S: "S \ I" then have "z \ rel_interior S" "x \ S" using z x by auto then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e)*\<^sub>R x + e *\<^sub>R z \ S)" using convex_rel_interior_if[of S z] S assms hull_subset[of S] by auto } then obtain mS where mS: "\S\I. mS S > 1 \ (\e. e > 1 \ e \ mS S \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" by metis define e where "e = Min (mS ` I)" then have "e \ mS ` I" using assms \I \ {}\ by simp then have "e > 1" using mS by auto moreover have "\S\I. e \ mS S" using e_def assms by auto ultimately have "\e > 1. (1 - e) *\<^sub>R x + e *\<^sub>R z \ \I" using mS by auto } then have "z \ rel_interior (\I)" using convex_rel_interior_iff[of "\I" z] \\I \ {}\ \convex (\I)\ by auto } then show ?thesis using convex_rel_interior_inter[of I] assms by auto qed qed lemma convex_closure_inter_two: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" assumes "rel_interior S \ rel_interior T \ {}" shows "closure (S \ T) = closure S \ closure T" using convex_closure_inter[of "{S,T}"] assms by auto lemma convex_rel_interior_inter_two: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" and "rel_interior S \ rel_interior T \ {}" shows "rel_interior (S \ T) = rel_interior S \ rel_interior T" using convex_rel_interior_finite_inter[of "{S,T}"] assms by auto lemma convex_affine_closure_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "rel_interior S \ T \ {}" shows "closure (S \ T) = closure S \ T" proof - have "affine hull T = T" using assms by auto then have "rel_interior T = T" using rel_interior_affine_hull[of T] by metis moreover have "closure T = T" using assms affine_closed[of T] by auto ultimately show ?thesis using convex_closure_inter_two[of S T] assms affine_imp_convex by auto qed lemma connected_component_1_gen: fixes S :: "'a :: euclidean_space set" assumes "DIM('a) = 1" shows "connected_component S a b \ closed_segment a b \ S" unfolding connected_component_def by (metis (no_types, lifting) assms subsetD subsetI convex_contains_segment convex_segment(1) ends_in_segment connected_convex_1_gen) lemma connected_component_1: fixes S :: "real set" shows "connected_component S a b \ closed_segment a b \ S" by (simp add: connected_component_1_gen) lemma convex_affine_rel_interior_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "rel_interior S \ T \ {}" shows "rel_interior (S \ T) = rel_interior S \ T" proof - have "affine hull T = T" using assms by auto then have "rel_interior T = T" using rel_interior_affine_hull[of T] by metis moreover have "closure T = T" using assms affine_closed[of T] by auto ultimately show ?thesis using convex_rel_interior_inter_two[of S T] assms affine_imp_convex by auto qed lemma convex_affine_rel_frontier_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "interior S \ T \ {}" shows "rel_frontier(S \ T) = frontier S \ T" using assms apply (simp add: rel_frontier_def convex_affine_closure_Int frontier_def) by (metis Diff_Int_distrib2 Int_emptyI convex_affine_closure_Int convex_affine_rel_interior_Int empty_iff interior_rel_interior_gen) lemma rel_interior_convex_Int_affine: fixes S :: "'a::euclidean_space set" assumes "convex S" "affine T" "interior S \ T \ {}" shows "rel_interior(S \ T) = interior S \ T" proof - obtain a where aS: "a \ interior S" and aT:"a \ T" using assms by force have "rel_interior S = interior S" by (metis (no_types) aS affine_hull_nonempty_interior equals0D rel_interior_interior) then show ?thesis by (metis (no_types) affine_imp_convex assms convex_rel_interior_inter_two hull_same rel_interior_affine_hull) qed lemma closure_convex_Int_affine: fixes S :: "'a::euclidean_space set" assumes "convex S" "affine T" "rel_interior S \ T \ {}" shows "closure(S \ T) = closure S \ T" proof have "closure (S \ T) \ closure T" by (simp add: closure_mono) also have "... \ T" by (simp add: affine_closed assms) finally show "closure(S \ T) \ closure S \ T" by (simp add: closure_mono) next obtain a where "a \ rel_interior S" "a \ T" using assms by auto then have ssT: "subspace ((\x. (-a)+x) ` T)" and "a \ S" using affine_diffs_subspace rel_interior_subset assms by blast+ show "closure S \ T \ closure (S \ T)" proof fix x assume "x \ closure S \ T" show "x \ closure (S \ T)" proof (cases "x = a") case True then show ?thesis using \a \ S\ \a \ T\ closure_subset by fastforce next case False then have "x \ closure(open_segment a x)" by auto then show ?thesis using \x \ closure S \ T\ assms convex_affine_closure_Int by blast qed qed qed lemma subset_rel_interior_convex: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" and "S \ closure T" and "\ S \ rel_frontier T" shows "rel_interior S \ rel_interior T" proof - have *: "S \ closure T = S" using assms by auto have "\ rel_interior S \ rel_frontier T" using closure_mono[of "rel_interior S" "rel_frontier T"] closed_rel_frontier[of T] closure_closed[of S] convex_closure_rel_interior[of S] closure_subset[of S] assms by auto then have "rel_interior S \ rel_interior (closure T) \ {}" using assms rel_frontier_def[of T] rel_interior_subset convex_rel_interior_closure[of T] by auto then have "rel_interior S \ rel_interior T = rel_interior (S \ closure T)" using assms convex_closure convex_rel_interior_inter_two[of S "closure T"] convex_rel_interior_closure[of T] by auto also have "\ = rel_interior S" using * by auto finally show ?thesis by auto qed lemma rel_interior_convex_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" shows "f ` (rel_interior S) = rel_interior (f ` S)" proof (cases "S = {}") case True then show ?thesis using assms by auto next case False interpret linear f by fact have *: "f ` (rel_interior S) \ f ` S" unfolding image_mono using rel_interior_subset by auto have "f ` S \ f ` (closure S)" unfolding image_mono using closure_subset by auto also have "\ = f ` (closure (rel_interior S))" using convex_closure_rel_interior assms by auto also have "\ \ closure (f ` (rel_interior S))" using closure_linear_image_subset assms by auto finally have "closure (f ` S) = closure (f ` rel_interior S)" using closure_mono[of "f ` S" "closure (f ` rel_interior S)"] closure_closure closure_mono[of "f ` rel_interior S" "f ` S"] * by auto then have "rel_interior (f ` S) = rel_interior (f ` rel_interior S)" using assms convex_rel_interior linear_conv_bounded_linear[of f] convex_linear_image[of _ S] convex_linear_image[of _ "rel_interior S"] closure_eq_rel_interior_eq[of "f ` S" "f ` rel_interior S"] by auto then have "rel_interior (f ` S) \ f ` rel_interior S" using rel_interior_subset by auto moreover { fix z assume "z \ f ` rel_interior S" then obtain z1 where z1: "z1 \ rel_interior S" "f z1 = z" by auto { fix x assume "x \ f ` S" then obtain x1 where x1: "x1 \ S" "f x1 = x" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1 \ S" using convex_rel_interior_iff[of S z1] \convex S\ x1 z1 by auto moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z" using x1 z1 by (simp add: linear_add linear_scale \linear f\) ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using e by auto } then have "z \ rel_interior (f ` S)" using convex_rel_interior_iff[of "f ` S" z] \convex S\ \linear f\ \S \ {}\ convex_linear_image[of f S] linear_conv_bounded_linear[of f] by auto } ultimately show ?thesis by auto qed lemma rel_interior_convex_linear_preimage: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "f -` (rel_interior S) \ {}" shows "rel_interior (f -` S) = f -` (rel_interior S)" proof - interpret linear f by fact have "S \ {}" using assms by auto have nonemp: "f -` S \ {}" by (metis assms(3) rel_interior_subset subset_empty vimage_mono) then have "S \ (range f) \ {}" by auto have conv: "convex (f -` S)" using convex_linear_vimage assms by auto then have "convex (S \ range f)" by (simp add: assms(2) convex_Int convex_linear_image linear_axioms) { fix z assume "z \ f -` (rel_interior S)" then have z: "f z \ rel_interior S" by auto { fix x assume "x \ f -` S" then have "f x \ S" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R f x + e *\<^sub>R f z \ S" using convex_rel_interior_iff[of S "f z"] z assms \S \ {}\ by auto moreover have "(1 - e) *\<^sub>R f x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R x + e *\<^sub>R z)" using \linear f\ by (simp add: linear_iff) ultimately have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f -` S" using e by auto } then have "z \ rel_interior (f -` S)" using convex_rel_interior_iff[of "f -` S" z] conv nonemp by auto } moreover { fix z assume z: "z \ rel_interior (f -` S)" { fix x assume "x \ S \ range f" then obtain y where y: "f y = x" "y \ f -` S" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R y + e *\<^sub>R z \ f -` S" using convex_rel_interior_iff[of "f -` S" z] z conv by auto moreover have "(1 - e) *\<^sub>R x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R y + e *\<^sub>R z)" using \linear f\ y by (simp add: linear_iff) ultimately have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R f z \ S \ range f" using e by auto } then have "f z \ rel_interior (S \ range f)" using \convex (S \ (range f))\ \S \ range f \ {}\ convex_rel_interior_iff[of "S \ (range f)" "f z"] by auto moreover have "affine (range f)" by (simp add: linear_axioms linear_subspace_image subspace_imp_affine) ultimately have "f z \ rel_interior S" using convex_affine_rel_interior_Int[of S "range f"] assms by auto then have "z \ f -` (rel_interior S)" by auto } ultimately show ?thesis by auto qed lemma rel_interior_Times: fixes S :: "'n::euclidean_space set" and T :: "'m::euclidean_space set" assumes "convex S" and "convex T" shows "rel_interior (S \ T) = rel_interior S \ rel_interior T" proof - { assume "S = {}" then have ?thesis by auto } moreover { assume "T = {}" then have ?thesis by auto } moreover { assume "S \ {}" "T \ {}" then have ri: "rel_interior S \ {}" "rel_interior T \ {}" using rel_interior_eq_empty assms by auto then have "fst -` rel_interior S \ {}" using fst_vimage_eq_Times[of "rel_interior S"] by auto then have "rel_interior ((fst :: 'n * 'm \ 'n) -` S) = fst -` rel_interior S" using fst_linear \convex S\ rel_interior_convex_linear_preimage[of fst S] by auto then have s: "rel_interior (S \ (UNIV :: 'm set)) = rel_interior S \ UNIV" by (simp add: fst_vimage_eq_Times) from ri have "snd -` rel_interior T \ {}" using snd_vimage_eq_Times[of "rel_interior T"] by auto then have "rel_interior ((snd :: 'n * 'm \ 'm) -` T) = snd -` rel_interior T" using snd_linear \convex T\ rel_interior_convex_linear_preimage[of snd T] by auto then have t: "rel_interior ((UNIV :: 'n set) \ T) = UNIV \ rel_interior T" by (simp add: snd_vimage_eq_Times) from s t have *: "rel_interior (S \ (UNIV :: 'm set)) \ rel_interior ((UNIV :: 'n set) \ T) = rel_interior S \ rel_interior T" by auto have "S \ T = S \ (UNIV :: 'm set) \ (UNIV :: 'n set) \ T" by auto then have "rel_interior (S \ T) = rel_interior ((S \ (UNIV :: 'm set)) \ ((UNIV :: 'n set) \ T))" by auto also have "\ = rel_interior (S \ (UNIV :: 'm set)) \ rel_interior ((UNIV :: 'n set) \ T)" apply (subst convex_rel_interior_inter_two[of "S \ (UNIV :: 'm set)" "(UNIV :: 'n set) \ T"]) using * ri assms convex_Times apply auto done finally have ?thesis using * by auto } ultimately show ?thesis by blast qed lemma rel_interior_scaleR: fixes S :: "'n::euclidean_space set" assumes "c \ 0" shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)" using rel_interior_injective_linear_image[of "((*\<^sub>R) c)" S] linear_conv_bounded_linear[of "(*\<^sub>R) c"] linear_scaleR injective_scaleR[of c] assms by auto lemma rel_interior_convex_scaleR: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)" by (metis assms linear_scaleR rel_interior_convex_linear_image) lemma convex_rel_open_scaleR: fixes S :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" shows "convex (((*\<^sub>R) c) ` S) \ rel_open (((*\<^sub>R) c) ` S)" by (metis assms convex_scaling rel_interior_convex_scaleR rel_open_def) lemma convex_rel_open_finite_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set) \ rel_open S" and "finite I" shows "convex (\I) \ rel_open (\I)" proof (cases "\{rel_interior S |S. S \ I} = {}") case True then have "\I = {}" using assms unfolding rel_open_def by auto then show ?thesis unfolding rel_open_def by auto next case False then have "rel_open (\I)" using assms unfolding rel_open_def using convex_rel_interior_finite_inter[of I] by auto then show ?thesis using convex_Inter assms by auto qed lemma convex_rel_open_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "rel_open S" shows "convex (f ` S) \ rel_open (f ` S)" by (metis assms convex_linear_image rel_interior_convex_linear_image rel_open_def) lemma convex_rel_open_linear_preimage: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "rel_open S" shows "convex (f -` S) \ rel_open (f -` S)" proof (cases "f -` (rel_interior S) = {}") case True then have "f -` S = {}" using assms unfolding rel_open_def by auto then show ?thesis unfolding rel_open_def by auto next case False then have "rel_open (f -` S)" using assms unfolding rel_open_def using rel_interior_convex_linear_preimage[of f S] by auto then show ?thesis using convex_linear_vimage assms by auto qed lemma rel_interior_projection: fixes S :: "('m::euclidean_space \ 'n::euclidean_space) set" and f :: "'m::euclidean_space \ 'n::euclidean_space set" assumes "convex S" and "f = (\y. {z. (y, z) \ S})" shows "(y, z) \ rel_interior S \ (y \ rel_interior {y. (f y \ {})} \ z \ rel_interior (f y))" proof - { fix y assume "y \ {y. f y \ {}}" then obtain z where "(y, z) \ S" using assms by auto then have "\x. x \ S \ y = fst x" apply (rule_tac x="(y, z)" in exI) apply auto done then obtain x where "x \ S" "y = fst x" by blast then have "y \ fst ` S" unfolding image_def by auto } then have "fst ` S = {y. f y \ {}}" unfolding fst_def using assms by auto then have h1: "fst ` rel_interior S = rel_interior {y. f y \ {}}" using rel_interior_convex_linear_image[of fst S] assms fst_linear by auto { fix y assume "y \ rel_interior {y. f y \ {}}" then have "y \ fst ` rel_interior S" using h1 by auto then have *: "rel_interior S \ fst -` {y} \ {}" by auto moreover have aff: "affine (fst -` {y})" unfolding affine_alt by (simp add: algebra_simps) ultimately have **: "rel_interior (S \ fst -` {y}) = rel_interior S \ fst -` {y}" using convex_affine_rel_interior_Int[of S "fst -` {y}"] assms by auto have conv: "convex (S \ fst -` {y})" using convex_Int assms aff affine_imp_convex by auto { fix x assume "x \ f y" then have "(y, x) \ S \ (fst -` {y})" using assms by auto moreover have "x = snd (y, x)" by auto ultimately have "x \ snd ` (S \ fst -` {y})" by blast } then have "snd ` (S \ fst -` {y}) = f y" using assms by auto then have ***: "rel_interior (f y) = snd ` rel_interior (S \ fst -` {y})" using rel_interior_convex_linear_image[of snd "S \ fst -` {y}"] snd_linear conv by auto { fix z assume "z \ rel_interior (f y)" then have "z \ snd ` rel_interior (S \ fst -` {y})" using *** by auto moreover have "{y} = fst ` rel_interior (S \ fst -` {y})" using * ** rel_interior_subset by auto ultimately have "(y, z) \ rel_interior (S \ fst -` {y})" by force then have "(y,z) \ rel_interior S" using ** by auto } moreover { fix z assume "(y, z) \ rel_interior S" then have "(y, z) \ rel_interior (S \ fst -` {y})" using ** by auto then have "z \ snd ` rel_interior (S \ fst -` {y})" by (metis Range_iff snd_eq_Range) then have "z \ rel_interior (f y)" using *** by auto } ultimately have "\z. (y, z) \ rel_interior S \ z \ rel_interior (f y)" by auto } then have h2: "\y z. y \ rel_interior {t. f t \ {}} \ (y, z) \ rel_interior S \ z \ rel_interior (f y)" by auto { fix y z assume asm: "(y, z) \ rel_interior S" then have "y \ fst ` rel_interior S" by (metis Domain_iff fst_eq_Domain) then have "y \ rel_interior {t. f t \ {}}" using h1 by auto then have "y \ rel_interior {t. f t \ {}}" and "(z \ rel_interior (f y))" using h2 asm by auto } then show ?thesis using h2 by blast qed lemma rel_frontier_Times: fixes S :: "'n::euclidean_space set" and T :: "'m::euclidean_space set" assumes "convex S" and "convex T" shows "rel_frontier S \ rel_frontier T \ rel_frontier (S \ T)" by (force simp: rel_frontier_def rel_interior_Times assms closure_Times) subsubsection\<^marker>\tag unimportant\ \Relative interior of convex cone\ lemma cone_rel_interior: fixes S :: "'m::euclidean_space set" assumes "cone S" shows "cone ({0} \ rel_interior S)" proof (cases "S = {}") case True then show ?thesis by (simp add: cone_0) next case False then have *: "0 \ S \ (\c. c > 0 \ (*\<^sub>R) c ` S = S)" using cone_iff[of S] assms by auto then have *: "0 \ ({0} \ rel_interior S)" and "\c. c > 0 \ (*\<^sub>R) c ` ({0} \ rel_interior S) = ({0} \ rel_interior S)" by (auto simp add: rel_interior_scaleR) then show ?thesis using cone_iff[of "{0} \ rel_interior S"] by auto qed lemma rel_interior_convex_cone_aux: fixes S :: "'m::euclidean_space set" assumes "convex S" shows "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) \ c > 0 \ x \ (((*\<^sub>R) c) ` (rel_interior S))" proof (cases "S = {}") case True then show ?thesis by (simp add: cone_hull_empty) next case False then obtain s where "s \ S" by auto have conv: "convex ({(1 :: real)} \ S)" using convex_Times[of "{(1 :: real)}" S] assms convex_singleton[of "1 :: real"] by auto define f where "f y = {z. (y, z) \ cone hull ({1 :: real} \ S)}" for y then have *: "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) = (c \ rel_interior {y. f y \ {}} \ x \ rel_interior (f c))" apply (subst rel_interior_projection[of "cone hull ({(1 :: real)} \ S)" f c x]) using convex_cone_hull[of "{(1 :: real)} \ S"] conv apply auto done { fix y :: real assume "y \ 0" then have "y *\<^sub>R (1,s) \ cone hull ({1 :: real} \ S)" using cone_hull_expl[of "{(1 :: real)} \ S"] \s \ S\ by auto then have "f y \ {}" using f_def by auto } then have "{y. f y \ {}} = {0..}" using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto then have **: "rel_interior {y. f y \ {}} = {0<..}" using rel_interior_real_semiline by auto { fix c :: real assume "c > 0" then have "f c = ((*\<^sub>R) c ` S)" using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto then have "rel_interior (f c) = (*\<^sub>R) c ` rel_interior S" using rel_interior_convex_scaleR[of S c] assms by auto } then show ?thesis using * ** by auto qed lemma rel_interior_convex_cone: fixes S :: "'m::euclidean_space set" assumes "convex S" shows "rel_interior (cone hull ({1 :: real} \ S)) = {(c, c *\<^sub>R x) | c x. c > 0 \ x \ rel_interior S}" (is "?lhs = ?rhs") proof - { fix z assume "z \ ?lhs" have *: "z = (fst z, snd z)" by auto then have "z \ ?rhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \z \ ?lhs\ by fastforce } moreover { fix z assume "z \ ?rhs" then have "z \ ?lhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms by auto } ultimately show ?thesis by blast qed lemma convex_hull_finite_union: assumes "finite I" assumes "\i\I. convex (S i) \ (S i) \ {}" shows "convex hull (\(S ` I)) = {sum (\i. c i *\<^sub>R s i) I | c s. (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. s i \ S i)}" (is "?lhs = ?rhs") proof - have "?lhs \ ?rhs" proof fix x assume "x \ ?rhs" then obtain c s where *: "sum (\i. c i *\<^sub>R s i) I = x" "sum c I = 1" "(\i\I. c i \ 0) \ (\i\I. s i \ S i)" by auto then have "\i\I. s i \ convex hull (\(S ` I))" using hull_subset[of "\(S ` I)" convex] by auto then show "x \ ?lhs" unfolding *(1)[symmetric] apply (subst convex_sum[of I "convex hull \(S ` I)" c s]) using * assms convex_convex_hull apply auto done qed { fix i assume "i \ I" with assms have "\p. p \ S i" by auto } then obtain p where p: "\i\I. p i \ S i" by metis { fix i assume "i \ I" { fix x assume "x \ S i" define c where "c j = (if j = i then 1::real else 0)" for j then have *: "sum c I = 1" using \finite I\ \i \ I\ sum.delta[of I i "\j::'a. 1::real"] by auto define s where "s j = (if j = i then x else p j)" for j then have "\j. c j *\<^sub>R s j = (if j = i then x else 0)" using c_def by (auto simp add: algebra_simps) then have "x = sum (\i. c i *\<^sub>R s i) I" using s_def c_def \finite I\ \i \ I\ sum.delta[of I i "\j::'a. x"] by auto then have "x \ ?rhs" apply auto apply (rule_tac x = c in exI) apply (rule_tac x = s in exI) using * c_def s_def p \x \ S i\ apply auto done } then have "?rhs \ S i" by auto } then have *: "?rhs \ \(S ` I)" by auto { fix u v :: real assume uv: "u \ 0 \ v \ 0 \ u + v = 1" fix x y assume xy: "x \ ?rhs \ y \ ?rhs" from xy obtain c s where xc: "x = sum (\i. c i *\<^sub>R s i) I \ (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. s i \ S i)" by auto from xy obtain d t where yc: "y = sum (\i. d i *\<^sub>R t i) I \ (\i\I. d i \ 0) \ sum d I = 1 \ (\i\I. t i \ S i)" by auto define e where "e i = u * c i + v * d i" for i have ge0: "\i\I. e i \ 0" using e_def xc yc uv by simp have "sum (\i. u * c i) I = u * sum c I" by (simp add: sum_distrib_left) moreover have "sum (\i. v * d i) I = v * sum d I" by (simp add: sum_distrib_left) ultimately have sum1: "sum e I = 1" using e_def xc yc uv by (simp add: sum.distrib) define q where "q i = (if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i)" for i { fix i assume i: "i \ I" have "q i \ S i" proof (cases "e i = 0") case True then show ?thesis using i p q_def by auto next case False then show ?thesis using mem_convex_alt[of "S i" "s i" "t i" "u * (c i)" "v * (d i)"] mult_nonneg_nonneg[of u "c i"] mult_nonneg_nonneg[of v "d i"] assms q_def e_def i False xc yc uv by (auto simp del: mult_nonneg_nonneg) qed } then have qs: "\i\I. q i \ S i" by auto { fix i assume i: "i \ I" have "(u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" proof (cases "e i = 0") case True have ge: "u * (c i) \ 0 \ v * d i \ 0" using xc yc uv i by simp moreover from ge have "u * c i \ 0 \ v * d i \ 0" using True e_def i by simp ultimately have "u * c i = 0 \ v * d i = 0" by auto with True show ?thesis by auto next case False then have "(u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i) = q i" using q_def by auto then have "e i *\<^sub>R ((u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i)) = (e i) *\<^sub>R (q i)" by auto with False show ?thesis by (simp add: algebra_simps) qed } then have *: "\i\I. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" by auto have "u *\<^sub>R x + v *\<^sub>R y = sum (\i. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i) I" using xc yc by (simp add: algebra_simps scaleR_right.sum sum.distrib) also have "\ = sum (\i. e i *\<^sub>R q i) I" using * by auto finally have "u *\<^sub>R x + v *\<^sub>R y = sum (\i. (e i) *\<^sub>R (q i)) I" by auto then have "u *\<^sub>R x + v *\<^sub>R y \ ?rhs" using ge0 sum1 qs by auto } then have "convex ?rhs" unfolding convex_def by auto then show ?thesis using \?lhs \ ?rhs\ * hull_minimal[of "\(S ` I)" ?rhs convex] by blast qed lemma convex_hull_union_two: fixes S T :: "'m::euclidean_space set" assumes "convex S" and "S \ {}" and "convex T" and "T \ {}" shows "convex hull (S \ T) = {u *\<^sub>R s + v *\<^sub>R t | u v s t. u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T}" (is "?lhs = ?rhs") proof define I :: "nat set" where "I = {1, 2}" define s where "s i = (if i = (1::nat) then S else T)" for i have "\(s ` I) = S \ T" using s_def I_def by auto then have "convex hull (\(s ` I)) = convex hull (S \ T)" by auto moreover have "convex hull \(s ` I) = {\ i\I. c i *\<^sub>R sa i | c sa. (\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. sa i \ s i)}" apply (subst convex_hull_finite_union[of I s]) using assms s_def I_def apply auto done moreover have "{\i\I. c i *\<^sub>R sa i | c sa. (\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. sa i \ s i)} \ ?rhs" using s_def I_def by auto ultimately show "?lhs \ ?rhs" by auto { fix x assume "x \ ?rhs" then obtain u v s t where *: "x = u *\<^sub>R s + v *\<^sub>R t \ u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T" by auto then have "x \ convex hull {s, t}" using convex_hull_2[of s t] by auto then have "x \ convex hull (S \ T)" using * hull_mono[of "{s, t}" "S \ T"] by auto } then show "?lhs \ ?rhs" by blast qed proposition ray_to_rel_frontier: fixes a :: "'a::real_inner" assumes "bounded S" and a: "a \ rel_interior S" and aff: "(a + l) \ affine hull S" and "l \ 0" obtains d where "0 < d" "(a + d *\<^sub>R l) \ rel_frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ rel_interior S" proof - have aaff: "a \ affine hull S" by (meson a hull_subset rel_interior_subset rev_subsetD) let ?D = "{d. 0 < d \ a + d *\<^sub>R l \ rel_interior S}" obtain B where "B > 0" and B: "S \ ball a B" using bounded_subset_ballD [OF \bounded S\] by blast have "a + (B / norm l) *\<^sub>R l \ ball a B" by (simp add: dist_norm \l \ 0\) with B have "a + (B / norm l) *\<^sub>R l \ rel_interior S" using rel_interior_subset subsetCE by blast with \B > 0\ \l \ 0\ have nonMT: "?D \ {}" using divide_pos_pos zero_less_norm_iff by fastforce have bdd: "bdd_below ?D" by (metis (no_types, lifting) bdd_belowI le_less mem_Collect_eq) have relin_Ex: "\x. x \ rel_interior S \ \e>0. \x'\affine hull S. dist x' x < e \ x' \ rel_interior S" using openin_rel_interior [of S] by (simp add: openin_euclidean_subtopology_iff) define d where "d = Inf ?D" obtain \ where "0 < \" and \: "\\. \0 \ \; \ < \\ \ (a + \ *\<^sub>R l) \ rel_interior S" proof - obtain e where "e>0" and e: "\x'. x' \ affine hull S \ dist x' a < e \ x' \ rel_interior S" using relin_Ex a by blast show thesis proof (rule_tac \ = "e / norm l" in that) show "0 < e / norm l" by (simp add: \0 < e\ \l \ 0\) next show "a + \ *\<^sub>R l \ rel_interior S" if "0 \ \" "\ < e / norm l" for \ proof (rule e) show "a + \ *\<^sub>R l \ affine hull S" by (metis (no_types) add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff) show "dist (a + \ *\<^sub>R l) a < e" using that by (simp add: \l \ 0\ dist_norm pos_less_divide_eq) qed qed qed have inint: "\e. \0 \ e; e < d\ \ a + e *\<^sub>R l \ rel_interior S" unfolding d_def using cInf_lower [OF _ bdd] by (metis (no_types, lifting) a add.right_neutral le_less mem_Collect_eq not_less real_vector.scale_zero_left) have "\ \ d" unfolding d_def apply (rule cInf_greatest [OF nonMT]) using \ dual_order.strict_implies_order le_less_linear by blast with \0 < \\ have "0 < d" by simp have "a + d *\<^sub>R l \ rel_interior S" proof assume adl: "a + d *\<^sub>R l \ rel_interior S" obtain e where "e > 0" and e: "\x'. x' \ affine hull S \ dist x' (a + d *\<^sub>R l) < e \ x' \ rel_interior S" using relin_Ex adl by blast have "d + e / norm l \ Inf {d. 0 < d \ a + d *\<^sub>R l \ rel_interior S}" proof (rule cInf_greatest [OF nonMT], clarsimp) fix x::real assume "0 < x" and nonrel: "a + x *\<^sub>R l \ rel_interior S" show "d + e / norm l \ x" proof (cases "x < d") case True with inint nonrel \0 < x\ show ?thesis by auto next case False then have dle: "x < d + e / norm l \ dist (a + x *\<^sub>R l) (a + d *\<^sub>R l) < e" by (simp add: field_simps \l \ 0\) have ain: "a + x *\<^sub>R l \ affine hull S" by (metis add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff) show ?thesis using e [OF ain] nonrel dle by force qed qed then show False using \0 < e\ \l \ 0\ by (simp add: d_def [symmetric] field_simps) qed moreover have "a + d *\<^sub>R l \ closure S" proof (clarsimp simp: closure_approachable) fix \::real assume "0 < \" have 1: "a + (d - min d (\ / 2 / norm l)) *\<^sub>R l \ S" apply (rule subsetD [OF rel_interior_subset inint]) using \l \ 0\ \0 < d\ \0 < \\ by auto have "norm l * min d (\ / (norm l * 2)) \ norm l * (\ / (norm l * 2))" by (metis min_def mult_left_mono norm_ge_zero order_refl) also have "... < \" using \l \ 0\ \0 < \\ by (simp add: field_simps) finally have 2: "norm l * min d (\ / (norm l * 2)) < \" . show "\y\S. dist y (a + d *\<^sub>R l) < \" apply (rule_tac x="a + (d - min d (\ / 2 / norm l)) *\<^sub>R l" in bexI) using 1 2 \0 < d\ \0 < \\ apply (auto simp: algebra_simps) done qed ultimately have infront: "a + d *\<^sub>R l \ rel_frontier S" by (simp add: rel_frontier_def) show ?thesis by (rule that [OF \0 < d\ infront inint]) qed corollary ray_to_frontier: fixes a :: "'a::euclidean_space" assumes "bounded S" and a: "a \ interior S" and "l \ 0" obtains d where "0 < d" "(a + d *\<^sub>R l) \ frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ interior S" proof - have "interior S = rel_interior S" using a rel_interior_nonempty_interior by auto then have "a \ rel_interior S" using a by simp then show ?thesis apply (rule ray_to_rel_frontier [OF \bounded S\ _ _ \l \ 0\]) using a affine_hull_nonempty_interior apply blast by (simp add: \interior S = rel_interior S\ frontier_def rel_frontier_def that) qed lemma segment_to_rel_frontier_aux: fixes x :: "'a::euclidean_space" assumes "convex S" "bounded S" and x: "x \ rel_interior S" and y: "y \ S" and xy: "x \ y" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof - have "x + (y - x) \ affine hull S" using hull_inc [OF y] by auto then obtain d where "0 < d" and df: "(x + d *\<^sub>R (y-x)) \ rel_frontier S" and di: "\e. \0 \ e; e < d\ \ (x + e *\<^sub>R (y-x)) \ rel_interior S" by (rule ray_to_rel_frontier [OF \bounded S\ x]) (use xy in auto) show ?thesis proof show "x + d *\<^sub>R (y - x) \ rel_frontier S" by (simp add: df) next have "open_segment x y \ rel_interior S" using rel_interior_closure_convex_segment [OF \convex S\ x] closure_subset y by blast moreover have "x + d *\<^sub>R (y - x) \ open_segment x y" if "d < 1" using xy apply (auto simp: in_segment) apply (rule_tac x="d" in exI) using \0 < d\ that apply (auto simp: algebra_simps) done ultimately have "1 \ d" using df rel_frontier_def by fastforce moreover have "x = (1 / d) *\<^sub>R x + ((d - 1) / d) *\<^sub>R x" by (metis \0 < d\ add.commute add_divide_distrib diff_add_cancel divide_self_if less_irrefl scaleR_add_left scaleR_one) ultimately show "y \ closed_segment x (x + d *\<^sub>R (y - x))" apply (auto simp: in_segment) apply (rule_tac x="1/d" in exI) apply (auto simp: algebra_simps) done next show "open_segment x (x + d *\<^sub>R (y - x)) \ rel_interior S" apply (rule rel_interior_closure_convex_segment [OF \convex S\ x]) using df rel_frontier_def by auto qed qed lemma segment_to_rel_frontier: fixes x :: "'a::euclidean_space" assumes S: "convex S" "bounded S" and x: "x \ rel_interior S" and y: "y \ S" and xy: "\(x = y \ S = {x})" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof (cases "x=y") case True with xy have "S \ {x}" by blast with True show ?thesis by (metis Set.set_insert all_not_in_conv ends_in_segment(1) insert_iff segment_to_rel_frontier_aux[OF S x] that y) next case False then show ?thesis using segment_to_rel_frontier_aux [OF S x y] that by blast qed proposition rel_frontier_not_sing: fixes a :: "'a::euclidean_space" assumes "bounded S" shows "rel_frontier S \ {a}" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain z where "z \ S" by blast then show ?thesis proof (cases "S = {z}") case True then show ?thesis by simp next case False then obtain w where "w \ S" "w \ z" using \z \ S\ by blast show ?thesis proof assume "rel_frontier S = {a}" then consider "w \ rel_frontier S" | "z \ rel_frontier S" using \w \ z\ by auto then show False proof cases case 1 then have w: "w \ rel_interior S" using \w \ S\ closure_subset rel_frontier_def by fastforce have "w + (w - z) \ affine hull S" by (metis \w \ S\ \z \ S\ affine_affine_hull hull_inc mem_affine_3_minus scaleR_one) then obtain e where "0 < e" "(w + e *\<^sub>R (w - z)) \ rel_frontier S" using \w \ z\ \z \ S\ by (metis assms ray_to_rel_frontier right_minus_eq w) moreover obtain d where "0 < d" "(w + d *\<^sub>R (z - w)) \ rel_frontier S" using ray_to_rel_frontier [OF \bounded S\ w, of "1 *\<^sub>R (z - w)"] \w \ z\ \z \ S\ by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one) ultimately have "d *\<^sub>R (z - w) = e *\<^sub>R (w - z)" using \rel_frontier S = {a}\ by force moreover have "e \ -d " using \0 < e\ \0 < d\ by force ultimately show False by (metis (no_types, lifting) \w \ z\ eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus) next case 2 then have z: "z \ rel_interior S" using \z \ S\ closure_subset rel_frontier_def by fastforce have "z + (z - w) \ affine hull S" by (metis \z \ S\ \w \ S\ affine_affine_hull hull_inc mem_affine_3_minus scaleR_one) then obtain e where "0 < e" "(z + e *\<^sub>R (z - w)) \ rel_frontier S" using \w \ z\ \w \ S\ by (metis assms ray_to_rel_frontier right_minus_eq z) moreover obtain d where "0 < d" "(z + d *\<^sub>R (w - z)) \ rel_frontier S" using ray_to_rel_frontier [OF \bounded S\ z, of "1 *\<^sub>R (w - z)"] \w \ z\ \w \ S\ by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one) ultimately have "d *\<^sub>R (w - z) = e *\<^sub>R (z - w)" using \rel_frontier S = {a}\ by force moreover have "e \ -d " using \0 < e\ \0 < d\ by force ultimately show False by (metis (no_types, lifting) \w \ z\ eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus) qed qed qed qed subsection\<^marker>\tag unimportant\ \Convexity on direct sums\ lemma closure_sum: fixes S T :: "'a::real_normed_vector set" shows "closure S + closure T \ closure (S + T)" unfolding set_plus_image closure_Times [symmetric] split_def by (intro closure_bounded_linear_image_subset bounded_linear_add bounded_linear_fst bounded_linear_snd) lemma rel_interior_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" shows "rel_interior (S + T) = rel_interior S + rel_interior T" proof - have "rel_interior S + rel_interior T = (\(x,y). x + y) ` (rel_interior S \ rel_interior T)" by (simp add: set_plus_image) also have "\ = (\(x,y). x + y) ` rel_interior (S \ T)" using rel_interior_Times assms by auto also have "\ = rel_interior (S + T)" using fst_snd_linear convex_Times assms rel_interior_convex_linear_image[of "(\(x,y). x + y)" "S \ T"] by (auto simp add: set_plus_image) finally show ?thesis .. qed lemma rel_interior_sum_gen: fixes S :: "'a \ 'n::euclidean_space set" assumes "\i\I. convex (S i)" shows "rel_interior (sum S I) = sum (\i. rel_interior (S i)) I" apply (subst sum_set_cond_linear[of convex]) using rel_interior_sum rel_interior_sing[of "0"] assms apply (auto simp add: convex_set_plus) done lemma convex_rel_open_direct_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" and "convex T" and "rel_open T" shows "convex (S \ T) \ rel_open (S \ T)" by (metis assms convex_Times rel_interior_Times rel_open_def) lemma convex_rel_open_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" and "convex T" and "rel_open T" shows "convex (S + T) \ rel_open (S + T)" by (metis assms convex_set_plus rel_interior_sum rel_open_def) lemma convex_hull_finite_union_cones: assumes "finite I" and "I \ {}" assumes "\i\I. convex (S i) \ cone (S i) \ S i \ {}" shows "convex hull (\(S ` I)) = sum S I" (is "?lhs = ?rhs") proof - { fix x assume "x \ ?lhs" then obtain c xs where x: "x = sum (\i. c i *\<^sub>R xs i) I \ (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. xs i \ S i)" using convex_hull_finite_union[of I S] assms by auto define s where "s i = c i *\<^sub>R xs i" for i { fix i assume "i \ I" then have "s i \ S i" using s_def x assms mem_cone[of "S i" "xs i" "c i"] by auto } then have "\i\I. s i \ S i" by auto moreover have "x = sum s I" using x s_def by auto ultimately have "x \ ?rhs" using set_sum_alt[of I S] assms by auto } moreover { fix x assume "x \ ?rhs" then obtain s where x: "x = sum s I \ (\i\I. s i \ S i)" using set_sum_alt[of I S] assms by auto define xs where "xs i = of_nat(card I) *\<^sub>R s i" for i then have "x = sum (\i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I" using x assms by auto moreover have "\i\I. xs i \ S i" using x xs_def assms by (simp add: cone_def) moreover have "\i\I. (1 :: real) / of_nat (card I) \ 0" by auto moreover have "sum (\i. (1 :: real) / of_nat (card I)) I = 1" using assms by auto ultimately have "x \ ?lhs" apply (subst convex_hull_finite_union[of I S]) using assms apply blast using assms apply blast apply rule apply (rule_tac x = "(\i. (1 :: real) / of_nat (card I))" in exI) apply auto done } ultimately show ?thesis by auto qed lemma convex_hull_union_cones_two: fixes S T :: "'m::euclidean_space set" assumes "convex S" and "cone S" and "S \ {}" assumes "convex T" and "cone T" and "T \ {}" shows "convex hull (S \ T) = S + T" proof - define I :: "nat set" where "I = {1, 2}" define A where "A i = (if i = (1::nat) then S else T)" for i have "\(A ` I) = S \ T" using A_def I_def by auto then have "convex hull (\(A ` I)) = convex hull (S \ T)" by auto moreover have "convex hull \(A ` I) = sum A I" apply (subst convex_hull_finite_union_cones[of I A]) using assms A_def I_def apply auto done moreover have "sum A I = S + T" using A_def I_def unfolding set_plus_def apply auto unfolding set_plus_def apply auto done ultimately show ?thesis by auto qed lemma rel_interior_convex_hull_union: fixes S :: "'a \ 'n::euclidean_space set" assumes "finite I" and "\i\I. convex (S i) \ S i \ {}" shows "rel_interior (convex hull (\(S ` I))) = {sum (\i. c i *\<^sub>R s i) I | c s. (\i\I. c i > 0) \ sum c I = 1 \ (\i\I. s i \ rel_interior(S i))}" (is "?lhs = ?rhs") proof (cases "I = {}") case True then show ?thesis using convex_hull_empty by auto next case False define C0 where "C0 = convex hull (\(S ` I))" have "\i\I. C0 \ S i" unfolding C0_def using hull_subset[of "\(S ` I)"] by auto define K0 where "K0 = cone hull ({1 :: real} \ C0)" define K where "K i = cone hull ({1 :: real} \ S i)" for i have "\i\I. K i \ {}" unfolding K_def using assms by (simp add: cone_hull_empty_iff[symmetric]) { fix i assume "i \ I" then have "convex (K i)" unfolding K_def apply (subst convex_cone_hull) apply (subst convex_Times) using assms apply auto done } then have convK: "\i\I. convex (K i)" by auto { fix i assume "i \ I" then have "K0 \ K i" unfolding K0_def K_def apply (subst hull_mono) using \\i\I. C0 \ S i\ apply auto done } then have "K0 \ \(K ` I)" by auto moreover have "convex K0" unfolding K0_def apply (subst convex_cone_hull) apply (subst convex_Times) unfolding C0_def using convex_convex_hull apply auto done ultimately have geq: "K0 \ convex hull (\(K ` I))" using hull_minimal[of _ "K0" "convex"] by blast have "\i\I. K i \ {1 :: real} \ S i" using K_def by (simp add: hull_subset) then have "\(K ` I) \ {1 :: real} \ \(S ` I)" by auto then have "convex hull \(K ` I) \ convex hull ({1 :: real} \ \(S ` I))" by (simp add: hull_mono) then have "convex hull \(K ` I) \ {1 :: real} \ C0" unfolding C0_def using convex_hull_Times[of "{(1 :: real)}" "\(S ` I)"] convex_hull_singleton by auto moreover have "cone (convex hull (\(K ` I)))" apply (subst cone_convex_hull) using cone_Union[of "K ` I"] apply auto unfolding K_def using cone_cone_hull apply auto done ultimately have "convex hull (\(K ` I)) \ K0" unfolding K0_def using hull_minimal[of _ "convex hull (\(K ` I))" "cone"] by blast then have "K0 = convex hull (\(K ` I))" using geq by auto also have "\ = sum K I" apply (subst convex_hull_finite_union_cones[of I K]) using assms apply blast using False apply blast unfolding K_def apply rule apply (subst convex_cone_hull) apply (subst convex_Times) using assms cone_cone_hull \\i\I. K i \ {}\ K_def apply auto done finally have "K0 = sum K I" by auto then have *: "rel_interior K0 = sum (\i. (rel_interior (K i))) I" using rel_interior_sum_gen[of I K] convK by auto { fix x assume "x \ ?lhs" then have "(1::real, x) \ rel_interior K0" using K0_def C0_def rel_interior_convex_cone_aux[of C0 "1::real" x] convex_convex_hull by auto then obtain k where k: "(1::real, x) = sum k I \ (\i\I. k i \ rel_interior (K i))" using \finite I\ * set_sum_alt[of I "\i. rel_interior (K i)"] by auto { fix i assume "i \ I" then have "convex (S i) \ k i \ rel_interior (cone hull {1} \ S i)" using k K_def assms by auto then have "\ci si. k i = (ci, ci *\<^sub>R si) \ 0 < ci \ si \ rel_interior (S i)" using rel_interior_convex_cone[of "S i"] by auto } then obtain c s where cs: "\i\I. k i = (c i, c i *\<^sub>R s i) \ 0 < c i \ s i \ rel_interior (S i)" by metis then have "x = (\i\I. c i *\<^sub>R s i) \ sum c I = 1" using k by (simp add: sum_prod) then have "x \ ?rhs" using k cs by auto } moreover { fix x assume "x \ ?rhs" then obtain c s where cs: "x = sum (\i. c i *\<^sub>R s i) I \ (\i\I. c i > 0) \ sum c I = 1 \ (\i\I. s i \ rel_interior (S i))" by auto define k where "k i = (c i, c i *\<^sub>R s i)" for i { fix i assume "i \ I" then have "k i \ rel_interior (K i)" using k_def K_def assms cs rel_interior_convex_cone[of "S i"] by auto } then have "(1::real, x) \ rel_interior K0" using K0_def * set_sum_alt[of I "(\i. rel_interior (K i))"] assms k_def cs apply auto apply (rule_tac x = k in exI) apply (simp add: sum_prod) done then have "x \ ?lhs" using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x] by auto } ultimately show ?thesis by blast qed lemma convex_le_Inf_differential: fixes f :: "real \ real" assumes "convex_on I f" and "x \ interior I" and "y \ I" shows "f y \ f x + Inf ((\t. (f x - f t) / (x - t)) ` ({x<..} \ I)) * (y - x)" (is "_ \ _ + Inf (?F x) * (y - x)") proof (cases rule: linorder_cases) assume "x < y" moreover have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . moreover define t where "t = min (x + e / 2) ((x + y) / 2)" ultimately have "x < t" "t < y" "t \ ball x e" by (auto simp: dist_real_def field_simps split: split_min) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where "0 < e" "ball x e \ interior I" . moreover define K where "K = x - e / 2" with \0 < e\ have "K \ ball x e" "K < x" by (auto simp: dist_real_def) ultimately have "K \ I" "K < x" "x \ I" using interior_subset[of I] \x \ interior I\ by auto have "Inf (?F x) \ (f x - f y) / (x - y)" proof (intro bdd_belowI cInf_lower2) show "(f x - f t) / (x - t) \ ?F x" using \t \ I\ \x < t\ by auto show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" using \convex_on I f\ \x \ I\ \y \ I\ \x < t\ \t < y\ by (rule convex_on_diff) next fix y assume "y \ ?F x" with order_trans[OF convex_on_diff[OF \convex_on I f\ \K \ I\ _ \K < x\ _]] show "(f K - f x) / (K - x) \ y" by auto qed then show ?thesis using \x < y\ by (simp add: field_simps) next assume "y < x" moreover have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . moreover define t where "t = x + e / 2" ultimately have "x < t" "t \ ball x e" by (auto simp: dist_real_def field_simps) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto have "(f x - f y) / (x - y) \ Inf (?F x)" proof (rule cInf_greatest) have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" using \y < x\ by (auto simp: field_simps) also fix z assume "z \ ?F x" with order_trans[OF convex_on_diff[OF \convex_on I f\ \y \ I\ _ \y < x\]] have "(f y - f x) / (y - x) \ z" by auto finally show "(f x - f y) / (x - y) \ z" . next have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . then have "x + e / 2 \ ball x e" by (auto simp: dist_real_def) with e interior_subset[of I] have "x + e / 2 \ {x<..} \ I" by auto then show "?F x \ {}" by blast qed then show ?thesis using \y < x\ by (simp add: field_simps) qed simp subsection\<^marker>\tag unimportant\\Explicit formulas for interior and relative interior of convex hull\ lemma at_within_cbox_finite: assumes "x \ box a b" "x \ S" "finite S" shows "(at x within cbox a b - S) = at x" proof - have "interior (cbox a b - S) = box a b - S" using \finite S\ by (simp add: interior_diff finite_imp_closed) then show ?thesis using at_within_interior assms by fastforce qed lemma affine_independent_convex_affine_hull: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" "t \ s" shows "convex hull t = affine hull t \ convex hull s" proof - have fin: "finite s" "finite t" using assms aff_independent_finite finite_subset by auto { fix u v x assume uv: "sum u t = 1" "\x\s. 0 \ v x" "sum v s = 1" "(\x\s. v x *\<^sub>R x) = (\v\t. u v *\<^sub>R v)" "x \ t" then have s: "s = (s - t) \ t" \ \split into separate cases\ using assms by auto have [simp]: "(\x\t. v x *\<^sub>R x) + (\x\s - t. v x *\<^sub>R x) = (\x\t. u x *\<^sub>R x)" "sum v t + sum v (s - t) = 1" using uv fin s by (auto simp: sum.union_disjoint [symmetric] Un_commute) have "(\x\s. if x \ t then v x - u x else v x) = 0" "(\x\s. (if x \ t then v x - u x else v x) *\<^sub>R x) = 0" using uv fin by (subst s, subst sum.union_disjoint, auto simp: algebra_simps sum_subtractf)+ } note [simp] = this have "convex hull t \ affine hull t" using convex_hull_subset_affine_hull by blast moreover have "convex hull t \ convex hull s" using assms hull_mono by blast moreover have "affine hull t \ convex hull s \ convex hull t" using assms apply (simp add: convex_hull_finite affine_hull_finite fin affine_dependent_explicit) apply (drule_tac x=s in spec) apply (auto simp: fin) apply (rule_tac x=u in exI) apply (rename_tac v) apply (drule_tac x="\x. if x \ t then v x - u x else v x" in spec) apply (force)+ done ultimately show ?thesis by blast qed lemma affine_independent_span_eq: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" "card s = Suc (DIM ('a))" shows "affine hull s = UNIV" proof (cases "s = {}") case True then show ?thesis using assms by simp next case False then obtain a t where t: "a \ t" "s = insert a t" by blast then have fin: "finite t" using assms by (metis finite_insert aff_independent_finite) show ?thesis using assms t fin apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen) apply (rule subset_antisym) apply force apply (rule Fun.vimage_subsetD) apply (metis add.commute diff_add_cancel surj_def) apply (rule card_ge_dim_independent) apply (auto simp: card_image inj_on_def dim_subset_UNIV) done qed lemma affine_independent_span_gt: fixes s :: "'a::euclidean_space set" assumes ind: "\ affine_dependent s" and dim: "DIM ('a) < card s" shows "affine hull s = UNIV" apply (rule affine_independent_span_eq [OF ind]) apply (rule antisym) using assms apply auto apply (metis add_2_eq_Suc' not_less_eq_eq affine_dependent_biggerset aff_independent_finite) done lemma empty_interior_affine_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" and dim: "card s \ DIM ('a)" shows "interior(affine hull s) = {}" using assms apply (induct s rule: finite_induct) apply (simp_all add: affine_dependent_iff_dependent affine_hull_insert_span_gen interior_translation) apply (rule empty_interior_lowdim) by (auto simp: Suc_le_lessD card_image_le dual_order.trans intro!: dim_le_card'[THEN le_less_trans]) lemma empty_interior_convex_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" and dim: "card s \ DIM ('a)" shows "interior(convex hull s) = {}" by (metis Diff_empty Diff_eq_empty_iff convex_hull_subset_affine_hull interior_mono empty_interior_affine_hull [OF assms]) lemma explicit_subset_rel_interior_convex_hull: fixes s :: "'a::euclidean_space set" shows "finite s \ {y. \u. (\x \ s. 0 < u x \ u x < 1) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y} \ rel_interior (convex hull s)" by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=s, simplified]) lemma explicit_subset_rel_interior_convex_hull_minimal: fixes s :: "'a::euclidean_space set" shows "finite s \ {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y} \ rel_interior (convex hull s)" by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=s, simplified]) lemma rel_interior_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "rel_interior(convex hull s) = {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" (is "?lhs = ?rhs") proof show "?rhs \ ?lhs" by (simp add: aff_independent_finite explicit_subset_rel_interior_convex_hull_minimal assms) next show "?lhs \ ?rhs" proof (cases "\a. s = {a}") case True then show "?lhs \ ?rhs" by force next case False have fs: "finite s" using assms by (simp add: aff_independent_finite) { fix a b and d::real assume ab: "a \ s" "b \ s" "a \ b" then have s: "s = (s - {a,b}) \ {a,b}" \ \split into separate cases\ by auto have "(\x\s. if x = a then - d else if x = b then d else 0) = 0" "(\x\s. (if x = a then - d else if x = b then d else 0) *\<^sub>R x) = d *\<^sub>R b - d *\<^sub>R a" using ab fs by (subst s, subst sum.union_disjoint, auto)+ } note [simp] = this { fix y assume y: "y \ convex hull s" "y \ ?rhs" { fix u T a assume ua: "\x\s. 0 \ u x" "sum u s = 1" "\ 0 < u a" "a \ s" and yT: "y = (\x\s. u x *\<^sub>R x)" "y \ T" "open T" and sb: "T \ affine hull s \ {w. \u. (\x\s. 0 \ u x) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = w}" have ua0: "u a = 0" using ua by auto obtain b where b: "b\s" "a \ b" using ua False by auto obtain e where e: "0 < e" "ball (\x\s. u x *\<^sub>R x) e \ T" using yT by (auto elim: openE) with b obtain d where d: "0 < d" "norm(d *\<^sub>R (a-b)) < e" by (auto intro: that [of "e / 2 / norm(a-b)"]) have "(\x\s. u x *\<^sub>R x) \ affine hull s" using yT y by (metis affine_hull_convex_hull hull_redundant_eq) then have "(\x\s. u x *\<^sub>R x) - d *\<^sub>R (a - b) \ affine hull s" using ua b by (auto simp: hull_inc intro: mem_affine_3_minus2) then have "y - d *\<^sub>R (a - b) \ T \ affine hull s" using d e yT by auto then obtain v where "\x\s. 0 \ v x" "sum v s = 1" "(\x\s. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x) - d *\<^sub>R (a - b)" using subsetD [OF sb] yT by auto then have False using assms apply (simp add: affine_dependent_explicit_finite fs) apply (drule_tac x="\x. (v x - u x) - (if x = a then -d else if x = b then d else 0)" in spec) using ua b d apply (auto simp: algebra_simps sum_subtractf sum.distrib) done } note * = this have "y \ rel_interior (convex hull s)" using y apply (simp add: mem_rel_interior) apply (auto simp: convex_hull_finite [OF fs]) apply (drule_tac x=u in spec) apply (auto intro: *) done } with rel_interior_subset show "?lhs \ ?rhs" by blast qed qed lemma interior_convex_hull_explicit_minimal: fixes s :: "'a::euclidean_space set" shows "\ affine_dependent s ==> interior(convex hull s) = (if card(s) \ DIM('a) then {} else {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = y})" apply (simp add: aff_independent_finite empty_interior_convex_hull, clarify) apply (rule trans [of _ "rel_interior(convex hull s)"]) apply (simp add: affine_independent_span_gt rel_interior_interior) by (simp add: rel_interior_convex_hull_explicit) lemma interior_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "interior(convex hull s) = (if card(s) \ DIM('a) then {} else {y. \u. (\x \ s. 0 < u x \ u x < 1) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = y})" proof - { fix u :: "'a \ real" and a assume "card Basis < card s" and u: "\x. x\s \ 0 < u x" "sum u s = 1" and a: "a \ s" then have cs: "Suc 0 < card s" by (metis DIM_positive less_trans_Suc) obtain b where b: "b \ s" "a \ b" proof (cases "s \ {a}") case True then show thesis using cs subset_singletonD by fastforce next case False then show thesis by (blast intro: that) qed have "u a + u b \ sum u {a,b}" using a b by simp also have "... \ sum u s" apply (rule Groups_Big.sum_mono2) using a b u apply (auto simp: less_imp_le aff_independent_finite assms) done finally have "u a < 1" using \b \ s\ u by fastforce } note [simp] = this show ?thesis using assms apply (auto simp: interior_convex_hull_explicit_minimal) apply (rule_tac x=u in exI) apply (auto simp: not_le) done qed lemma interior_closed_segment_ge2: fixes a :: "'a::euclidean_space" assumes "2 \ DIM('a)" shows "interior(closed_segment a b) = {}" using assms unfolding segment_convex_hull proof - have "card {a, b} \ DIM('a)" using assms by (simp add: card_insert_if linear not_less_eq_eq numeral_2_eq_2) then show "interior (convex hull {a, b}) = {}" by (metis empty_interior_convex_hull finite.insertI finite.emptyI) qed lemma interior_open_segment: fixes a :: "'a::euclidean_space" shows "interior(open_segment a b) = (if 2 \ DIM('a) then {} else open_segment a b)" proof (simp add: not_le, intro conjI impI) assume "2 \ DIM('a)" then show "interior (open_segment a b) = {}" apply (simp add: segment_convex_hull open_segment_def) apply (metis Diff_subset interior_mono segment_convex_hull subset_empty interior_closed_segment_ge2) done next assume le2: "DIM('a) < 2" show "interior (open_segment a b) = open_segment a b" proof (cases "a = b") case True then show ?thesis by auto next case False with le2 have "affine hull (open_segment a b) = UNIV" apply simp apply (rule affine_independent_span_gt) apply (simp_all add: affine_dependent_def insert_Diff_if) done then show "interior (open_segment a b) = open_segment a b" using rel_interior_interior rel_interior_open_segment by blast qed qed lemma interior_closed_segment: fixes a :: "'a::euclidean_space" shows "interior(closed_segment a b) = (if 2 \ DIM('a) then {} else open_segment a b)" proof (cases "a = b") case True then show ?thesis by simp next case False then have "closure (open_segment a b) = closed_segment a b" by simp then show ?thesis by (metis (no_types) convex_interior_closure convex_open_segment interior_open_segment) qed lemmas interior_segment = interior_closed_segment interior_open_segment lemma closed_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "closed_segment a b = closed_segment c d \ {a,b} = {c,d}" proof assume abcd: "closed_segment a b = closed_segment c d" show "{a,b} = {c,d}" proof (cases "a=b \ c=d") case True with abcd show ?thesis by force next case False then have neq: "a \ b \ c \ d" by force have *: "closed_segment c d - {a, b} = rel_interior (closed_segment c d)" using neq abcd by (metis (no_types) open_segment_def rel_interior_closed_segment) have "b \ {c, d}" proof - have "insert b (closed_segment c d) = closed_segment c d" using abcd by blast then show ?thesis by (metis DiffD2 Diff_insert2 False * insertI1 insert_Diff_if open_segment_def rel_interior_closed_segment) qed moreover have "a \ {c, d}" by (metis Diff_iff False * abcd ends_in_segment(1) insertI1 open_segment_def rel_interior_closed_segment) ultimately show "{a, b} = {c, d}" using neq by fastforce qed next assume "{a,b} = {c,d}" then show "closed_segment a b = closed_segment c d" by (simp add: segment_convex_hull) qed lemma closed_open_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "closed_segment a b \ open_segment c d" by (metis DiffE closed_segment_neq_empty closure_closed_segment closure_open_segment ends_in_segment(1) insertI1 open_segment_def) lemma open_closed_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "open_segment a b \ closed_segment c d" using closed_open_segment_eq by blast lemma open_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "open_segment a b = open_segment c d \ a = b \ c = d \ {a,b} = {c,d}" (is "?lhs = ?rhs") proof assume abcd: ?lhs show ?rhs proof (cases "a=b \ c=d") case True with abcd show ?thesis using finite_open_segment by fastforce next case False then have a2: "a \ b \ c \ d" by force with abcd show ?rhs unfolding open_segment_def by (metis (no_types) abcd closed_segment_eq closure_open_segment) qed next assume ?rhs then show ?lhs by (metis Diff_cancel convex_hull_singleton insert_absorb2 open_segment_def segment_convex_hull) qed subsection\<^marker>\tag unimportant\\Similar results for closure and (relative or absolute) frontier\ lemma closure_convex_hull [simp]: fixes s :: "'a::euclidean_space set" shows "compact s ==> closure(convex hull s) = convex hull s" by (simp add: compact_imp_closed compact_convex_hull) lemma rel_frontier_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "rel_frontier(convex hull s) = {y. \u. (\x \ s. 0 \ u x) \ (\x \ s. u x = 0) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" proof - have fs: "finite s" using assms by (simp add: aff_independent_finite) show ?thesis apply (simp add: rel_frontier_def finite_imp_compact rel_interior_convex_hull_explicit assms fs) apply (auto simp: convex_hull_finite fs) apply (drule_tac x=u in spec) apply (rule_tac x=u in exI) apply force apply (rename_tac v) apply (rule notE [OF assms]) apply (simp add: affine_dependent_explicit) apply (rule_tac x=s in exI) apply (auto simp: fs) apply (rule_tac x = "\x. u x - v x" in exI) apply (force simp: sum_subtractf scaleR_diff_left) done qed lemma frontier_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "frontier(convex hull s) = {y. \u. (\x \ s. 0 \ u x) \ (DIM ('a) < card s \ (\x \ s. u x = 0)) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" proof - have fs: "finite s" using assms by (simp add: aff_independent_finite) show ?thesis proof (cases "DIM ('a) < card s") case True with assms fs show ?thesis by (simp add: rel_frontier_def frontier_def rel_frontier_convex_hull_explicit [symmetric] interior_convex_hull_explicit_minimal rel_interior_convex_hull_explicit) next case False then have "card s \ DIM ('a)" by linarith then show ?thesis using assms fs apply (simp add: frontier_def interior_convex_hull_explicit finite_imp_compact) apply (simp add: convex_hull_finite) done qed qed lemma rel_frontier_convex_hull_cases: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "rel_frontier(convex hull s) = \{convex hull (s - {x}) |x. x \ s}" proof - have fs: "finite s" using assms by (simp add: aff_independent_finite) { fix u a have "\x\s. 0 \ u x \ a \ s \ u a = 0 \ sum u s = 1 \ \x v. x \ s \ (\x\s - {x}. 0 \ v x) \ sum v (s - {x}) = 1 \ (\x\s - {x}. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x)" apply (rule_tac x=a in exI) apply (rule_tac x=u in exI) apply (simp add: Groups_Big.sum_diff1 fs) done } moreover { fix a u have "a \ s \ \x\s - {a}. 0 \ u x \ sum u (s - {a}) = 1 \ \v. (\x\s. 0 \ v x) \ (\x\s. v x = 0) \ sum v s = 1 \ (\x\s. v x *\<^sub>R x) = (\x\s - {a}. u x *\<^sub>R x)" apply (rule_tac x="\x. if x = a then 0 else u x" in exI) apply (auto simp: sum.If_cases Diff_eq if_smult fs) done } ultimately show ?thesis using assms apply (simp add: rel_frontier_convex_hull_explicit) apply (simp add: convex_hull_finite fs Union_SetCompr_eq, auto) done qed lemma frontier_convex_hull_eq_rel_frontier: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "frontier(convex hull s) = (if card s \ DIM ('a) then convex hull s else rel_frontier(convex hull s))" using assms unfolding rel_frontier_def frontier_def by (simp add: affine_independent_span_gt rel_interior_interior finite_imp_compact empty_interior_convex_hull aff_independent_finite) lemma frontier_convex_hull_cases: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "frontier(convex hull s) = (if card s \ DIM ('a) then convex hull s else \{convex hull (s - {x}) |x. x \ s})" by (simp add: assms frontier_convex_hull_eq_rel_frontier rel_frontier_convex_hull_cases) lemma in_frontier_convex_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" "card s \ Suc (DIM ('a))" "x \ s" shows "x \ frontier(convex hull s)" proof (cases "affine_dependent s") case True with assms show ?thesis apply (auto simp: affine_dependent_def frontier_def finite_imp_compact hull_inc) by (metis card.insert_remove convex_hull_subset_affine_hull empty_interior_affine_hull finite_Diff hull_redundant insert_Diff insert_Diff_single insert_not_empty interior_mono not_less_eq_eq subset_empty) next case False { assume "card s = Suc (card Basis)" then have cs: "Suc 0 < card s" by (simp) with subset_singletonD have "\y \ s. y \ x" by (cases "s \ {x}") fastforce+ } note [dest!] = this show ?thesis using assms unfolding frontier_convex_hull_cases [OF False] Union_SetCompr_eq by (auto simp: le_Suc_eq hull_inc) qed lemma not_in_interior_convex_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" "card s \ Suc (DIM ('a))" "x \ s" shows "x \ interior(convex hull s)" using in_frontier_convex_hull [OF assms] by (metis Diff_iff frontier_def) lemma interior_convex_hull_eq_empty: fixes s :: "'a::euclidean_space set" assumes "card s = Suc (DIM ('a))" shows "interior(convex hull s) = {} \ affine_dependent s" proof - { fix a b assume ab: "a \ interior (convex hull s)" "b \ s" "b \ affine hull (s - {b})" then have "interior(affine hull s) = {}" using assms by (metis DIM_positive One_nat_def Suc_mono card.remove card_infinite empty_interior_affine_hull eq_iff hull_redundant insert_Diff not_less zero_le_one) then have False using ab by (metis convex_hull_subset_affine_hull equals0D interior_mono subset_eq) } then show ?thesis using assms apply auto apply (metis UNIV_I affine_hull_convex_hull affine_hull_empty affine_independent_span_eq convex_convex_hull empty_iff rel_interior_interior rel_interior_same_affine_hull) apply (auto simp: affine_dependent_def) done qed subsection \Coplanarity, and collinearity in terms of affine hull\ definition\<^marker>\tag important\ coplanar where "coplanar s \ \u v w. s \ affine hull {u,v,w}" lemma collinear_affine_hull: "collinear s \ (\u v. s \ affine hull {u,v})" proof (cases "s={}") case True then show ?thesis by simp next case False then obtain x where x: "x \ s" by auto { fix u assume *: "\x y. \x\s; y\s\ \ \c. x - y = c *\<^sub>R u" have "\u v. s \ {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" apply (rule_tac x=x in exI) apply (rule_tac x="x+u" in exI, clarify) apply (erule exE [OF * [OF x]]) apply (rename_tac c) apply (rule_tac x="1+c" in exI) apply (rule_tac x="-c" in exI) apply (simp add: algebra_simps) done } moreover { fix u v x y assume *: "s \ {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" have "x\s \ y\s \ \c. x - y = c *\<^sub>R (v-u)" apply (drule subsetD [OF *])+ apply simp apply clarify apply (rename_tac r1 r2) apply (rule_tac x="r1-r2" in exI) apply (simp add: algebra_simps) apply (metis scaleR_left.add) done } ultimately show ?thesis unfolding collinear_def affine_hull_2 by blast qed lemma collinear_closed_segment [simp]: "collinear (closed_segment a b)" by (metis affine_hull_convex_hull collinear_affine_hull hull_subset segment_convex_hull) lemma collinear_open_segment [simp]: "collinear (open_segment a b)" unfolding open_segment_def by (metis convex_hull_subset_affine_hull segment_convex_hull dual_order.trans convex_hull_subset_affine_hull Diff_subset collinear_affine_hull) lemma collinear_between_cases: fixes c :: "'a::euclidean_space" shows "collinear {a,b,c} \ between (b,c) a \ between (c,a) b \ between (a,b) c" (is "?lhs = ?rhs") proof assume ?lhs then obtain u v where uv: "\x. x \ {a, b, c} \ \c. x = u + c *\<^sub>R v" by (auto simp: collinear_alt) show ?rhs using uv [of a] uv [of b] uv [of c] by (auto simp: between_1) next assume ?rhs then show ?lhs unfolding between_mem_convex_hull by (metis (no_types, hide_lams) collinear_closed_segment collinear_subset hull_redundant hull_subset insert_commute segment_convex_hull) qed lemma subset_continuous_image_segment_1: fixes f :: "'a::euclidean_space \ real" assumes "continuous_on (closed_segment a b) f" shows "closed_segment (f a) (f b) \ image f (closed_segment a b)" by (metis connected_segment convex_contains_segment ends_in_segment imageI is_interval_connected_1 is_interval_convex connected_continuous_image [OF assms]) lemma continuous_injective_image_segment_1: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and injf: "inj_on f (closed_segment a b)" shows "f ` (closed_segment a b) = closed_segment (f a) (f b)" proof show "closed_segment (f a) (f b) \ f ` closed_segment a b" by (metis subset_continuous_image_segment_1 contf) show "f ` closed_segment a b \ closed_segment (f a) (f b)" proof (cases "a = b") case True then show ?thesis by auto next case False then have fnot: "f a \ f b" using inj_onD injf by fastforce moreover have "f a \ open_segment (f c) (f b)" if c: "c \ closed_segment a b" for c proof (clarsimp simp add: open_segment_def) assume fa: "f a \ closed_segment (f c) (f b)" moreover have "closed_segment (f c) (f b) \ f ` closed_segment c b" by (meson closed_segment_subset contf continuous_on_subset convex_closed_segment ends_in_segment(2) subset_continuous_image_segment_1 that) ultimately have "f a \ f ` closed_segment c b" by blast then have a: "a \ closed_segment c b" by (meson ends_in_segment inj_on_image_mem_iff_alt injf subset_closed_segment that) have cb: "closed_segment c b \ closed_segment a b" by (simp add: closed_segment_subset that) show "f a = f c" proof (rule between_antisym) show "between (f c, f b) (f a)" by (simp add: between_mem_segment fa) show "between (f a, f b) (f c)" by (metis a cb between_antisym between_mem_segment between_triv1 subset_iff) qed qed moreover have "f b \ open_segment (f a) (f c)" if c: "c \ closed_segment a b" for c proof (clarsimp simp add: open_segment_def fnot eq_commute) assume fb: "f b \ closed_segment (f a) (f c)" moreover have "closed_segment (f a) (f c) \ f ` closed_segment a c" by (meson contf continuous_on_subset ends_in_segment(1) subset_closed_segment subset_continuous_image_segment_1 that) ultimately have "f b \ f ` closed_segment a c" by blast then have b: "b \ closed_segment a c" by (meson ends_in_segment inj_on_image_mem_iff_alt injf subset_closed_segment that) have ca: "closed_segment a c \ closed_segment a b" by (simp add: closed_segment_subset that) show "f b = f c" proof (rule between_antisym) show "between (f c, f a) (f b)" by (simp add: between_commute between_mem_segment fb) show "between (f b, f a) (f c)" by (metis b between_antisym between_commute between_mem_segment between_triv2 that) qed qed ultimately show ?thesis by (force simp: closed_segment_eq_real_ivl open_segment_eq_real_ivl split: if_split_asm) qed qed lemma continuous_injective_image_open_segment_1: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and injf: "inj_on f (closed_segment a b)" shows "f ` (open_segment a b) = open_segment (f a) (f b)" proof - have "f ` (open_segment a b) = f ` (closed_segment a b) - {f a, f b}" by (metis (no_types, hide_lams) empty_subsetI ends_in_segment image_insert image_is_empty inj_on_image_set_diff injf insert_subset open_segment_def segment_open_subset_closed) also have "... = open_segment (f a) (f b)" using continuous_injective_image_segment_1 [OF assms] by (simp add: open_segment_def inj_on_image_set_diff [OF injf]) finally show ?thesis . qed lemma collinear_imp_coplanar: "collinear s ==> coplanar s" by (metis collinear_affine_hull coplanar_def insert_absorb2) lemma collinear_small: assumes "finite s" "card s \ 2" shows "collinear s" proof - have "card s = 0 \ card s = 1 \ card s = 2" using assms by linarith then show ?thesis using assms using card_eq_SucD by auto (metis collinear_2 numeral_2_eq_2) qed lemma coplanar_small: assumes "finite s" "card s \ 3" shows "coplanar s" proof - have "card s \ 2 \ card s = Suc (Suc (Suc 0))" using assms by linarith then show ?thesis using assms apply safe apply (simp add: collinear_small collinear_imp_coplanar) apply (safe dest!: card_eq_SucD) apply (auto simp: coplanar_def) apply (metis hull_subset insert_subset) done qed lemma coplanar_empty: "coplanar {}" by (simp add: coplanar_small) lemma coplanar_sing: "coplanar {a}" by (simp add: coplanar_small) lemma coplanar_2: "coplanar {a,b}" by (auto simp: card_insert_if coplanar_small) lemma coplanar_3: "coplanar {a,b,c}" by (auto simp: card_insert_if coplanar_small) lemma collinear_affine_hull_collinear: "collinear(affine hull s) \ collinear s" unfolding collinear_affine_hull by (metis affine_affine_hull subset_hull hull_hull hull_mono) lemma coplanar_affine_hull_coplanar: "coplanar(affine hull s) \ coplanar s" unfolding coplanar_def by (metis affine_affine_hull subset_hull hull_hull hull_mono) lemma coplanar_linear_image: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "coplanar s" "linear f" shows "coplanar(f ` s)" proof - { fix u v w assume "s \ affine hull {u, v, w}" then have "f ` s \ f ` (affine hull {u, v, w})" by (simp add: image_mono) then have "f ` s \ affine hull (f ` {u, v, w})" by (metis assms(2) linear_conv_bounded_linear affine_hull_linear_image) } then show ?thesis by auto (meson assms(1) coplanar_def) qed lemma coplanar_translation_imp: "coplanar s \ coplanar ((\x. a + x) ` s)" unfolding coplanar_def apply clarify apply (rule_tac x="u+a" in exI) apply (rule_tac x="v+a" in exI) apply (rule_tac x="w+a" in exI) using affine_hull_translation [of a "{u,v,w}" for u v w] apply (force simp: add.commute) done lemma coplanar_translation_eq: "coplanar((\x. a + x) ` s) \ coplanar s" by (metis (no_types) coplanar_translation_imp translation_galois) lemma coplanar_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "coplanar(f ` s) = coplanar s" proof assume "coplanar s" then show "coplanar (f ` s)" unfolding coplanar_def using affine_hull_linear_image [of f "{u,v,w}" for u v w] assms by (meson coplanar_def coplanar_linear_image) next obtain g where g: "linear g" "g \ f = id" using linear_injective_left_inverse [OF assms] by blast assume "coplanar (f ` s)" then obtain u v w where "f ` s \ affine hull {u, v, w}" by (auto simp: coplanar_def) then have "g ` f ` s \ g ` (affine hull {u, v, w})" by blast then have "s \ g ` (affine hull {u, v, w})" using g by (simp add: Fun.image_comp) then show "coplanar s" unfolding coplanar_def using affine_hull_linear_image [of g "{u,v,w}" for u v w] \linear g\ linear_conv_bounded_linear by fastforce qed (*The HOL Light proof is simply MATCH_ACCEPT_TAC(LINEAR_INVARIANT_RULE COPLANAR_LINEAR_IMAGE));; *) lemma coplanar_subset: "\coplanar t; s \ t\ \ coplanar s" by (meson coplanar_def order_trans) lemma affine_hull_3_imp_collinear: "c \ affine hull {a,b} \ collinear {a,b,c}" by (metis collinear_2 collinear_affine_hull_collinear hull_redundant insert_commute) lemma collinear_3_imp_in_affine_hull: "\collinear {a,b,c}; a \ b\ \ c \ affine hull {a,b}" unfolding collinear_def apply clarify apply (frule_tac x=b in bspec, blast, drule_tac x=a in bspec, blast, erule exE) apply (drule_tac x=c in bspec, blast, drule_tac x=a in bspec, blast, erule exE) apply (rename_tac y x) apply (simp add: affine_hull_2) apply (rule_tac x="1 - x/y" in exI) apply (simp add: algebra_simps) done lemma collinear_3_affine_hull: assumes "a \ b" shows "collinear {a,b,c} \ c \ affine hull {a,b}" using affine_hull_3_imp_collinear assms collinear_3_imp_in_affine_hull by blast lemma collinear_3_eq_affine_dependent: "collinear{a,b,c} \ a = b \ a = c \ b = c \ affine_dependent {a,b,c}" apply (case_tac "a=b", simp) apply (case_tac "a=c") apply (simp add: insert_commute) apply (case_tac "b=c") apply (simp add: insert_commute) apply (auto simp: affine_dependent_def collinear_3_affine_hull insert_Diff_if) apply (metis collinear_3_affine_hull insert_commute)+ done lemma affine_dependent_imp_collinear_3: "affine_dependent {a,b,c} \ collinear{a,b,c}" by (simp add: collinear_3_eq_affine_dependent) lemma collinear_3: "NO_MATCH 0 x \ collinear {x,y,z} \ collinear {0, x-y, z-y}" by (auto simp add: collinear_def) lemma collinear_3_expand: "collinear{a,b,c} \ a = c \ (\u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" proof - have "collinear{a,b,c} = collinear{a,c,b}" by (simp add: insert_commute) also have "... = collinear {0, a - c, b - c}" by (simp add: collinear_3) also have "... \ (a = c \ b = c \ (\ca. b - c = ca *\<^sub>R (a - c)))" by (simp add: collinear_lemma) also have "... \ a = c \ (\u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" by (cases "a = c \ b = c") (auto simp: algebra_simps) finally show ?thesis . qed lemma collinear_aff_dim: "collinear S \ aff_dim S \ 1" proof assume "collinear S" then obtain u and v :: "'a" where "aff_dim S \ aff_dim {u,v}" by (metis \collinear S\ aff_dim_affine_hull aff_dim_subset collinear_affine_hull) then show "aff_dim S \ 1" using order_trans by fastforce next assume "aff_dim S \ 1" then have le1: "aff_dim (affine hull S) \ 1" by simp obtain B where "B \ S" and B: "\ affine_dependent B" "affine hull S = affine hull B" using affine_basis_exists [of S] by auto then have "finite B" "card B \ 2" using B le1 by (auto simp: affine_independent_iff_card) then have "collinear B" by (rule collinear_small) then show "collinear S" by (metis \affine hull S = affine hull B\ collinear_affine_hull_collinear) qed lemma collinear_midpoint: "collinear{a,midpoint a b,b}" apply (auto simp: collinear_3 collinear_lemma) apply (drule_tac x="-1" in spec) apply (simp add: algebra_simps) done lemma midpoint_collinear: fixes a b c :: "'a::real_normed_vector" assumes "a \ c" shows "b = midpoint a c \ collinear{a,b,c} \ dist a b = dist b c" proof - have *: "a - (u *\<^sub>R a + (1 - u) *\<^sub>R c) = (1 - u) *\<^sub>R (a - c)" "u *\<^sub>R a + (1 - u) *\<^sub>R c - c = u *\<^sub>R (a - c)" "\1 - u\ = \u\ \ u = 1/2" for u::real by (auto simp: algebra_simps) have "b = midpoint a c \ collinear{a,b,c} " using collinear_midpoint by blast moreover have "collinear{a,b,c} \ b = midpoint a c \ dist a b = dist b c" apply (auto simp: collinear_3_expand assms dist_midpoint) apply (simp add: dist_norm * assms midpoint_def del: divide_const_simps) apply (simp add: algebra_simps) done ultimately show ?thesis by blast qed lemma between_imp_collinear: fixes x :: "'a :: euclidean_space" assumes "between (a,b) x" shows "collinear {a,x,b}" proof (cases "x = a \ x = b \ a = b") case True with assms show ?thesis by (auto simp: dist_commute) next case False with assms show ?thesis apply (auto simp: collinear_3 collinear_lemma between_norm) apply (drule_tac x="-(norm(b - x) / norm(x - a))" in spec) apply (simp add: vector_add_divide_simps real_vector.scale_minus_right [symmetric]) done qed lemma midpoint_between: fixes a b :: "'a::euclidean_space" shows "b = midpoint a c \ between (a,c) b \ dist a b = dist b c" proof (cases "a = c") case True then show ?thesis by (auto simp: dist_commute) next case False show ?thesis apply (rule iffI) apply (simp add: between_midpoint(1) dist_midpoint) using False between_imp_collinear midpoint_collinear by blast qed lemma collinear_triples: assumes "a \ b" shows "collinear(insert a (insert b S)) \ (\x \ S. collinear{a,b,x})" (is "?lhs = ?rhs") proof safe fix x assume ?lhs and "x \ S" then show "collinear {a, b, x}" using collinear_subset by force next assume ?rhs then have "\x \ S. collinear{a,x,b}" by (simp add: insert_commute) then have *: "\u. x = u *\<^sub>R a + (1 - u) *\<^sub>R b" if "x \ (insert a (insert b S))" for x using that assms collinear_3_expand by fastforce+ show ?lhs unfolding collinear_def apply (rule_tac x="b-a" in exI) apply (clarify dest!: *) by (metis (no_types, hide_lams) add.commute diff_add_cancel diff_diff_eq2 real_vector.scale_right_diff_distrib scaleR_left.diff) qed lemma collinear_4_3: assumes "a \ b" shows "collinear {a,b,c,d} \ collinear{a,b,c} \ collinear{a,b,d}" using collinear_triples [OF assms, of "{c,d}"] by (force simp:) lemma collinear_3_trans: assumes "collinear{a,b,c}" "collinear{b,c,d}" "b \ c" shows "collinear{a,b,d}" proof - have "collinear{b,c,a,d}" by (metis (full_types) assms collinear_4_3 insert_commute) then show ?thesis by (simp add: collinear_subset) qed lemma affine_hull_eq_empty [simp]: "affine hull S = {} \ S = {}" using affine_hull_nonempty by blast lemma affine_hull_2_alt: fixes a b :: "'a::real_vector" shows "affine hull {a,b} = range (\u. a + u *\<^sub>R (b - a))" apply (simp add: affine_hull_2, safe) apply (rule_tac x=v in image_eqI) apply (simp add: algebra_simps) apply (metis scaleR_add_left scaleR_one, simp) apply (rule_tac x="1-u" in exI) apply (simp add: algebra_simps) done lemma interior_convex_hull_3_minimal: fixes a :: "'a::euclidean_space" shows "\\ collinear{a,b,c}; DIM('a) = 2\ \ interior(convex hull {a,b,c}) = {v. \x y z. 0 < x \ 0 < y \ 0 < z \ x + y + z = 1 \ x *\<^sub>R a + y *\<^sub>R b + z *\<^sub>R c = v}" apply (simp add: collinear_3_eq_affine_dependent interior_convex_hull_explicit_minimal, safe) apply (rule_tac x="u a" in exI, simp) apply (rule_tac x="u b" in exI, simp) apply (rule_tac x="u c" in exI, simp) apply (rename_tac uu x y z) apply (rule_tac x="\r. (if r=a then x else if r=b then y else if r=c then z else 0)" in exI) apply simp done subsection\<^marker>\tag unimportant\\Basic lemmas about hyperplanes and halfspaces\ lemma halfspace_Int_eq: "{x. a \ x \ b} \ {x. b \ a \ x} = {x. a \ x = b}" "{x. b \ a \ x} \ {x. a \ x \ b} = {x. a \ x = b}" by auto lemma hyperplane_eq_Ex: assumes "a \ 0" obtains x where "a \ x = b" by (rule_tac x = "(b / (a \ a)) *\<^sub>R a" in that) (simp add: assms) lemma hyperplane_eq_empty: "{x. a \ x = b} = {} \ a = 0 \ b \ 0" using hyperplane_eq_Ex apply auto[1] using inner_zero_right by blast lemma hyperplane_eq_UNIV: "{x. a \ x = b} = UNIV \ a = 0 \ b = 0" proof - have "UNIV \ {x. a \ x = b} \ a = 0 \ b = 0" apply (drule_tac c = "((b+1) / (a \ a)) *\<^sub>R a" in subsetD) apply simp_all by (metis add_cancel_right_right zero_neq_one) then show ?thesis by force qed lemma halfspace_eq_empty_lt: "{x. a \ x < b} = {} \ a = 0 \ b \ 0" proof - have "{x. a \ x < b} \ {} \ a = 0 \ b \ 0" apply (rule ccontr) apply (drule_tac c = "((b-1) / (a \ a)) *\<^sub>R a" in subsetD) apply force+ done then show ?thesis by force qed lemma halfspace_eq_empty_gt: "{x. a \ x > b} = {} \ a = 0 \ b \ 0" using halfspace_eq_empty_lt [of "-a" "-b"] by simp lemma halfspace_eq_empty_le: "{x. a \ x \ b} = {} \ a = 0 \ b < 0" proof - have "{x. a \ x \ b} \ {} \ a = 0 \ b < 0" apply (rule ccontr) apply (drule_tac c = "((b-1) / (a \ a)) *\<^sub>R a" in subsetD) apply force+ done then show ?thesis by force qed lemma halfspace_eq_empty_ge: "{x. a \ x \ b} = {} \ a = 0 \ b > 0" using halfspace_eq_empty_le [of "-a" "-b"] by simp subsection\<^marker>\tag unimportant\\Use set distance for an easy proof of separation properties\ proposition\<^marker>\tag unimportant\ separation_closures: fixes S :: "'a::euclidean_space set" assumes "S \ closure T = {}" "T \ closure S = {}" obtains U V where "U \ V = {}" "open U" "open V" "S \ U" "T \ V" proof (cases "S = {} \ T = {}") case True with that show ?thesis by auto next case False define f where "f \ \x. setdist {x} T - setdist {x} S" have contf: "continuous_on UNIV f" unfolding f_def by (intro continuous_intros continuous_on_setdist) show ?thesis proof (rule_tac U = "{x. f x > 0}" and V = "{x. f x < 0}" in that) show "{x. 0 < f x} \ {x. f x < 0} = {}" by auto show "open {x. 0 < f x}" by (simp add: open_Collect_less contf) show "open {x. f x < 0}" by (simp add: open_Collect_less contf) show "S \ {x. 0 < f x}" apply (clarsimp simp add: f_def setdist_sing_in_set) using assms by (metis False IntI empty_iff le_less setdist_eq_0_sing_2 setdist_pos_le setdist_sym) show "T \ {x. f x < 0}" apply (clarsimp simp add: f_def setdist_sing_in_set) using assms by (metis False IntI empty_iff le_less setdist_eq_0_sing_2 setdist_pos_le setdist_sym) qed qed lemma separation_normal: fixes S :: "'a::euclidean_space set" assumes "closed S" "closed T" "S \ T = {}" obtains U V where "open U" "open V" "S \ U" "T \ V" "U \ V = {}" using separation_closures [of S T] by (metis assms closure_closed disjnt_def inf_commute) lemma separation_normal_local: fixes S :: "'a::euclidean_space set" assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "S \ T = {}" obtains S' T' where "openin (top_of_set U) S'" "openin (top_of_set U) T'" "S \ S'" "T \ T'" "S' \ T' = {}" proof (cases "S = {} \ T = {}") case True with that show ?thesis using UT US by (blast dest: closedin_subset) next case False define f where "f \ \x. setdist {x} T - setdist {x} S" have contf: "continuous_on U f" unfolding f_def by (intro continuous_intros) show ?thesis proof (rule_tac S' = "(U \ f -` {0<..})" and T' = "(U \ f -` {..<0})" in that) show "(U \ f -` {0<..}) \ (U \ f -` {..<0}) = {}" by auto show "openin (top_of_set U) (U \ f -` {0<..})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next show "openin (top_of_set U) (U \ f -` {..<0})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next have "S \ U" "T \ U" using closedin_imp_subset assms by blast+ then show "S \ U \ f -` {0<..}" "T \ U \ f -` {..<0}" using assms False by (force simp add: f_def setdist_sing_in_set intro!: setdist_gt_0_closedin)+ qed qed lemma separation_normal_compact: fixes S :: "'a::euclidean_space set" assumes "compact S" "closed T" "S \ T = {}" obtains U V where "open U" "compact(closure U)" "open V" "S \ U" "T \ V" "U \ V = {}" proof - have "closed S" "bounded S" using assms by (auto simp: compact_eq_bounded_closed) then obtain r where "r>0" and r: "S \ ball 0 r" by (auto dest!: bounded_subset_ballD) have **: "closed (T \ - ball 0 r)" "S \ (T \ - ball 0 r) = {}" using assms r by blast+ then show ?thesis apply (rule separation_normal [OF \closed S\]) apply (rule_tac U=U and V=V in that) by auto (meson bounded_ball bounded_subset compl_le_swap2 disjoint_eq_subset_Compl) qed subsection\Connectedness of the intersection of a chain\ proposition connected_chain: fixes \ :: "'a :: euclidean_space set set" assumes cc: "\S. S \ \ \ compact S \ connected S" and linear: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "connected(\\)" proof (cases "\ = {}") case True then show ?thesis by auto next case False then have cf: "compact(\\)" by (simp add: cc compact_Inter) have False if AB: "closed A" "closed B" "A \ B = {}" and ABeq: "A \ B = \\" and "A \ {}" "B \ {}" for A B proof - obtain U V where "open U" "open V" "A \ U" "B \ V" "U \ V = {}" using separation_normal [OF AB] by metis obtain K where "K \ \" "compact K" using cc False by blast then obtain N where "open N" and "K \ N" by blast let ?\ = "insert (U \ V) ((\S. N - S) ` \)" obtain \ where "\ \ ?\" "finite \" "K \ \\" proof (rule compactE [OF \compact K\]) show "K \ \(insert (U \ V) ((-) N ` \))" using \K \ N\ ABeq \A \ U\ \B \ V\ by auto show "\B. B \ insert (U \ V) ((-) N ` \) \ open B" by (auto simp: \open U\ \open V\ open_Un \open N\ cc compact_imp_closed open_Diff) qed then have "finite(\ - {U \ V})" by blast moreover have "\ - {U \ V} \ (\S. N - S) ` \" using \\ \ ?\\ by blast ultimately obtain \ where "\ \ \" "finite \" and Deq: "\ - {U \ V} = (\S. N-S) ` \" using finite_subset_image by metis obtain J where "J \ \" and J: "(\S\\. N - S) \ N - J" proof (cases "\ = {}") case True with \\ \ {}\ that show ?thesis by auto next case False have "\S T. \S \ \; T \ \\ \ S \ T \ T \ S" by (meson \\ \ \\ in_mono local.linear) with \finite \\ \\ \ {}\ have "\J \ \. (\S\\. N - S) \ N - J" proof induction case (insert X \) show ?case proof (cases "\ = {}") case True then show ?thesis by auto next case False then have "\S T. \S \ \; T \ \\ \ S \ T \ T \ S" by (simp add: insert.prems) with insert.IH False obtain J where "J \ \" and J: "(\Y\\. N - Y) \ N - J" by metis have "N - J \ N - X \ N - X \ N - J" by (meson Diff_mono \J \ \\ insert.prems(2) insert_iff order_refl) then show ?thesis proof assume "N - J \ N - X" with J show ?thesis by auto next assume "N - X \ N - J" with J have "N - X \ \ ((-) N ` \) \ N - J" by auto with \J \ \\ show ?thesis by blast qed qed qed simp with \\ \ \\ show ?thesis by (blast intro: that) qed have "K \ \(insert (U \ V) (\ - {U \ V}))" using \K \ \\\ by auto also have "... \ (U \ V) \ (N - J)" by (metis (no_types, hide_lams) Deq Un_subset_iff Un_upper2 J Union_insert order_trans sup_ge1) finally have "J \ K \ U \ V" by blast moreover have "connected(J \ K)" by (metis Int_absorb1 \J \ \\ \K \ \\ cc inf.orderE local.linear) moreover have "U \ (J \ K) \ {}" using ABeq \J \ \\ \K \ \\ \A \ {}\ \A \ U\ by blast moreover have "V \ (J \ K) \ {}" using ABeq \J \ \\ \K \ \\ \B \ {}\ \B \ V\ by blast ultimately show False using connectedD [of "J \ K" U V] \open U\ \open V\ \U \ V = {}\ by auto qed with cf show ?thesis by (auto simp: connected_closed_set compact_imp_closed) qed lemma connected_chain_gen: fixes \ :: "'a :: euclidean_space set set" assumes X: "X \ \" "compact X" and cc: "\T. T \ \ \ closed T \ connected T" and linear: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "connected(\\)" proof - have "\\ = (\T\\. X \ T)" using X by blast moreover have "connected (\T\\. X \ T)" proof (rule connected_chain) show "\T. T \ (\) X ` \ \ compact T \ connected T" using cc X by auto (metis inf.absorb2 inf.orderE local.linear) show "\S T. S \ (\) X ` \ \ T \ (\) X ` \ \ S \ T \ T \ S" using local.linear by blast qed ultimately show ?thesis by metis qed lemma connected_nest: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. compact(S n)" "\n. connected(S n)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" apply (rule connected_chain) using S apply blast by (metis image_iff le_cases nest) lemma connected_nest_gen: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. closed(S n)" "\n. connected(S n)" "compact(S k)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" apply (rule connected_chain_gen [of "S k"]) using S apply auto by (meson le_cases nest subsetCE) subsection\Proper maps, including projections out of compact sets\ lemma finite_indexed_bound: assumes A: "finite A" "\x. x \ A \ \n::'a::linorder. P x n" shows "\m. \x \ A. \k\m. P x k" using A proof (induction A) case empty then show ?case by force next case (insert a A) then obtain m n where "\x \ A. \k\m. P x k" "P a n" by force then show ?case apply (rule_tac x="max m n" in exI, safe) using max.cobounded2 apply blast by (meson le_max_iff_disj) qed proposition proper_map: fixes f :: "'a::heine_borel \ 'b::heine_borel" assumes "closedin (top_of_set S) K" and com: "\U. \U \ T; compact U\ \ compact (S \ f -` U)" and "f ` S \ T" shows "closedin (top_of_set T) (f ` K)" proof - have "K \ S" using assms closedin_imp_subset by metis obtain C where "closed C" and Keq: "K = S \ C" using assms by (auto simp: closedin_closed) have *: "y \ f ` K" if "y \ T" and y: "y islimpt f ` K" for y proof - obtain h where "\n. (\x\K. h n = f x) \ h n \ y" "inj h" and hlim: "(h \ y) sequentially" using \y \ T\ y by (force simp: limpt_sequential_inj) then obtain X where X: "\n. X n \ K \ h n = f (X n) \ h n \ y" by metis then have fX: "\n. f (X n) = h n" by metis have "compact (C \ (S \ f -` insert y (range (\i. f(X(n + i))))))" for n apply (rule closed_Int_compact [OF \closed C\]) apply (rule com) using X \K \ S\ \f ` S \ T\ \y \ T\ apply blast apply (rule compact_sequence_with_limit) apply (simp add: fX add.commute [of n] LIMSEQ_ignore_initial_segment [OF hlim]) done then have comf: "compact {a \ K. f a \ insert y (range (\i. f(X(n + i))))}" for n by (simp add: Keq Int_def conj_commute) have ne: "\\ \ {}" if "finite \" and \: "\t. t \ \ \ (\n. t = {a \ K. f a \ insert y (range (\i. f (X (n + i))))})" for \ proof - obtain m where m: "\t. t \ \ \ \k\m. t = {a \ K. f a \ insert y (range (\i. f (X (k + i))))}" apply (rule exE) apply (rule finite_indexed_bound [OF \finite \\ \], assumption, force) done have "X m \ \\" using X le_Suc_ex by (fastforce dest: m) then show ?thesis by blast qed have "\{{a. a \ K \ f a \ insert y (range (\i. f(X(n + i))))} |n. n \ UNIV} \ {}" apply (rule compact_fip_Heine_Borel) using comf apply force using ne apply (simp add: subset_iff del: insert_iff) done then have "\x. x \ (\n. {a \ K. f a \ insert y (range (\i. f (X (n + i))))})" by blast then show ?thesis apply (simp add: image_iff fX) by (metis \inj h\ le_add1 not_less_eq_eq rangeI range_ex1_eq) qed with assms closedin_subset show ?thesis by (force simp: closedin_limpt) qed lemma compact_continuous_image_eq: fixes f :: "'a::heine_borel \ 'b::heine_borel" assumes f: "inj_on f S" shows "continuous_on S f \ (\T. compact T \ T \ S \ compact(f ` T))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis continuous_on_subset compact_continuous_image) next assume RHS: ?rhs obtain g where gf: "\x. x \ S \ g (f x) = x" by (metis inv_into_f_f f) then have *: "(S \ f -` U) = g ` U" if "U \ f ` S" for U using that by fastforce have gfim: "g ` f ` S \ S" using gf by auto have **: "compact (f ` S \ g -` C)" if C: "C \ S" "compact C" for C proof - obtain h where "h C \ C \ h C \ S \ compact (f ` C)" by (force simp: C RHS) moreover have "f ` C = (f ` S \ g -` C)" using C gf by auto ultimately show ?thesis using C by auto qed show ?lhs using proper_map [OF _ _ gfim] ** by (simp add: continuous_on_closed * closedin_imp_subset) qed subsection\<^marker>\tag unimportant\\Trivial fact: convexity equals connectedness for collinear sets\ lemma convex_connected_collinear: fixes S :: "'a::euclidean_space set" assumes "collinear S" shows "convex S \ connected S" proof assume "convex S" then show "connected S" using convex_connected by blast next assume S: "connected S" show "convex S" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain a where "a \ S" by auto have "collinear (affine hull S)" by (simp add: assms collinear_affine_hull_collinear) then obtain z where "z \ 0" "\x. x \ affine hull S \ \c. x - a = c *\<^sub>R z" by (meson \a \ S\ collinear hull_inc) then obtain f where f: "\x. x \ affine hull S \ x - a = f x *\<^sub>R z" by metis then have inj_f: "inj_on f (affine hull S)" by (metis diff_add_cancel inj_onI) have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \ affine hull S" and y: "y \ affine hull S" for x y proof - have "f x *\<^sub>R z = x - a" by (simp add: f hull_inc x) moreover have "f y *\<^sub>R z = y - a" by (simp add: f hull_inc y) ultimately show ?thesis by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" apply (clarsimp simp: dist_norm continuous_on_iff diff) by (metis \z \ 0\ mult.commute mult_less_cancel_left_pos norm_minus_commute real_norm_def zero_less_mult_iff zero_less_norm_iff) then have conn_fS: "connected (f ` S)" by (meson S connected_continuous_image continuous_on_subset hull_subset) show ?thesis proof (clarsimp simp: convex_contains_segment) fix x y z assume "x \ S" "y \ S" "z \ closed_segment x y" have False if "z \ S" proof - have "f ` (closed_segment x y) = closed_segment (f x) (f y)" apply (rule continuous_injective_image_segment_1) apply (meson \x \ S\ \y \ S\ convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f]) by (meson \x \ S\ \y \ S\ convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f]) then have fz: "f z \ closed_segment (f x) (f y)" using \z \ closed_segment x y\ by blast have "z \ affine hull S" by (meson \x \ S\ \y \ S\ \z \ closed_segment x y\ convex_affine_hull convex_contains_segment hull_inc subset_eq) then have fz_notin: "f z \ f ` S" using hull_subset inj_f inj_onD that by fastforce moreover have "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" proof - have "{.. f ` {x,y} \ {}" "{f z<..} \ f ` {x,y} \ {}" using fz fz_notin \x \ S\ \y \ S\ apply (auto simp: closed_segment_eq_real_ivl split: if_split_asm) apply (metis image_eqI less_eq_real_def)+ done then show "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" using \x \ S\ \y \ S\ by blast+ qed ultimately show False using connectedD [OF conn_fS, of "{.. S" by meson qed qed qed lemma compact_convex_collinear_segment_alt: fixes S :: "'a::euclidean_space set" assumes "S \ {}" "compact S" "connected S" "collinear S" obtains a b where "S = closed_segment a b" proof - obtain \ where "\ \ S" using \S \ {}\ by auto have "collinear (affine hull S)" by (simp add: assms collinear_affine_hull_collinear) then obtain z where "z \ 0" "\x. x \ affine hull S \ \c. x - \ = c *\<^sub>R z" by (meson \\ \ S\ collinear hull_inc) then obtain f where f: "\x. x \ affine hull S \ x - \ = f x *\<^sub>R z" by metis let ?g = "\r. r *\<^sub>R z + \" have gf: "?g (f x) = x" if "x \ affine hull S" for x by (metis diff_add_cancel f that) then have inj_f: "inj_on f (affine hull S)" by (metis inj_onI) have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \ affine hull S" and y: "y \ affine hull S" for x y proof - have "f x *\<^sub>R z = x - \" by (simp add: f hull_inc x) moreover have "f y *\<^sub>R z = y - \" by (simp add: f hull_inc y) ultimately show ?thesis by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" apply (clarsimp simp: dist_norm continuous_on_iff diff) by (metis \z \ 0\ mult.commute mult_less_cancel_left_pos norm_minus_commute real_norm_def zero_less_mult_iff zero_less_norm_iff) then have "connected (f ` S)" by (meson \connected S\ connected_continuous_image continuous_on_subset hull_subset) moreover have "compact (f ` S)" by (meson \compact S\ compact_continuous_image_eq cont_f hull_subset inj_f) ultimately obtain x y where "f ` S = {x..y}" by (meson connected_compact_interval_1) then have fS_eq: "f ` S = closed_segment x y" using \S \ {}\ closed_segment_eq_real_ivl by auto obtain a b where "a \ S" "f a = x" "b \ S" "f b = y" by (metis (full_types) ends_in_segment fS_eq imageE) have "f ` (closed_segment a b) = closed_segment (f a) (f b)" apply (rule continuous_injective_image_segment_1) apply (meson \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f]) by (meson \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f]) then have "f ` (closed_segment a b) = f ` S" by (simp add: \f a = x\ \f b = y\ fS_eq) then have "?g ` f ` (closed_segment a b) = ?g ` f ` S" by simp moreover have "(\x. f x *\<^sub>R z + \) ` closed_segment a b = closed_segment a b" apply safe apply (metis (mono_tags, hide_lams) \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment gf hull_inc subsetCE) by (metis (mono_tags, lifting) \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment gf hull_subset image_iff subsetCE) ultimately have "closed_segment a b = S" using gf by (simp add: image_comp o_def hull_inc cong: image_cong) then show ?thesis using that by blast qed lemma compact_convex_collinear_segment: fixes S :: "'a::euclidean_space set" assumes "S \ {}" "compact S" "convex S" "collinear S" obtains a b where "S = closed_segment a b" using assms convex_connected_collinear compact_convex_collinear_segment_alt by blast lemma proper_map_from_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and imf: "f ` S \ T" and "compact S" "closedin (top_of_set T) K" shows "compact (S \ f -` K)" by (rule closedin_compact [OF \compact S\] continuous_closedin_preimage_gen assms)+ lemma proper_map_fst: assumes "compact T" "K \ S" "compact K" shows "compact (S \ T \ fst -` K)" proof - have "(S \ T \ fst -` K) = K \ T" using assms by auto then show ?thesis by (simp add: assms compact_Times) qed lemma closed_map_fst: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact T" "closedin (top_of_set (S \ T)) c" shows "closedin (top_of_set S) (fst ` c)" proof - have *: "fst ` (S \ T) \ S" by auto show ?thesis using proper_map [OF _ _ *] by (simp add: proper_map_fst assms) qed lemma proper_map_snd: assumes "compact S" "K \ T" "compact K" shows "compact (S \ T \ snd -` K)" proof - have "(S \ T \ snd -` K) = S \ K" using assms by auto then show ?thesis by (simp add: assms compact_Times) qed lemma closed_map_snd: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" "closedin (top_of_set (S \ T)) c" shows "closedin (top_of_set T) (snd ` c)" proof - have *: "snd ` (S \ T) \ T" by auto show ?thesis using proper_map [OF _ _ *] by (simp add: proper_map_snd assms) qed lemma closedin_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" and clo: "closedin (top_of_set (S \ T)) U" shows "closedin (top_of_set T) {y. \x. x \ S \ (x, y) \ U}" proof - have "U \ S \ T" by (metis clo closedin_imp_subset) then have "{y. \x. x \ S \ (x, y) \ U} = snd ` U" by force moreover have "closedin (top_of_set T) (snd ` U)" by (rule closed_map_snd [OF assms]) ultimately show ?thesis by simp qed lemma closed_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "('a * 'b::euclidean_space) set" assumes "compact S" and clo: "closed T" shows "closed {y. \x. x \ S \ (x, y) \ T}" proof - have *: "{y. \x. x \ S \ Pair x y \ T} = {y. \x. x \ S \ Pair x y \ ((S \ UNIV) \ T)}" by auto show ?thesis apply (subst *) apply (rule closedin_closed_trans [OF _ closed_UNIV]) apply (rule closedin_compact_projection [OF \compact S\]) by (simp add: clo closedin_closed_Int) qed subsubsection\<^marker>\tag unimportant\\Representing affine hull as a finite intersection of hyperplanes\ proposition\<^marker>\tag unimportant\ affine_hull_convex_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "convex S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" proof show "affine hull (S \ T) \ affine hull S" by (simp add: hull_mono) next obtain a where "a \ S" "a \ T" and at: "a \ interior T" using assms interior_subset by blast then obtain e where "e > 0" and e: "cball a e \ T" using mem_interior_cball by blast have *: "x \ (+) a ` span ((\x. x - a) ` (S \ T))" if "x \ S" for x proof (cases "x = a") case True with that span_0 eq_add_iff image_def mem_Collect_eq show ?thesis by blast next case False define k where "k = min (1/2) (e / norm (x-a))" have k: "0 < k" "k < 1" using \e > 0\ False by (auto simp: k_def) then have xa: "(x-a) = inverse k *\<^sub>R k *\<^sub>R (x-a)" by simp have "e / norm (x - a) \ k" using k_def by linarith then have "a + k *\<^sub>R (x - a) \ cball a e" using \0 < k\ False by (simp add: dist_norm) (simp add: field_simps) then have T: "a + k *\<^sub>R (x - a) \ T" using e by blast have S: "a + k *\<^sub>R (x - a) \ S" using k \a \ S\ convexD [OF \convex S\ \a \ S\ \x \ S\, of "1-k" k] by (simp add: algebra_simps) have "inverse k *\<^sub>R k *\<^sub>R (x-a) \ span ((\x. x - a) ` (S \ T))" apply (rule span_mul) apply (rule span_base) apply (rule image_eqI [where x = "a + k *\<^sub>R (x - a)"]) apply (auto simp: S T) done with xa image_iff show ?thesis by fastforce qed show "affine hull S \ affine hull (S \ T)" apply (simp add: subset_hull) apply (simp add: \a \ S\ \a \ T\ hull_inc affine_hull_span_gen [of a]) apply (force simp: *) done qed corollary affine_hull_convex_Int_open: fixes S :: "'a::real_normed_vector set" assumes "convex S" "open T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using affine_hull_convex_Int_nonempty_interior assms interior_eq by blast corollary affine_hull_affine_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "affine S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" by (simp add: affine_hull_convex_Int_nonempty_interior affine_imp_convex assms) corollary affine_hull_affine_Int_open: fixes S :: "'a::real_normed_vector set" assumes "affine S" "open T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" by (simp add: affine_hull_convex_Int_open affine_imp_convex assms) corollary affine_hull_convex_Int_openin: fixes S :: "'a::real_normed_vector set" assumes "convex S" "openin (top_of_set (affine hull S)) T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using assms unfolding openin_open by (metis affine_hull_convex_Int_open hull_subset inf.orderE inf_assoc) corollary affine_hull_openin: fixes S :: "'a::real_normed_vector set" assumes "openin (top_of_set (affine hull T)) S" "S \ {}" shows "affine hull S = affine hull T" using assms unfolding openin_open by (metis affine_affine_hull affine_hull_affine_Int_open hull_hull) corollary affine_hull_open: fixes S :: "'a::real_normed_vector set" assumes "open S" "S \ {}" shows "affine hull S = UNIV" by (metis affine_hull_convex_Int_nonempty_interior assms convex_UNIV hull_UNIV inf_top.left_neutral interior_open) lemma aff_dim_convex_Int_nonempty_interior: fixes S :: "'a::euclidean_space set" shows "\convex S; S \ interior T \ {}\ \ aff_dim(S \ T) = aff_dim S" using aff_dim_affine_hull2 affine_hull_convex_Int_nonempty_interior by blast lemma aff_dim_convex_Int_open: fixes S :: "'a::euclidean_space set" shows "\convex S; open T; S \ T \ {}\ \ aff_dim(S \ T) = aff_dim S" using aff_dim_convex_Int_nonempty_interior interior_eq by blast lemma affine_hull_Diff: fixes S:: "'a::real_normed_vector set" assumes ope: "openin (top_of_set (affine hull S)) S" and "finite F" "F \ S" shows "affine hull (S - F) = affine hull S" proof - have clo: "closedin (top_of_set S) F" using assms finite_imp_closedin by auto moreover have "S - F \ {}" using assms by auto ultimately show ?thesis by (metis ope closedin_def topspace_euclidean_subtopology affine_hull_openin openin_trans) qed lemma affine_hull_halfspace_lt: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x < r} = (if a = 0 \ r \ 0 then {} else UNIV)" using halfspace_eq_empty_lt [of a r] by (simp add: open_halfspace_lt affine_hull_open) lemma affine_hull_halfspace_le: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x \ r} = (if a = 0 \ r < 0 then {} else UNIV)" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "affine hull closure {x. a \ x < r} = UNIV" using affine_hull_halfspace_lt closure_same_affine_hull by fastforce moreover have "{x. a \ x < r} \ {x. a \ x \ r}" by (simp add: Collect_mono) ultimately show ?thesis using False antisym_conv hull_mono top_greatest by (metis affine_hull_halfspace_lt) qed lemma affine_hull_halfspace_gt: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x > r} = (if a = 0 \ r \ 0 then {} else UNIV)" using halfspace_eq_empty_gt [of r a] by (simp add: open_halfspace_gt affine_hull_open) lemma affine_hull_halfspace_ge: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x \ r} = (if a = 0 \ r > 0 then {} else UNIV)" using affine_hull_halfspace_le [of "-a" "-r"] by simp lemma aff_dim_halfspace_lt: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x < r} = (if a = 0 \ r \ 0 then -1 else DIM('a))" by simp (metis aff_dim_open halfspace_eq_empty_lt open_halfspace_lt) lemma aff_dim_halfspace_le: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x \ r} = (if a = 0 \ r < 0 then -1 else DIM('a))" proof - have "int (DIM('a)) = aff_dim (UNIV::'a set)" by (simp) then have "aff_dim (affine hull {x. a \ x \ r}) = DIM('a)" if "(a = 0 \ r \ 0)" using that by (simp add: affine_hull_halfspace_le not_less) then show ?thesis by (force) qed lemma aff_dim_halfspace_gt: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x > r} = (if a = 0 \ r \ 0 then -1 else DIM('a))" by simp (metis aff_dim_open halfspace_eq_empty_gt open_halfspace_gt) lemma aff_dim_halfspace_ge: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x \ r} = (if a = 0 \ r > 0 then -1 else DIM('a))" using aff_dim_halfspace_le [of "-a" "-r"] by simp proposition aff_dim_eq_hyperplane: fixes S :: "'a::euclidean_space set" shows "aff_dim S = DIM('a) - 1 \ (\a b. a \ 0 \ affine hull S = {x. a \ x = b})" proof (cases "S = {}") case True then show ?thesis by (auto simp: dest: hyperplane_eq_Ex) next case False then obtain c where "c \ S" by blast show ?thesis proof (cases "c = 0") case True show ?thesis using span_zero [of S] apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def) apply (auto simp add: \c = 0\) done next case False have xc_im: "x \ (+) c ` {y. a \ y = 0}" if "a \ x = a \ c" for a x proof - have "\y. a \ y = 0 \ c + y = x" by (metis that add.commute diff_add_cancel inner_commute inner_diff_left right_minus_eq) then show "x \ (+) c ` {y. a \ y = 0}" by blast qed have 2: "span ((\x. x - c) ` S) = {x. a \ x = 0}" if "(+) c ` span ((\x. x - c) ` S) = {x. a \ x = b}" for a b proof - have "b = a \ c" using span_0 that by fastforce with that have "(+) c ` span ((\x. x - c) ` S) = {x. a \ x = a \ c}" by simp then have "span ((\x. x - c) ` S) = (\x. x - c) ` {x. a \ x = a \ c}" by (metis (no_types) image_cong translation_galois uminus_add_conv_diff) also have "... = {x. a \ x = 0}" by (force simp: inner_distrib inner_diff_right intro: image_eqI [where x="x+c" for x]) finally show ?thesis . qed show ?thesis apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def cong: image_cong_simp, safe) apply (fastforce simp add: inner_distrib intro: xc_im) apply (force simp: intro!: 2) done qed qed corollary aff_dim_hyperplane [simp]: fixes a :: "'a::euclidean_space" shows "a \ 0 \ aff_dim {x. a \ x = r} = DIM('a) - 1" by (metis aff_dim_eq_hyperplane affine_hull_eq affine_hyperplane) subsection\<^marker>\tag unimportant\\Some stepping theorems\ lemma aff_dim_insert: fixes a :: "'a::euclidean_space" shows "aff_dim (insert a S) = (if a \ affine hull S then aff_dim S else aff_dim S + 1)" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain x s' where S: "S = insert x s'" "x \ s'" by (meson Set.set_insert all_not_in_conv) show ?thesis using S apply (simp add: hull_redundant cong: aff_dim_affine_hull2) apply (simp add: affine_hull_insert_span_gen hull_inc) by (force simp add: span_zero insert_commute [of a] hull_inc aff_dim_eq_dim [of x] dim_insert cong: image_cong_simp) qed lemma affine_dependent_choose: fixes a :: "'a :: euclidean_space" assumes "\(affine_dependent S)" shows "affine_dependent(insert a S) \ a \ S \ a \ affine hull S" (is "?lhs = ?rhs") proof safe assume "affine_dependent (insert a S)" and "a \ S" then show "False" using \a \ S\ assms insert_absorb by fastforce next assume lhs: "affine_dependent (insert a S)" then have "a \ S" by (metis (no_types) assms insert_absorb) moreover have "finite S" using affine_independent_iff_card assms by blast moreover have "aff_dim (insert a S) \ int (card S)" using \finite S\ affine_independent_iff_card \a \ S\ lhs by fastforce ultimately show "a \ affine hull S" by (metis aff_dim_affine_independent aff_dim_insert assms) next assume "a \ S" and "a \ affine hull S" show "affine_dependent (insert a S)" by (simp add: \a \ affine hull S\ \a \ S\ affine_dependent_def) qed lemma affine_independent_insert: fixes a :: "'a :: euclidean_space" shows "\\ affine_dependent S; a \ affine hull S\ \ \ affine_dependent(insert a S)" by (simp add: affine_dependent_choose) lemma subspace_bounded_eq_trivial: fixes S :: "'a::real_normed_vector set" assumes "subspace S" shows "bounded S \ S = {0}" proof - have "False" if "bounded S" "x \ S" "x \ 0" for x proof - obtain B where B: "\y. y \ S \ norm y < B" "B > 0" using \bounded S\ by (force simp: bounded_pos_less) have "(B / norm x) *\<^sub>R x \ S" using assms subspace_mul \x \ S\ by auto moreover have "norm ((B / norm x) *\<^sub>R x) = B" using that B by (simp add: algebra_simps) ultimately show False using B by force qed then have "bounded S \ S = {0}" using assms subspace_0 by fastforce then show ?thesis by blast qed lemma affine_bounded_eq_trivial: fixes S :: "'a::real_normed_vector set" assumes "affine S" shows "bounded S \ S = {} \ (\a. S = {a})" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain b where "b \ S" by blast with False assms show ?thesis apply safe using affine_diffs_subspace [OF assms \b \ S\] apply (metis (no_types, lifting) subspace_bounded_eq_trivial ab_left_minus bounded_translation image_empty image_insert translation_invert) apply force done qed lemma affine_bounded_eq_lowdim: fixes S :: "'a::euclidean_space set" assumes "affine S" shows "bounded S \ aff_dim S \ 0" apply safe using affine_bounded_eq_trivial assms apply fastforce by (metis aff_dim_sing aff_dim_subset affine_dim_equal affine_sing all_not_in_conv assms bounded_empty bounded_insert dual_order.antisym empty_subsetI insert_subset) lemma bounded_hyperplane_eq_trivial_0: fixes a :: "'a::euclidean_space" assumes "a \ 0" shows "bounded {x. a \ x = 0} \ DIM('a) = 1" proof assume "bounded {x. a \ x = 0}" then have "aff_dim {x. a \ x = 0} \ 0" by (simp add: affine_bounded_eq_lowdim affine_hyperplane) with assms show "DIM('a) = 1" by (simp add: le_Suc_eq) next assume "DIM('a) = 1" then show "bounded {x. a \ x = 0}" by (simp add: affine_bounded_eq_lowdim affine_hyperplane assms) qed lemma bounded_hyperplane_eq_trivial: fixes a :: "'a::euclidean_space" shows "bounded {x. a \ x = r} \ (if a = 0 then r \ 0 else DIM('a) = 1)" proof (simp add: bounded_hyperplane_eq_trivial_0, clarify) assume "r \ 0" "a \ 0" have "aff_dim {x. y \ x = 0} = aff_dim {x. a \ x = r}" if "y \ 0" for y::'a by (metis that \a \ 0\ aff_dim_hyperplane) then show "bounded {x. a \ x = r} = (DIM('a) = Suc 0)" by (metis One_nat_def \a \ 0\ affine_bounded_eq_lowdim affine_hyperplane bounded_hyperplane_eq_trivial_0) qed subsection\<^marker>\tag unimportant\\General case without assuming closure and getting non-strict separation\ proposition\<^marker>\tag unimportant\ separating_hyperplane_closed_point_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "z \ S" obtains a b where "a \ S" "(a - z) \ z < b" "\x. x \ S \ b < (a - z) \ x" proof - obtain y where "y \ S" and y: "\u. u \ S \ dist z y \ dist z u" using distance_attains_inf [of S z] assms by auto then have *: "(y - z) \ z < (y - z) \ z + (norm (y - z))\<^sup>2 / 2" using \y \ S\ \z \ S\ by auto show ?thesis proof (rule that [OF \y \ S\ *]) fix x assume "x \ S" have yz: "0 < (y - z) \ (y - z)" using \y \ S\ \z \ S\ by auto { assume 0: "0 < ((z - y) \ (x - y))" with any_closest_point_dot [OF \convex S\ \closed S\] have False using y \x \ S\ \y \ S\ not_less by blast } then have "0 \ ((y - z) \ (x - y))" by (force simp: not_less inner_diff_left) with yz have "0 < 2 * ((y - z) \ (x - y)) + (y - z) \ (y - z)" by (simp add: algebra_simps) then show "(y - z) \ z + (norm (y - z))\<^sup>2 / 2 < (y - z) \ x" by (simp add: field_simps inner_diff_left inner_diff_right dot_square_norm [symmetric]) qed qed lemma separating_hyperplane_closed_0_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "0 \ S" obtains a b where "a \ S" "a \ 0" "0 < b" "\x. x \ S \ a \ x > b" using separating_hyperplane_closed_point_inset [OF assms] by simp (metis \0 \ S\) proposition\<^marker>\tag unimportant\ separating_hyperplane_set_0_inspan: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" "0 \ S" obtains a where "a \ span S" "a \ 0" "\x. x \ S \ 0 \ a \ x" proof - define k where [abs_def]: "k c = {x. 0 \ c \ x}" for c :: 'a have *: "span S \ frontier (cball 0 1) \ \f' \ {}" if f': "finite f'" "f' \ k ` S" for f' proof - obtain C where "C \ S" "finite C" and C: "f' = k ` C" using finite_subset_image [OF f'] by blast obtain a where "a \ S" "a \ 0" using \S \ {}\ \0 \ S\ ex_in_conv by blast then have "norm (a /\<^sub>R (norm a)) = 1" by simp moreover have "a /\<^sub>R (norm a) \ span S" by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp show ?thesis proof (cases "C = {}") case True with C ass show ?thesis by auto next case False have "closed (convex hull C)" using \finite C\ compact_eq_bounded_closed finite_imp_compact_convex_hull by auto moreover have "convex hull C \ {}" by (simp add: False) moreover have "0 \ convex hull C" by (metis \C \ S\ \convex S\ \0 \ S\ convex_hull_subset hull_same insert_absorb insert_subset) ultimately obtain a b where "a \ convex hull C" "a \ 0" "0 < b" and ab: "\x. x \ convex hull C \ a \ x > b" using separating_hyperplane_closed_0_inset by blast then have "a \ S" by (metis \C \ S\ assms(1) subsetCE subset_hull) moreover have "norm (a /\<^sub>R (norm a)) = 1" using \a \ 0\ by simp moreover have "a /\<^sub>R (norm a) \ span S" by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp have aa: "a /\<^sub>R (norm a) \ (\c\C. {x. 0 \ c \ x})" apply (clarsimp simp add: field_split_simps) using ab \0 < b\ by (metis hull_inc inner_commute less_eq_real_def less_trans) show ?thesis apply (simp add: C k_def) using ass aa Int_iff empty_iff by blast qed qed have "(span S \ frontier(cball 0 1)) \ (\ (k ` S)) \ {}" apply (rule compact_imp_fip) apply (blast intro: compact_cball) using closed_halfspace_ge k_def apply blast apply (metis *) done then show ?thesis unfolding set_eq_iff k_def by simp (metis inner_commute norm_eq_zero that zero_neq_one) qed lemma separating_hyperplane_set_point_inaff: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" and zno: "z \ S" obtains a b where "(z + a) \ affine hull (insert z S)" and "a \ 0" and "a \ z \ b" and "\x. x \ S \ a \ x \ b" proof - from separating_hyperplane_set_0_inspan [of "image (\x. -z + x) S"] have "convex ((+) (- z) ` S)" using \convex S\ by simp moreover have "(+) (- z) ` S \ {}" by (simp add: \S \ {}\) moreover have "0 \ (+) (- z) ` S" using zno by auto ultimately obtain a where "a \ span ((+) (- z) ` S)" "a \ 0" and a: "\x. x \ ((+) (- z) ` S) \ 0 \ a \ x" using separating_hyperplane_set_0_inspan [of "image (\x. -z + x) S"] by blast then have szx: "\x. x \ S \ a \ z \ a \ x" by (metis (no_types, lifting) imageI inner_minus_right inner_right_distrib minus_add neg_le_0_iff_le neg_le_iff_le real_add_le_0_iff) show ?thesis apply (rule_tac a=a and b = "a \ z" in that, simp_all) using \a \ span ((+) (- z) ` S)\ affine_hull_insert_span_gen apply blast apply (simp_all add: \a \ 0\ szx) done qed proposition\<^marker>\tag unimportant\ supporting_hyperplane_rel_boundary: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ S" and xno: "x \ rel_interior S" obtains a where "a \ 0" and "\y. y \ S \ a \ x \ a \ y" and "\y. y \ rel_interior S \ a \ x < a \ y" proof - obtain a b where aff: "(x + a) \ affine hull (insert x (rel_interior S))" and "a \ 0" and "a \ x \ b" and ageb: "\u. u \ (rel_interior S) \ a \ u \ b" using separating_hyperplane_set_point_inaff [of "rel_interior S" x] assms by (auto simp: rel_interior_eq_empty convex_rel_interior) have le_ay: "a \ x \ a \ y" if "y \ S" for y proof - have con: "continuous_on (closure (rel_interior S)) ((\) a)" by (rule continuous_intros continuous_on_subset | blast)+ have y: "y \ closure (rel_interior S)" using \convex S\ closure_def convex_closure_rel_interior \y \ S\ by fastforce show ?thesis using continuous_ge_on_closure [OF con y] ageb \a \ x \ b\ by fastforce qed have 3: "a \ x < a \ y" if "y \ rel_interior S" for y proof - obtain e where "0 < e" "y \ S" and e: "cball y e \ affine hull S \ S" using \y \ rel_interior S\ by (force simp: rel_interior_cball) define y' where "y' = y - (e / norm a) *\<^sub>R ((x + a) - x)" have "y' \ cball y e" unfolding y'_def using \0 < e\ by force moreover have "y' \ affine hull S" unfolding y'_def by (metis \x \ S\ \y \ S\ \convex S\ aff affine_affine_hull hull_redundant rel_interior_same_affine_hull hull_inc mem_affine_3_minus2) ultimately have "y' \ S" using e by auto have "a \ x \ a \ y" using le_ay \a \ 0\ \y \ S\ by blast moreover have "a \ x \ a \ y" using le_ay [OF \y' \ S\] \a \ 0\ apply (simp add: y'_def inner_diff dot_square_norm power2_eq_square) by (metis \0 < e\ add_le_same_cancel1 inner_commute inner_real_def inner_zero_left le_diff_eq norm_le_zero_iff real_mult_le_cancel_iff2) ultimately show ?thesis by force qed show ?thesis by (rule that [OF \a \ 0\ le_ay 3]) qed lemma supporting_hyperplane_relative_frontier: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ closure S" "x \ rel_interior S" obtains a where "a \ 0" and "\y. y \ closure S \ a \ x \ a \ y" and "\y. y \ rel_interior S \ a \ x < a \ y" using supporting_hyperplane_rel_boundary [of "closure S" x] by (metis assms convex_closure convex_rel_interior_closure) subsection\<^marker>\tag unimportant\\ Some results on decomposing convex hulls: intersections, simplicial subdivision\ lemma fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent(s \ t)" shows convex_hull_Int_subset: "convex hull s \ convex hull t \ convex hull (s \ t)" (is ?C) and affine_hull_Int_subset: "affine hull s \ affine hull t \ affine hull (s \ t)" (is ?A) proof - have [simp]: "finite s" "finite t" using aff_independent_finite assms by blast+ have "sum u (s \ t) = 1 \ (\v\s \ t. u v *\<^sub>R v) = (\v\s. u v *\<^sub>R v)" if [simp]: "sum u s = 1" "sum v t = 1" and eq: "(\x\t. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x)" for u v proof - define f where "f x = (if x \ s then u x else 0) - (if x \ t then v x else 0)" for x have "sum f (s \ t) = 0" apply (simp add: f_def sum_Un sum_subtractf) apply (simp add: sum.inter_restrict [symmetric] Int_commute) done moreover have "(\x\(s \ t). f x *\<^sub>R x) = 0" apply (simp add: f_def sum_Un scaleR_left_diff_distrib sum_subtractf) apply (simp add: if_smult sum.inter_restrict [symmetric] Int_commute eq cong del: if_weak_cong) done ultimately have "\v. v \ s \ t \ f v = 0" using aff_independent_finite assms unfolding affine_dependent_explicit by blast then have u [simp]: "\x. x \ s \ u x = (if x \ t then v x else 0)" by (simp add: f_def) presburger have "sum u (s \ t) = sum u s" by (simp add: sum.inter_restrict) then have "sum u (s \ t) = 1" using that by linarith moreover have "(\v\s \ t. u v *\<^sub>R v) = (\v\s. u v *\<^sub>R v)" by (auto simp: if_smult sum.inter_restrict intro: sum.cong) ultimately show ?thesis by force qed then show ?A ?C by (auto simp: convex_hull_finite affine_hull_finite) qed proposition\<^marker>\tag unimportant\ affine_hull_Int: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent(s \ t)" shows "affine hull (s \ t) = affine hull s \ affine hull t" apply (rule subset_antisym) apply (simp add: hull_mono) by (simp add: affine_hull_Int_subset assms) proposition\<^marker>\tag unimportant\ convex_hull_Int: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent(s \ t)" shows "convex hull (s \ t) = convex hull s \ convex hull t" apply (rule subset_antisym) apply (simp add: hull_mono) by (simp add: convex_hull_Int_subset assms) proposition\<^marker>\tag unimportant\ fixes s :: "'a::euclidean_space set set" assumes "\ affine_dependent (\s)" shows affine_hull_Inter: "affine hull (\s) = (\t\s. affine hull t)" (is "?A") and convex_hull_Inter: "convex hull (\s) = (\t\s. convex hull t)" (is "?C") proof - have "finite s" using aff_independent_finite assms finite_UnionD by blast then have "?A \ ?C" using assms proof (induction s rule: finite_induct) case empty then show ?case by auto next case (insert t F) then show ?case proof (cases "F={}") case True then show ?thesis by simp next case False with "insert.prems" have [simp]: "\ affine_dependent (t \ \F)" by (auto intro: affine_dependent_subset) have [simp]: "\ affine_dependent (\F)" using affine_independent_subset insert.prems by fastforce show ?thesis by (simp add: affine_hull_Int convex_hull_Int insert.IH) qed qed then show "?A" "?C" by auto qed proposition\<^marker>\tag unimportant\ in_convex_hull_exchange_unique: fixes S :: "'a::euclidean_space set" assumes naff: "\ affine_dependent S" and a: "a \ convex hull S" and S: "T \ S" "T' \ S" and x: "x \ convex hull (insert a T)" and x': "x \ convex hull (insert a T')" shows "x \ convex hull (insert a (T \ T'))" proof (cases "a \ S") case True then have "\ affine_dependent (insert a T \ insert a T')" using affine_dependent_subset assms by auto then have "x \ convex hull (insert a T \ insert a T')" by (metis IntI convex_hull_Int x x') then show ?thesis by simp next case False then have anot: "a \ T" "a \ T'" using assms by auto have [simp]: "finite S" by (simp add: aff_independent_finite assms) then obtain b where b0: "\s. s \ S \ 0 \ b s" and b1: "sum b S = 1" and aeq: "a = (\s\S. b s *\<^sub>R s)" using a by (auto simp: convex_hull_finite) have fin [simp]: "finite T" "finite T'" using assms infinite_super \finite S\ by blast+ then obtain c c' where c0: "\t. t \ insert a T \ 0 \ c t" and c1: "sum c (insert a T) = 1" and xeq: "x = (\t \ insert a T. c t *\<^sub>R t)" and c'0: "\t. t \ insert a T' \ 0 \ c' t" and c'1: "sum c' (insert a T') = 1" and x'eq: "x = (\t \ insert a T'. c' t *\<^sub>R t)" using x x' by (auto simp: convex_hull_finite) with fin anot have sumTT': "sum c T = 1 - c a" "sum c' T' = 1 - c' a" and wsumT: "(\t \ T. c t *\<^sub>R t) = x - c a *\<^sub>R a" by simp_all have wsumT': "(\t \ T'. c' t *\<^sub>R t) = x - c' a *\<^sub>R a" using x'eq fin anot by simp define cc where "cc \ \x. if x \ T then c x else 0" define cc' where "cc' \ \x. if x \ T' then c' x else 0" define dd where "dd \ \x. cc x - cc' x + (c a - c' a) * b x" have sumSS': "sum cc S = 1 - c a" "sum cc' S = 1 - c' a" unfolding cc_def cc'_def using S by (simp_all add: Int_absorb1 Int_absorb2 sum_subtractf sum.inter_restrict [symmetric] sumTT') have wsumSS: "(\t \ S. cc t *\<^sub>R t) = x - c a *\<^sub>R a" "(\t \ S. cc' t *\<^sub>R t) = x - c' a *\<^sub>R a" unfolding cc_def cc'_def using S by (simp_all add: Int_absorb1 Int_absorb2 if_smult sum.inter_restrict [symmetric] wsumT wsumT' cong: if_cong) have sum_dd0: "sum dd S = 0" unfolding dd_def using S by (simp add: sumSS' comm_monoid_add_class.sum.distrib sum_subtractf algebra_simps sum_distrib_right [symmetric] b1) have "(\v\S. (b v * x) *\<^sub>R v) = x *\<^sub>R (\v\S. b v *\<^sub>R v)" for x by (simp add: pth_5 real_vector.scale_sum_right mult.commute) then have *: "(\v\S. (b v * x) *\<^sub>R v) = x *\<^sub>R a" for x using aeq by blast have "(\v \ S. dd v *\<^sub>R v) = 0" unfolding dd_def using S by (simp add: * wsumSS sum.distrib sum_subtractf algebra_simps) then have dd0: "dd v = 0" if "v \ S" for v using naff that \finite S\ sum_dd0 unfolding affine_dependent_explicit apply (simp only: not_ex) apply (drule_tac x=S in spec) apply (drule_tac x=dd in spec, simp) done consider "c' a \ c a" | "c a \ c' a" by linarith then show ?thesis proof cases case 1 then have "sum cc S \ sum cc' S" by (simp add: sumSS') then have le: "cc x \ cc' x" if "x \ S" for x using dd0 [OF that] 1 b0 mult_left_mono that by (fastforce simp add: dd_def algebra_simps) have cc0: "cc x = 0" if "x \ S" "x \ T \ T'" for x using le [OF \x \ S\] that c0 by (force simp: cc_def cc'_def split: if_split_asm) show ?thesis proof (simp add: convex_hull_finite, intro exI conjI) show "\x\T \ T'. 0 \ (cc(a := c a)) x" by (simp add: c0 cc_def) show "0 \ (cc(a := c a)) a" by (simp add: c0) have "sum (cc(a := c a)) (insert a (T \ T')) = c a + sum (cc(a := c a)) (T \ T')" by (simp add: anot) also have "... = c a + sum (cc(a := c a)) S" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c a + (1 - c a)" by (metis \a \ S\ fun_upd_other sum.cong sumSS') finally show "sum (cc(a := c a)) (insert a (T \ T')) = 1" by simp have "(\x\insert a (T \ T'). (cc(a := c a)) x *\<^sub>R x) = c a *\<^sub>R a + (\x \ T \ T'. (cc(a := c a)) x *\<^sub>R x)" by (simp add: anot) also have "... = c a *\<^sub>R a + (\x \ S. (cc(a := c a)) x *\<^sub>R x)" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" by (simp add: wsumSS \a \ S\ if_smult sum_delta_notmem) finally show "(\x\insert a (T \ T'). (cc(a := c a)) x *\<^sub>R x) = x" by simp qed next case 2 then have "sum cc' S \ sum cc S" by (simp add: sumSS') then have le: "cc' x \ cc x" if "x \ S" for x using dd0 [OF that] 2 b0 mult_left_mono that by (fastforce simp add: dd_def algebra_simps) have cc0: "cc' x = 0" if "x \ S" "x \ T \ T'" for x using le [OF \x \ S\] that c'0 by (force simp: cc_def cc'_def split: if_split_asm) show ?thesis proof (simp add: convex_hull_finite, intro exI conjI) show "\x\T \ T'. 0 \ (cc'(a := c' a)) x" by (simp add: c'0 cc'_def) show "0 \ (cc'(a := c' a)) a" by (simp add: c'0) have "sum (cc'(a := c' a)) (insert a (T \ T')) = c' a + sum (cc'(a := c' a)) (T \ T')" by (simp add: anot) also have "... = c' a + sum (cc'(a := c' a)) S" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c' a + (1 - c' a)" by (metis \a \ S\ fun_upd_other sum.cong sumSS') finally show "sum (cc'(a := c' a)) (insert a (T \ T')) = 1" by simp have "(\x\insert a (T \ T'). (cc'(a := c' a)) x *\<^sub>R x) = c' a *\<^sub>R a + (\x \ T \ T'. (cc'(a := c' a)) x *\<^sub>R x)" by (simp add: anot) also have "... = c' a *\<^sub>R a + (\x \ S. (cc'(a := c' a)) x *\<^sub>R x)" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" by (simp add: wsumSS \a \ S\ if_smult sum_delta_notmem) finally show "(\x\insert a (T \ T'). (cc'(a := c' a)) x *\<^sub>R x) = x" by simp qed qed qed corollary\<^marker>\tag unimportant\ convex_hull_exchange_Int: fixes a :: "'a::euclidean_space" assumes "\ affine_dependent S" "a \ convex hull S" "T \ S" "T' \ S" shows "(convex hull (insert a T)) \ (convex hull (insert a T')) = convex hull (insert a (T \ T'))" apply (rule subset_antisym) using in_convex_hull_exchange_unique assms apply blast by (metis hull_mono inf_le1 inf_le2 insert_inter_insert le_inf_iff) lemma Int_closed_segment: fixes b :: "'a::euclidean_space" assumes "b \ closed_segment a c \ \ collinear{a,b,c}" shows "closed_segment a b \ closed_segment b c = {b}" proof (cases "c = a") case True then show ?thesis using assms collinear_3_eq_affine_dependent by fastforce next case False from assms show ?thesis proof assume "b \ closed_segment a c" moreover have "\ affine_dependent {a, c}" by (simp) ultimately show ?thesis using False convex_hull_exchange_Int [of "{a,c}" b "{a}" "{c}"] by (simp add: segment_convex_hull insert_commute) next assume ncoll: "\ collinear {a, b, c}" have False if "closed_segment a b \ closed_segment b c \ {b}" proof - have "b \ closed_segment a b" and "b \ closed_segment b c" by auto with that obtain d where "b \ d" "d \ closed_segment a b" "d \ closed_segment b c" by force then have d: "collinear {a, d, b}" "collinear {b, d, c}" by (auto simp: between_mem_segment between_imp_collinear) have "collinear {a, b, c}" apply (rule collinear_3_trans [OF _ _ \b \ d\]) using d by (auto simp: insert_commute) with ncoll show False .. qed then show ?thesis by blast qed qed lemma affine_hull_finite_intersection_hyperplanes: fixes s :: "'a::euclidean_space set" obtains f where "finite f" "of_nat (card f) + aff_dim s = DIM('a)" "affine hull s = \f" "\h. h \ f \ \a b. a \ 0 \ h = {x. a \ x = b}" proof - obtain b where "b \ s" and indb: "\ affine_dependent b" and eq: "affine hull s = affine hull b" using affine_basis_exists by blast obtain c where indc: "\ affine_dependent c" and "b \ c" and affc: "affine hull c = UNIV" by (metis extend_to_affine_basis affine_UNIV hull_same indb subset_UNIV) then have "finite c" by (simp add: aff_independent_finite) then have fbc: "finite b" "card b \ card c" using \b \ c\ infinite_super by (auto simp: card_mono) have imeq: "(\x. affine hull x) ` ((\a. c - {a}) ` (c - b)) = ((\a. affine hull (c - {a})) ` (c - b))" by blast have card1: "card ((\a. affine hull (c - {a})) ` (c - b)) = card (c - b)" apply (rule card_image [OF inj_onI]) by (metis Diff_eq_empty_iff Diff_iff indc affine_dependent_def hull_subset insert_iff) have card2: "(card (c - b)) + aff_dim s = DIM('a)" proof - have aff: "aff_dim (UNIV::'a set) = aff_dim c" by (metis aff_dim_affine_hull affc) have "aff_dim b = aff_dim s" by (metis (no_types) aff_dim_affine_hull eq) then have "int (card b) = 1 + aff_dim s" by (simp add: aff_dim_affine_independent indb) then show ?thesis using fbc aff by (simp add: \\ affine_dependent c\ \b \ c\ aff_dim_affine_independent card_Diff_subset of_nat_diff) qed show ?thesis proof (cases "c = b") case True show ?thesis apply (rule_tac f="{}" in that) using True affc apply (simp_all add: eq [symmetric]) by (metis aff_dim_UNIV aff_dim_affine_hull) next case False have ind: "\ affine_dependent (\a\c - b. c - {a})" by (rule affine_independent_subset [OF indc]) auto have affeq: "affine hull s = (\x\(\a. c - {a}) ` (c - b). affine hull x)" using \b \ c\ False apply (subst affine_hull_Inter [OF ind, symmetric]) apply (simp add: eq double_diff) done have *: "1 + aff_dim (c - {t}) = int (DIM('a))" if t: "t \ c" for t proof - have "insert t c = c" using t by blast then show ?thesis by (metis (full_types) add.commute aff_dim_affine_hull aff_dim_insert aff_dim_UNIV affc affine_dependent_def indc insert_Diff_single t) qed show ?thesis apply (rule_tac f = "(\x. affine hull x) ` ((\a. c - {a}) ` (c - b))" in that) using \finite c\ apply blast apply (simp add: imeq card1 card2) apply (simp add: affeq, clarify) apply (metis DIM_positive One_nat_def Suc_leI add_diff_cancel_left' of_nat_1 aff_dim_eq_hyperplane of_nat_diff *) done qed qed lemma affine_hyperplane_sums_eq_UNIV_0: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "0 \ S" and "w \ S" and "a \ w \ 0" shows "{x + y| x y. x \ S \ a \ y = 0} = UNIV" proof - have "subspace S" by (simp add: assms subspace_affine) have span1: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" apply (rule span_mono) using \0 \ S\ add.left_neutral by force have "w \ span {y. a \ y = 0}" using \a \ w \ 0\ span_induct subspace_hyperplane by auto moreover have "w \ span {x + y |x y. x \ S \ a \ y = 0}" using \w \ S\ by (metis (mono_tags, lifting) inner_zero_right mem_Collect_eq pth_d span_base) ultimately have span2: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" by blast have "a \ 0" using assms inner_zero_left by blast then have "DIM('a) - 1 = dim {y. a \ y = 0}" by (simp add: dim_hyperplane) also have "... < dim {x + y |x y. x \ S \ a \ y = 0}" using span1 span2 by (blast intro: dim_psubset) finally have DIM_lt: "DIM('a) - 1 < dim {x + y |x y. x \ S \ a \ y = 0}" . have subs: "subspace {x + y| x y. x \ S \ a \ y = 0}" using subspace_sums [OF \subspace S\ subspace_hyperplane] by simp moreover have "span {x + y| x y. x \ S \ a \ y = 0} = UNIV" apply (rule dim_eq_full [THEN iffD1]) apply (rule antisym [OF dim_subset_UNIV]) using DIM_lt apply simp done ultimately show ?thesis by (simp add: subs) (metis (lifting) span_eq_iff subs) qed proposition\<^marker>\tag unimportant\ affine_hyperplane_sums_eq_UNIV: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {v. a \ v = b} \ {}" and "S - {v. a \ v = b} \ {}" shows "{x + y| x y. x \ S \ a \ y = b} = UNIV" proof (cases "a = 0") case True with assms show ?thesis by (auto simp: if_splits) next case False obtain c where "c \ S" and c: "a \ c = b" using assms by force with affine_diffs_subspace [OF \affine S\] have "subspace ((+) (- c) ` S)" by blast then have aff: "affine ((+) (- c) ` S)" by (simp add: subspace_imp_affine) have 0: "0 \ (+) (- c) ` S" by (simp add: \c \ S\) obtain d where "d \ S" and "a \ d \ b" and dc: "d-c \ (+) (- c) ` S" using assms by auto then have adc: "a \ (d - c) \ 0" by (simp add: c inner_diff_right) let ?U = "(+) (c+c) ` {x + y |x y. x \ (+) (- c) ` S \ a \ y = 0}" have "u + v \ (+) (c + c) ` {x + v |x v. x \ (+) (- c) ` S \ a \ v = 0}" if "u \ S" "b = a \ v" for u v apply (rule_tac x="u+v-c-c" in image_eqI) apply (simp_all add: algebra_simps) apply (rule_tac x="u-c" in exI) apply (rule_tac x="v-c" in exI) apply (simp add: algebra_simps that c) done moreover have "\a \ v = 0; u \ S\ \ \x ya. v + (u + c) = x + ya \ x \ S \ a \ ya = b" for v u by (metis add.left_commute c inner_right_distrib pth_d) ultimately have "{x + y |x y. x \ S \ a \ y = b} = ?U" by (fastforce simp: algebra_simps) also have "... = range ((+) (c + c))" by (simp only: affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc]) also have "... = UNIV" by simp finally show ?thesis . qed lemma aff_dim_sums_Int_0: assumes "affine S" and "affine T" and "0 \ S" "0 \ T" shows "aff_dim {x + y| x y. x \ S \ y \ T} = (aff_dim S + aff_dim T) - aff_dim(S \ T)" proof - have "0 \ {x + y |x y. x \ S \ y \ T}" using assms by force then have 0: "0 \ affine hull {x + y |x y. x \ S \ y \ T}" by (metis (lifting) hull_inc) have sub: "subspace S" "subspace T" using assms by (auto simp: subspace_affine) show ?thesis using dim_sums_Int [OF sub] by (simp add: aff_dim_zero assms 0 hull_inc) qed proposition aff_dim_sums_Int: assumes "affine S" and "affine T" and "S \ T \ {}" shows "aff_dim {x + y| x y. x \ S \ y \ T} = (aff_dim S + aff_dim T) - aff_dim(S \ T)" proof - obtain a where a: "a \ S" "a \ T" using assms by force have aff: "affine ((+) (-a) ` S)" "affine ((+) (-a) ` T)" using affine_translation [symmetric, of "- a"] assms by (simp_all cong: image_cong_simp) have zero: "0 \ ((+) (-a) ` S)" "0 \ ((+) (-a) ` T)" using a assms by auto have "{x + y |x y. x \ (+) (- a) ` S \ y \ (+) (- a) ` T} = (+) (- 2 *\<^sub>R a) ` {x + y| x y. x \ S \ y \ T}" by (force simp: algebra_simps scaleR_2) moreover have "(+) (- a) ` S \ (+) (- a) ` T = (+) (- a) ` (S \ T)" by auto ultimately show ?thesis using aff_dim_sums_Int_0 [OF aff zero] aff_dim_translation_eq by (metis (lifting)) qed lemma aff_dim_affine_Int_hyperplane: fixes a :: "'a::euclidean_space" assumes "affine S" shows "aff_dim(S \ {x. a \ x = b}) = (if S \ {v. a \ v = b} = {} then - 1 else if S \ {v. a \ v = b} then aff_dim S else aff_dim S - 1)" proof (cases "a = 0") case True with assms show ?thesis by auto next case False then have "aff_dim (S \ {x. a \ x = b}) = aff_dim S - 1" if "x \ S" "a \ x \ b" and non: "S \ {v. a \ v = b} \ {}" for x proof - have [simp]: "{x + y| x y. x \ S \ a \ y = b} = UNIV" using affine_hyperplane_sums_eq_UNIV [OF assms non] that by blast show ?thesis using aff_dim_sums_Int [OF assms affine_hyperplane non] by (simp add: of_nat_diff False) qed then show ?thesis by (metis (mono_tags, lifting) inf.orderE aff_dim_empty_eq mem_Collect_eq subsetI) qed lemma aff_dim_lt_full: fixes S :: "'a::euclidean_space set" shows "aff_dim S < DIM('a) \ (affine hull S \ UNIV)" by (metis (no_types) aff_dim_affine_hull aff_dim_le_DIM aff_dim_UNIV affine_hull_UNIV less_le) lemma aff_dim_openin: fixes S :: "'a::euclidean_space set" assumes ope: "openin (top_of_set T) S" and "affine T" "S \ {}" shows "aff_dim S = aff_dim T" proof - show ?thesis proof (rule order_antisym) show "aff_dim S \ aff_dim T" by (blast intro: aff_dim_subset [OF openin_imp_subset] ope) next obtain a where "a \ S" using \S \ {}\ by blast have "S \ T" using ope openin_imp_subset by auto then have "a \ T" using \a \ S\ by auto then have subT': "subspace ((\x. - a + x) ` T)" using affine_diffs_subspace \affine T\ by auto then obtain B where Bsub: "B \ ((\x. - a + x) ` T)" and po: "pairwise orthogonal B" and eq1: "\x. x \ B \ norm x = 1" and "independent B" and cardB: "card B = dim ((\x. - a + x) ` T)" and spanB: "span B = ((\x. - a + x) ` T)" by (rule orthonormal_basis_subspace) auto obtain e where "0 < e" and e: "cball a e \ T \ S" by (meson \a \ S\ openin_contains_cball ope) have "aff_dim T = aff_dim ((\x. - a + x) ` T)" by (metis aff_dim_translation_eq) also have "... = dim ((\x. - a + x) ` T)" using aff_dim_subspace subT' by blast also have "... = card B" by (simp add: cardB) also have "... = card ((\x. e *\<^sub>R x) ` B)" using \0 < e\ by (force simp: inj_on_def card_image) also have "... \ dim ((\x. - a + x) ` S)" proof (simp, rule independent_card_le_dim) have e': "cball 0 e \ (\x. x - a) ` T \ (\x. x - a) ` S" using e by (auto simp: dist_norm norm_minus_commute subset_eq) have "(\x. e *\<^sub>R x) ` B \ cball 0 e \ (\x. x - a) ` T" using Bsub \0 < e\ eq1 subT' \a \ T\ by (auto simp: subspace_def) then show "(\x. e *\<^sub>R x) ` B \ (\x. x - a) ` S" using e' by blast show "independent ((\x. e *\<^sub>R x) ` B)" using linear_scale_self \independent B\ apply (rule linear_independent_injective_image) using \0 < e\ inj_on_def by fastforce qed also have "... = aff_dim S" using \a \ S\ aff_dim_eq_dim hull_inc by (force cong: image_cong_simp) finally show "aff_dim T \ aff_dim S" . qed qed lemma dim_openin: fixes S :: "'a::euclidean_space set" assumes ope: "openin (top_of_set T) S" and "subspace T" "S \ {}" shows "dim S = dim T" proof (rule order_antisym) show "dim S \ dim T" by (metis ope dim_subset openin_subset topspace_euclidean_subtopology) next have "dim T = aff_dim S" using aff_dim_openin by (metis aff_dim_subspace \subspace T\ \S \ {}\ ope subspace_affine) also have "... \ dim S" by (metis aff_dim_subset aff_dim_subspace dim_span span_superset subspace_span) finally show "dim T \ dim S" by simp qed subsection\Lower-dimensional affine subsets are nowhere dense\ proposition dense_complement_subspace: fixes S :: "'a :: euclidean_space set" assumes dim_less: "dim T < dim S" and "subspace S" shows "closure(S - T) = S" proof - have "closure(S - U) = S" if "dim U < dim S" "U \ S" for U proof - have "span U \ span S" by (metis neq_iff psubsetI span_eq_dim span_mono that) then obtain a where "a \ 0" "a \ span S" and a: "\y. y \ span U \ orthogonal a y" using orthogonal_to_subspace_exists_gen by metis show ?thesis proof have "closed S" by (simp add: \subspace S\ closed_subspace) then show "closure (S - U) \ S" by (simp add: closure_minimal) show "S \ closure (S - U)" proof (clarsimp simp: closure_approachable) fix x and e::real assume "x \ S" "0 < e" show "\y\S - U. dist y x < e" proof (cases "x \ U") case True let ?y = "x + (e/2 / norm a) *\<^sub>R a" show ?thesis proof show "dist ?y x < e" using \0 < e\ by (simp add: dist_norm) next have "?y \ S" by (metis \a \ span S\ \x \ S\ assms(2) span_eq_iff subspace_add subspace_scale) moreover have "?y \ U" proof - have "e/2 / norm a \ 0" using \0 < e\ \a \ 0\ by auto then show ?thesis by (metis True \a \ 0\ a orthogonal_scaleR orthogonal_self real_vector.scale_eq_0_iff span_add_eq span_base) qed ultimately show "?y \ S - U" by blast qed next case False with \0 < e\ \x \ S\ show ?thesis by force qed qed qed qed moreover have "S - S \ T = S-T" by blast moreover have "dim (S \ T) < dim S" by (metis dim_less dim_subset inf.cobounded2 inf.orderE inf.strict_boundedE not_le) ultimately show ?thesis by force qed corollary\<^marker>\tag unimportant\ dense_complement_affine: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" and "affine S" shows "closure(S - T) = S" proof (cases "S \ T = {}") case True then show ?thesis by (metis Diff_triv affine_hull_eq \affine S\ closure_same_affine_hull closure_subset hull_subset subset_antisym) next case False then obtain z where z: "z \ S \ T" by blast then have "subspace ((+) (- z) ` S)" by (meson IntD1 affine_diffs_subspace \affine S\) moreover have "int (dim ((+) (- z) ` T)) < int (dim ((+) (- z) ` S))" thm aff_dim_eq_dim using z less by (simp add: aff_dim_eq_dim_subtract [of z] hull_inc cong: image_cong_simp) ultimately have "closure(((+) (- z) ` S) - ((+) (- z) ` T)) = ((+) (- z) ` S)" by (simp add: dense_complement_subspace) then show ?thesis by (metis closure_translation translation_diff translation_invert) qed corollary\<^marker>\tag unimportant\ dense_complement_openin_affine_hull: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" and ope: "openin (top_of_set (affine hull S)) S" shows "closure(S - T) = closure S" proof - have "affine hull S - T \ affine hull S" by blast then have "closure (S \ closure (affine hull S - T)) = closure (S \ (affine hull S - T))" by (rule closure_openin_Int_closure [OF ope]) then show ?thesis by (metis Int_Diff aff_dim_affine_hull affine_affine_hull dense_complement_affine hull_subset inf.orderE less) qed corollary\<^marker>\tag unimportant\ dense_complement_convex: fixes S :: "'a :: euclidean_space set" assumes "aff_dim T < aff_dim S" "convex S" shows "closure(S - T) = closure S" proof show "closure (S - T) \ closure S" by (simp add: closure_mono) have "closure (rel_interior S - T) = closure (rel_interior S)" apply (rule dense_complement_openin_affine_hull) apply (simp add: assms rel_interior_aff_dim) using \convex S\ rel_interior_rel_open rel_open by blast then show "closure S \ closure (S - T)" by (metis Diff_mono \convex S\ closure_mono convex_closure_rel_interior order_refl rel_interior_subset) qed corollary\<^marker>\tag unimportant\ dense_complement_convex_closed: fixes S :: "'a :: euclidean_space set" assumes "aff_dim T < aff_dim S" "convex S" "closed S" shows "closure(S - T) = S" by (simp add: assms dense_complement_convex) subsection\<^marker>\tag unimportant\\Parallel slices, etc\ text\ If we take a slice out of a set, we can do it perpendicularly, with the normal vector to the slice parallel to the affine hull.\ proposition\<^marker>\tag unimportant\ affine_parallel_slice: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {x. a \ x \ b} \ {}" and "\ (S \ {x. a \ x \ b})" obtains a' b' where "a' \ 0" "S \ {x. a' \ x \ b'} = S \ {x. a \ x \ b}" "S \ {x. a' \ x = b'} = S \ {x. a \ x = b}" "\w. w \ S \ (w + a') \ S" proof (cases "S \ {x. a \ x = b} = {}") case True then obtain u v where "u \ S" "v \ S" "a \ u \ b" "a \ v > b" using assms by (auto simp: not_le) define \ where "\ = u + ((b - a \ u) / (a \ v - a \ u)) *\<^sub>R (v - u)" have "\ \ S" by (simp add: \_def \u \ S\ \v \ S\ \affine S\ mem_affine_3_minus) moreover have "a \ \ = b" using \a \ u \ b\ \b < a \ v\ by (simp add: \_def algebra_simps) (simp add: field_simps) ultimately have False using True by force then show ?thesis .. next case False then obtain z where "z \ S" and z: "a \ z = b" using assms by auto with affine_diffs_subspace [OF \affine S\] have sub: "subspace ((+) (- z) ` S)" by blast then have aff: "affine ((+) (- z) ` S)" and span: "span ((+) (- z) ` S) = ((+) (- z) ` S)" by (auto simp: subspace_imp_affine) obtain a' a'' where a': "a' \ span ((+) (- z) ` S)" and a: "a = a' + a''" and "\w. w \ span ((+) (- z) ` S) \ orthogonal a'' w" using orthogonal_subspace_decomp_exists [of "(+) (- z) ` S" "a"] by metis then have "\w. w \ S \ a'' \ (w-z) = 0" by (simp add: span_base orthogonal_def) then have a'': "\w. w \ S \ a'' \ w = (a - a') \ z" by (simp add: a inner_diff_right) then have ba'': "\w. w \ S \ a'' \ w = b - a' \ z" by (simp add: inner_diff_left z) have "\w. w \ (+) (- z) ` S \ (w + a') \ (+) (- z) ` S" by (metis subspace_add a' span_eq_iff sub) then have Sclo: "\w. w \ S \ (w + a') \ S" by fastforce show ?thesis proof (cases "a' = 0") case True with a assms True a'' diff_zero less_irrefl show ?thesis by auto next case False show ?thesis apply (rule_tac a' = "a'" and b' = "a' \ z" in that) apply (auto simp: a ba'' inner_left_distrib False Sclo) done qed qed lemma diffs_affine_hull_span: assumes "a \ S" shows "{x - a |x. x \ affine hull S} = span {x - a |x. x \ S}" proof - have *: "((\x. x - a) ` (S - {a})) = {x. x + a \ S} - {0}" by (auto simp: algebra_simps) show ?thesis apply (simp add: affine_hull_span2 [OF assms] *) apply (auto simp: algebra_simps) done qed lemma aff_dim_dim_affine_diffs: fixes S :: "'a :: euclidean_space set" assumes "affine S" "a \ S" shows "aff_dim S = dim {x - a |x. x \ S}" proof - obtain B where aff: "affine hull B = affine hull S" and ind: "\ affine_dependent B" and card: "of_nat (card B) = aff_dim S + 1" using aff_dim_basis_exists by blast then have "B \ {}" using assms by (metis affine_hull_eq_empty ex_in_conv) then obtain c where "c \ B" by auto then have "c \ S" by (metis aff affine_hull_eq \affine S\ hull_inc) have xy: "x - c = y - a \ y = x + 1 *\<^sub>R (a - c)" for x y c and a::'a by (auto simp: algebra_simps) have *: "{x - c |x. x \ S} = {x - a |x. x \ S}" apply safe apply (simp_all only: xy) using mem_affine_3_minus [OF \affine S\] \a \ S\ \c \ S\ apply blast+ done have affS: "affine hull S = S" by (simp add: \affine S\) have "aff_dim S = of_nat (card B) - 1" using card by simp also have "... = dim {x - c |x. x \ B}" by (simp add: affine_independent_card_dim_diffs [OF ind \c \ B\]) also have "... = dim {x - c | x. x \ affine hull B}" by (simp add: diffs_affine_hull_span \c \ B\) also have "... = dim {x - a |x. x \ S}" by (simp add: affS aff *) finally show ?thesis . qed lemma aff_dim_linear_image_le: assumes "linear f" shows "aff_dim(f ` S) \ aff_dim S" proof - have "aff_dim (f ` T) \ aff_dim T" if "affine T" for T proof (cases "T = {}") case True then show ?thesis by (simp add: aff_dim_geq) next case False then obtain a where "a \ T" by auto have 1: "((\x. x - f a) ` f ` T) = {x - f a |x. x \ f ` T}" by auto have 2: "{x - f a| x. x \ f ` T} = f ` {x - a| x. x \ T}" by (force simp: linear_diff [OF assms]) have "aff_dim (f ` T) = int (dim {x - f a |x. x \ f ` T})" by (simp add: \a \ T\ hull_inc aff_dim_eq_dim [of "f a"] 1 cong: image_cong_simp) also have "... = int (dim (f ` {x - a| x. x \ T}))" by (force simp: linear_diff [OF assms] 2) also have "... \ int (dim {x - a| x. x \ T})" by (simp add: dim_image_le [OF assms]) also have "... \ aff_dim T" by (simp add: aff_dim_dim_affine_diffs [symmetric] \a \ T\ \affine T\) finally show ?thesis . qed then have "aff_dim (f ` (affine hull S)) \ aff_dim (affine hull S)" using affine_affine_hull [of S] by blast then show ?thesis using affine_hull_linear_image assms linear_conv_bounded_linear by fastforce qed lemma aff_dim_injective_linear_image [simp]: assumes "linear f" "inj f" shows "aff_dim (f ` S) = aff_dim S" proof (rule antisym) show "aff_dim (f ` S) \ aff_dim S" by (simp add: aff_dim_linear_image_le assms(1)) next obtain g where "linear g" "g \ f = id" using assms(1) assms(2) linear_injective_left_inverse by blast then have "aff_dim S \ aff_dim(g ` f ` S)" by (simp add: image_comp) also have "... \ aff_dim (f ` S)" by (simp add: \linear g\ aff_dim_linear_image_le) finally show "aff_dim S \ aff_dim (f ` S)" . qed lemma choose_affine_subset: assumes "affine S" "-1 \ d" and dle: "d \ aff_dim S" obtains T where "affine T" "T \ S" "aff_dim T = d" proof (cases "d = -1 \ S={}") case True with assms show ?thesis by (metis aff_dim_empty affine_empty bot.extremum that eq_iff) next case False with assms obtain a where "a \ S" "0 \ d" by auto with assms have ss: "subspace ((+) (- a) ` S)" by (simp add: affine_diffs_subspace_subtract cong: image_cong_simp) have "nat d \ dim ((+) (- a) ` S)" by (metis aff_dim_subspace aff_dim_translation_eq dle nat_int nat_mono ss) then obtain T where "subspace T" and Tsb: "T \ span ((+) (- a) ` S)" and Tdim: "dim T = nat d" using choose_subspace_of_subspace [of "nat d" "(+) (- a) ` S"] by blast then have "affine T" using subspace_affine by blast then have "affine ((+) a ` T)" by (metis affine_hull_eq affine_hull_translation) moreover have "(+) a ` T \ S" proof - have "T \ (+) (- a) ` S" by (metis (no_types) span_eq_iff Tsb ss) then show "(+) a ` T \ S" using add_ac by auto qed moreover have "aff_dim ((+) a ` T) = d" by (simp add: aff_dim_subspace Tdim \0 \ d\ \subspace T\ aff_dim_translation_eq) ultimately show ?thesis by (rule that) qed subsection\Paracompactness\ proposition paracompact: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "S \ \\" and opC: "\T. T \ \ \ open T" obtains \' where "S \ \ \'" and "\U. U \ \' \ open U \ (\T. T \ \ \ U \ T)" and "\x. x \ S \ \V. open V \ x \ V \ finite {U. U \ \' \ (U \ V \ {})}" proof (cases "S = {}") case True with that show ?thesis by blast next case False have "\T U. x \ U \ open U \ closure U \ T \ T \ \" if "x \ S" for x proof - obtain T where "x \ T" "T \ \" "open T" using assms \x \ S\ by blast then obtain e where "e > 0" "cball x e \ T" by (force simp: open_contains_cball) then show ?thesis apply (rule_tac x = T in exI) apply (rule_tac x = "ball x e" in exI) using \T \ \\ apply (simp add: closure_minimal) using closed_cball closure_minimal by blast qed then obtain F G where Gin: "x \ G x" and oG: "open (G x)" and clos: "closure (G x) \ F x" and Fin: "F x \ \" if "x \ S" for x by metis then obtain \ where "\ \ G ` S" "countable \" "\\ = \(G ` S)" using Lindelof [of "G ` S"] by (metis image_iff) then obtain K where K: "K \ S" "countable K" and eq: "\(G ` K) = \(G ` S)" by (metis countable_subset_image) with False Gin have "K \ {}" by force then obtain a :: "nat \ 'a" where "range a = K" by (metis range_from_nat_into \countable K\) then have odif: "\n. open (F (a n) - \{closure (G (a m)) |m. m < n})" using \K \ S\ Fin opC by (fastforce simp add:) let ?C = "range (\n. F(a n) - \{closure(G(a m)) |m. m < n})" have enum_S: "\n. x \ F(a n) \ x \ G(a n)" if "x \ S" for x proof - have "\y \ K. x \ G y" using eq that Gin by fastforce then show ?thesis using clos K \range a = K\ closure_subset by blast qed have 1: "S \ Union ?C" proof fix x assume "x \ S" define n where "n \ LEAST n. x \ F(a n)" have n: "x \ F(a n)" using enum_S [OF \x \ S\] by (force simp: n_def intro: LeastI) have notn: "x \ F(a m)" if "m < n" for m using that not_less_Least by (force simp: n_def) then have "x \ \{closure (G (a m)) |m. m < n}" using n \K \ S\ \range a = K\ clos notn by fastforce with n show "x \ Union ?C" by blast qed have 3: "\V. open V \ x \ V \ finite {U. U \ ?C \ (U \ V \ {})}" if "x \ S" for x proof - obtain n where n: "x \ F(a n)" "x \ G(a n)" using \x \ S\ enum_S by auto have "{U \ ?C. U \ G (a n) \ {}} \ (\n. F(a n) - \{closure(G(a m)) |m. m < n}) ` atMost n" proof clarsimp fix k assume "(F (a k) - \{closure (G (a m)) |m. m < k}) \ G (a n) \ {}" then have "k \ n" by auto (metis closure_subset not_le subsetCE) then show "F (a k) - \{closure (G (a m)) |m. m < k} \ (\n. F (a n) - \{closure (G (a m)) |m. m < n}) ` {..n}" by force qed moreover have "finite ((\n. F(a n) - \{closure(G(a m)) |m. m < n}) ` atMost n)" by force ultimately have *: "finite {U \ ?C. U \ G (a n) \ {}}" using finite_subset by blast show ?thesis apply (rule_tac x="G (a n)" in exI) apply (intro conjI oG n *) using \K \ S\ \range a = K\ apply blast done qed show ?thesis apply (rule that [OF 1 _ 3]) using Fin \K \ S\ \range a = K\ apply (auto simp: odif) done qed corollary paracompact_closedin: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes cin: "closedin (top_of_set U) S" and oin: "\T. T \ \ \ openin (top_of_set U) T" and "S \ \\" obtains \' where "S \ \ \'" and "\V. V \ \' \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" and "\x. x \ U \ \V. openin (top_of_set U) V \ x \ V \ finite {X. X \ \' \ (X \ V \ {})}" proof - have "\Z. open Z \ (T = U \ Z)" if "T \ \" for T using oin [OF that] by (auto simp: openin_open) then obtain F where opF: "open (F T)" and intF: "U \ F T = T" if "T \ \" for T by metis obtain K where K: "closed K" "U \ K = S" using cin by (auto simp: closedin_closed) have 1: "U \ \(insert (- K) (F ` \))" by clarsimp (metis Int_iff Union_iff \U \ K = S\ \S \ \\\ subsetD intF) have 2: "\T. T \ insert (- K) (F ` \) \ open T" using \closed K\ by (auto simp: opF) obtain \ where "U \ \\" and D1: "\U. U \ \ \ open U \ (\T. T \ insert (- K) (F ` \) \ U \ T)" and D2: "\x. x \ U \ \V. open V \ x \ V \ finite {U \ \. U \ V \ {}}" by (blast intro: paracompact [OF 1 2]) let ?C = "{U \ V |V. V \ \ \ (V \ K \ {})}" show ?thesis proof (rule_tac \' = "{U \ V |V. V \ \ \ (V \ K \ {})}" in that) show "S \ \?C" using \U \ K = S\ \U \ \\\ K by (blast dest!: subsetD) show "\V. V \ ?C \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" using D1 intF by fastforce have *: "{X. (\V. X = U \ V \ V \ \ \ V \ K \ {}) \ X \ (U \ V) \ {}} \ (\x. U \ x) ` {U \ \. U \ V \ {}}" for V by blast show "\V. openin (top_of_set U) V \ x \ V \ finite {X \ ?C. X \ V \ {}}" if "x \ U" for x using D2 [OF that] apply clarify apply (rule_tac x="U \ V" in exI) apply (auto intro: that finite_subset [OF *]) done qed qed corollary\<^marker>\tag unimportant\ paracompact_closed: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "closed S" and opC: "\T. T \ \ \ open T" and "S \ \\" obtains \' where "S \ \\'" and "\U. U \ \' \ open U \ (\T. T \ \ \ U \ T)" and "\x. \V. open V \ x \ V \ finite {U. U \ \' \ (U \ V \ {})}" by (rule paracompact_closedin [of UNIV S \]) (auto simp: assms) subsection\<^marker>\tag unimportant\\Closed-graph characterization of continuity\ lemma continuous_closed_graph_gen: fixes T :: "'b::real_normed_vector set" assumes contf: "continuous_on S f" and fim: "f ` S \ T" shows "closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" proof - have eq: "((\x. Pair x (f x)) ` S) =(S \ T \ (\z. (f \ fst)z - snd z) -` {0})" using fim by auto show ?thesis apply (subst eq) apply (intro continuous_intros continuous_closedin_preimage continuous_on_subset [OF contf]) by auto qed lemma continuous_closed_graph_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" shows "continuous_on S f \ closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" (is "?lhs = ?rhs") proof - have "?lhs" if ?rhs proof (clarsimp simp add: continuous_on_closed_gen [OF fim]) fix U assume U: "closedin (top_of_set T) U" have eq: "(S \ f -` U) = fst ` (((\x. Pair x (f x)) ` S) \ (S \ U))" by (force simp: image_iff) show "closedin (top_of_set S) (S \ f -` U)" by (simp add: U closedin_Int closedin_Times closed_map_fst [OF \compact T\] that eq) qed with continuous_closed_graph_gen assms show ?thesis by blast qed lemma continuous_closed_graph: fixes f :: "'a::topological_space \ 'b::real_normed_vector" assumes "closed S" and contf: "continuous_on S f" shows "closed ((\x. Pair x (f x)) ` S)" apply (rule closedin_closed_trans) apply (rule continuous_closed_graph_gen [OF contf subset_UNIV]) by (simp add: \closed S\ closed_Times) lemma continuous_from_closed_graph: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" and clo: "closed ((\x. Pair x (f x)) ` S)" shows "continuous_on S f" using fim clo by (auto intro: closed_subset simp: continuous_closed_graph_eq [OF \compact T\ fim]) lemma continuous_on_Un_local_open: assumes opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T f" shows "continuous_on (S \ T) f" using pasting_lemma [of "{S,T}" "top_of_set (S \ T)" id euclidean "\i. f" f] contf contg opS opT by (simp add: subtopology_subtopology) (metis inf.absorb2 openin_imp_subset) lemma continuous_on_cases_local_open: assumes opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T g" and fg: "\x. x \ S \ \P x \ x \ T \ P x \ f x = g x" shows "continuous_on (S \ T) (\x. if P x then f x else g x)" proof - have "\x. x \ S \ (if P x then f x else g x) = f x" "\x. x \ T \ (if P x then f x else g x) = g x" by (simp_all add: fg) then have "continuous_on S (\x. if P x then f x else g x)" "continuous_on T (\x. if P x then f x else g x)" by (simp_all add: contf contg cong: continuous_on_cong) then show ?thesis by (rule continuous_on_Un_local_open [OF opS opT]) qed lemma continuous_map_cases_le: assumes contp: "continuous_map X euclideanreal p" and contq: "continuous_map X euclideanreal q" and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" shows "continuous_map X Y (\x. if p x \ q x then f x else g x)" proof - have "continuous_map X Y (\x. if q x - p x \ {0..} then f x else g x)" proof (rule continuous_map_cases_function) show "continuous_map X euclideanreal (\x. q x - p x)" by (intro contp contq continuous_intros) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0..}}) Y f" by (simp add: contf) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0..})}) Y g" by (simp add: contg flip: Compl_eq_Diff_UNIV) qed (auto simp: fg) then show ?thesis by simp qed lemma continuous_map_cases_lt: assumes contp: "continuous_map X euclideanreal p" and contq: "continuous_map X euclideanreal q" and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" shows "continuous_map X Y (\x. if p x < q x then f x else g x)" proof - have "continuous_map X Y (\x. if q x - p x \ {0<..} then f x else g x)" proof (rule continuous_map_cases_function) show "continuous_map X euclideanreal (\x. q x - p x)" by (intro contp contq continuous_intros) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0<..}}) Y f" by (simp add: contf) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0<..})}) Y g" by (simp add: contg flip: Compl_eq_Diff_UNIV) qed (auto simp: fg) then show ?thesis by simp qed subsection\<^marker>\tag unimportant\\The union of two collinear segments is another segment\ proposition\<^marker>\tag unimportant\ in_convex_hull_exchange: fixes a :: "'a::euclidean_space" assumes a: "a \ convex hull S" and xS: "x \ convex hull S" obtains b where "b \ S" "x \ convex hull (insert a (S - {b}))" proof (cases "a \ S") case True with xS insert_Diff that show ?thesis by fastforce next case False show ?thesis proof (cases "finite S \ card S \ Suc (DIM('a))") case True then obtain u where u0: "\i. i \ S \ 0 \ u i" and u1: "sum u S = 1" and ua: "(\i\S. u i *\<^sub>R i) = a" using a by (auto simp: convex_hull_finite) obtain v where v0: "\i. i \ S \ 0 \ v i" and v1: "sum v S = 1" and vx: "(\i\S. v i *\<^sub>R i) = x" using True xS by (auto simp: convex_hull_finite) show ?thesis proof (cases "\b. b \ S \ v b = 0") case True then obtain b where b: "b \ S" "v b = 0" by blast show ?thesis proof have fin: "finite (insert a (S - {b}))" using sum.infinite v1 by fastforce show "x \ convex hull insert a (S - {b})" unfolding convex_hull_finite [OF fin] mem_Collect_eq proof (intro conjI exI ballI) have "(\x \ insert a (S - {b}). if x = a then 0 else v x) = (\x \ S - {b}. if x = a then 0 else v x)" apply (rule sum.mono_neutral_right) using fin by auto also have "... = (\x \ S - {b}. v x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x)" by (metis \v b = 0\ diff_zero sum.infinite sum_diff1 u1 zero_neq_one) finally show "(\x\insert a (S - {b}). if x = a then 0 else v x) = 1" by (simp add: v1) show "\x. x \ insert a (S - {b}) \ 0 \ (if x = a then 0 else v x)" by (auto simp: v0) have "(\x \ insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = (\x \ S - {b}. (if x = a then 0 else v x) *\<^sub>R x)" apply (rule sum.mono_neutral_right) using fin by auto also have "... = (\x \ S - {b}. v x *\<^sub>R x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x *\<^sub>R x)" by (metis (no_types, lifting) b(2) diff_zero fin finite.emptyI finite_Diff2 finite_insert scale_eq_0_iff sum_diff1) finally show "(\x\insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = x" by (simp add: vx) qed qed (rule \b \ S\) next case False have le_Max: "u i / v i \ Max ((\i. u i / v i) ` S)" if "i \ S" for i by (simp add: True that) have "Max ((\i. u i / v i) ` S) \ (\i. u i / v i) ` S" using True v1 by (auto intro: Max_in) then obtain b where "b \ S" and beq: "Max ((\b. u b / v b) ` S) = u b / v b" by blast then have "0 \ u b / v b" using le_Max beq divide_le_0_iff le_numeral_extra(2) sum_nonpos u1 by (metis False eq_iff v0) then have "0 < u b" "0 < v b" using False \b \ S\ u0 v0 by force+ have fin: "finite (insert a (S - {b}))" using sum.infinite v1 by fastforce show ?thesis proof show "x \ convex hull insert a (S - {b})" unfolding convex_hull_finite [OF fin] mem_Collect_eq proof (intro conjI exI ballI) have "(\x \ insert a (S - {b}). if x=a then v b / u b else v x - (v b / u b) * u x) = v b / u b + (\x \ S - {b}. v x - (v b / u b) * u x)" using \a \ S\ \b \ S\ True apply simp apply (rule sum.cong, auto) done also have "... = v b / u b + (\x \ S - {b}. v x) - (v b / u b) * (\x \ S - {b}. u x)" by (simp add: Groups_Big.sum_subtractf sum_distrib_left) also have "... = (\x\S. v x)" using \0 < u b\ True by (simp add: Groups_Big.sum_diff1 u1 field_simps) finally show "sum (\x. if x=a then v b / u b else v x - (v b / u b) * u x) (insert a (S - {b})) = 1" by (simp add: v1) show "0 \ (if i = a then v b / u b else v i - v b / u b * u i)" if "i \ insert a (S - {b})" for i using \0 < u b\ \0 < v b\ v0 [of i] le_Max [of i] beq that False by (auto simp: field_simps split: if_split_asm) have "(\x\insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = (v b / u b) *\<^sub>R a + (\x\S - {b}. (v x - v b / u b * u x) *\<^sub>R x)" using \a \ S\ \b \ S\ True apply simp apply (rule sum.cong, auto) done also have "... = (v b / u b) *\<^sub>R a + (\x \ S - {b}. v x *\<^sub>R x) - (v b / u b) *\<^sub>R (\x \ S - {b}. u x *\<^sub>R x)" by (simp add: Groups_Big.sum_subtractf scaleR_left_diff_distrib sum_distrib_left scale_sum_right) also have "... = (\x\S. v x *\<^sub>R x)" using \0 < u b\ True by (simp add: ua vx Groups_Big.sum_diff1 algebra_simps) finally show "(\x\insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = x" by (simp add: vx) qed qed (rule \b \ S\) qed next case False obtain T where "finite T" "T \ S" and caT: "card T \ Suc (DIM('a))" and xT: "x \ convex hull T" using xS by (auto simp: caratheodory [of S]) with False obtain b where b: "b \ S" "b \ T" by (metis antisym subsetI) show ?thesis proof show "x \ convex hull insert a (S - {b})" using \T \ S\ b by (blast intro: subsetD [OF hull_mono xT]) qed (rule \b \ S\) qed qed lemma convex_hull_exchange_Union: fixes a :: "'a::euclidean_space" assumes "a \ convex hull S" shows "convex hull S = (\b \ S. convex hull (insert a (S - {b})))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (blast intro: in_convex_hull_exchange [OF assms]) show "?rhs \ ?lhs" proof clarify fix x b assume"b \ S" "x \ convex hull insert a (S - {b})" then show "x \ convex hull S" if "b \ S" by (metis (no_types) that assms order_refl hull_mono hull_redundant insert_Diff_single insert_subset subsetCE) qed qed lemma Un_closed_segment: fixes a :: "'a::euclidean_space" assumes "b \ closed_segment a c" shows "closed_segment a b \ closed_segment b c = closed_segment a c" proof (cases "c = a") case True with assms show ?thesis by simp next case False with assms have "convex hull {a, b} \ convex hull {b, c} = (\ba\{a, c}. convex hull insert b ({a, c} - {ba}))" by (auto simp: insert_Diff_if insert_commute) then show ?thesis using convex_hull_exchange_Union by (metis assms segment_convex_hull) qed lemma Un_open_segment: fixes a :: "'a::euclidean_space" assumes "b \ open_segment a c" shows "open_segment a b \ {b} \ open_segment b c = open_segment a c" proof - have b: "b \ closed_segment a c" by (simp add: assms open_closed_segment) have *: "open_segment a c \ insert b (open_segment a b \ open_segment b c)" if "{b,c,a} \ open_segment a b \ open_segment b c = {c,a} \ open_segment a c" proof - have "insert a (insert c (insert b (open_segment a b \ open_segment b c))) = insert a (insert c (open_segment a c))" using that by (simp add: insert_commute) then show ?thesis by (metis (no_types) Diff_cancel Diff_eq_empty_iff Diff_insert2 open_segment_def) qed show ?thesis using Un_closed_segment [OF b] apply (simp add: closed_segment_eq_open) apply (rule equalityI) using assms apply (simp add: b subset_open_segment) using * by (simp add: insert_commute) qed subsection\Covering an open set by a countable chain of compact sets\ proposition open_Union_compact_subsets: fixes S :: "'a::euclidean_space set" assumes "open S" obtains C where "\n. compact(C n)" "\n. C n \ S" "\n. C n \ interior(C(Suc n))" "\(range C) = S" "\K. \compact K; K \ S\ \ \N. \n\N. K \ (C n)" proof (cases "S = {}") case True then show ?thesis by (rule_tac C = "\n. {}" in that) auto next case False then obtain a where "a \ S" by auto let ?C = "\n. cball a (real n) - (\x \ -S. \e \ ball 0 (1 / real(Suc n)). {x + e})" have "\N. \n\N. K \ (f n)" if "\n. compact(f n)" and sub_int: "\n. f n \ interior (f(Suc n))" and eq: "\(range f) = S" and "compact K" "K \ S" for f K proof - have *: "\n. f n \ (\n. interior (f n))" by (meson Sup_upper2 UNIV_I \\n. f n \ interior (f (Suc n))\ image_iff) have mono: "\m n. m \ n \f m \ f n" by (meson dual_order.trans interior_subset lift_Suc_mono_le sub_int) obtain I where "finite I" and I: "K \ (\i\I. interior (f i))" proof (rule compactE_image [OF \compact K\]) show "K \ (\n. interior (f n))" using \K \ S\ \\(f ` UNIV) = S\ * by blast qed auto { fix n assume n: "Max I \ n" have "(\i\I. interior (f i)) \ f n" by (rule UN_least) (meson dual_order.trans interior_subset mono I Max_ge [OF \finite I\] n) then have "K \ f n" using I by auto } then show ?thesis by blast qed moreover have "\f. (\n. compact(f n)) \ (\n. (f n) \ S) \ (\n. (f n) \ interior(f(Suc n))) \ ((\(range f) = S))" proof (intro exI conjI allI) show "\n. compact (?C n)" by (auto simp: compact_diff open_sums) show "\n. ?C n \ S" by auto show "?C n \ interior (?C (Suc n))" for n proof (simp add: interior_diff, rule Diff_mono) show "cball a (real n) \ ball a (1 + real n)" by (simp add: cball_subset_ball_iff) have cl: "closed (\x\- S. \e\cball 0 (1 / (2 + real n)). {x + e})" using assms by (auto intro: closed_compact_sums) have "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \e \ cball 0 (1 / (2 + real n)). {x + e})" by (intro closure_minimal UN_mono ball_subset_cball order_refl cl) also have "... \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" apply (intro UN_mono order_refl) apply (simp add: cball_subset_ball_iff field_split_simps) done finally show "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" . qed have "S \ \ (range ?C)" proof fix x assume x: "x \ S" then obtain e where "e > 0" and e: "ball x e \ S" using assms open_contains_ball by blast then obtain N1 where "N1 > 0" and N1: "real N1 > 1/e" using reals_Archimedean2 by (metis divide_less_0_iff less_eq_real_def neq0_conv not_le of_nat_0 of_nat_1 of_nat_less_0_iff) obtain N2 where N2: "norm(x - a) \ real N2" by (meson real_arch_simple) have N12: "inverse((N1 + N2) + 1) \ inverse(N1)" using \N1 > 0\ by (auto simp: field_split_simps) have "x \ y + z" if "y \ S" "norm z < 1 / (1 + (real N1 + real N2))" for y z proof - have "e * real N1 < e * (1 + (real N1 + real N2))" by (simp add: \0 < e\) then have "1 / (1 + (real N1 + real N2)) < e" using N1 \e > 0\ by (metis divide_less_eq less_trans mult.commute of_nat_add of_nat_less_0_iff of_nat_Suc) then have "x - z \ ball x e" using that by simp then have "x - z \ S" using e by blast with that show ?thesis by auto qed with N2 show "x \ \ (range ?C)" by (rule_tac a = "N1+N2" in UN_I) (auto simp: dist_norm norm_minus_commute) qed then show "\ (range ?C) = S" by auto qed ultimately show ?thesis using that by metis qed subsection\Orthogonal complement\ definition\<^marker>\tag important\ orthogonal_comp ("_\<^sup>\" [80] 80) where "orthogonal_comp W \ {x. \y \ W. orthogonal y x}" proposition subspace_orthogonal_comp: "subspace (W\<^sup>\)" unfolding subspace_def orthogonal_comp_def orthogonal_def by (auto simp: inner_right_distrib) lemma orthogonal_comp_anti_mono: assumes "A \ B" shows "B\<^sup>\ \ A\<^sup>\" proof fix x assume x: "x \ B\<^sup>\" show "x \ orthogonal_comp A" using x unfolding orthogonal_comp_def by (simp add: orthogonal_def, metis assms in_mono) qed lemma orthogonal_comp_null [simp]: "{0}\<^sup>\ = UNIV" by (auto simp: orthogonal_comp_def orthogonal_def) lemma orthogonal_comp_UNIV [simp]: "UNIV\<^sup>\ = {0}" unfolding orthogonal_comp_def orthogonal_def by auto (use inner_eq_zero_iff in blast) lemma orthogonal_comp_subset: "U \ U\<^sup>\\<^sup>\" by (auto simp: orthogonal_comp_def orthogonal_def inner_commute) lemma subspace_sum_minimal: assumes "S \ U" "T \ U" "subspace U" shows "S + T \ U" proof fix x assume "x \ S + T" then obtain xs xt where "xs \ S" "xt \ T" "x = xs+xt" by (meson set_plus_elim) then show "x \ U" by (meson assms subsetCE subspace_add) qed proposition subspace_sum_orthogonal_comp: fixes U :: "'a :: euclidean_space set" assumes "subspace U" shows "U + U\<^sup>\ = UNIV" proof - obtain B where "B \ U" and ortho: "pairwise orthogonal B" "\x. x \ B \ norm x = 1" and "independent B" "card B = dim U" "span B = U" using orthonormal_basis_subspace [OF assms] by metis then have "finite B" by (simp add: indep_card_eq_dim_span) have *: "\x\B. \y\B. x \ y = (if x=y then 1 else 0)" using ortho norm_eq_1 by (auto simp: orthogonal_def pairwise_def) { fix v let ?u = "\b\B. (v \ b) *\<^sub>R b" have "v = ?u + (v - ?u)" by simp moreover have "?u \ U" by (metis (no_types, lifting) \span B = U\ assms subspace_sum span_base span_mul) moreover have "(v - ?u) \ U\<^sup>\" proof (clarsimp simp: orthogonal_comp_def orthogonal_def) fix y assume "y \ U" with \span B = U\ span_finite [OF \finite B\] obtain u where u: "y = (\b\B. u b *\<^sub>R b)" by auto have "b \ (v - ?u) = 0" if "b \ B" for b using that \finite B\ by (simp add: * algebra_simps inner_sum_right if_distrib [of "(*)v" for v] inner_commute cong: if_cong) then show "y \ (v - ?u) = 0" by (simp add: u inner_sum_left) qed ultimately have "v \ U + U\<^sup>\" using set_plus_intro by fastforce } then show ?thesis by auto qed lemma orthogonal_Int_0: assumes "subspace U" shows "U \ U\<^sup>\ = {0}" using orthogonal_comp_def orthogonal_self by (force simp: assms subspace_0 subspace_orthogonal_comp) lemma orthogonal_comp_self: fixes U :: "'a :: euclidean_space set" assumes "subspace U" shows "U\<^sup>\\<^sup>\ = U" proof have ssU': "subspace (U\<^sup>\)" by (simp add: subspace_orthogonal_comp) have "u \ U" if "u \ U\<^sup>\\<^sup>\" for u proof - obtain v w where "u = v+w" "v \ U" "w \ U\<^sup>\" using subspace_sum_orthogonal_comp [OF assms] set_plus_elim by blast then have "u-v \ U\<^sup>\" by simp moreover have "v \ U\<^sup>\\<^sup>\" using \v \ U\ orthogonal_comp_subset by blast then have "u-v \ U\<^sup>\\<^sup>\" by (simp add: subspace_diff subspace_orthogonal_comp that) ultimately have "u-v = 0" using orthogonal_Int_0 ssU' by blast with \v \ U\ show ?thesis by auto qed then show "U\<^sup>\\<^sup>\ \ U" by auto qed (use orthogonal_comp_subset in auto) lemma ker_orthogonal_comp_adjoint: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "f -` {0} = (range (adjoint f))\<^sup>\" apply (auto simp: orthogonal_comp_def orthogonal_def) apply (simp add: adjoint_works assms(1) inner_commute) by (metis adjoint_works all_zero_iff assms(1) inner_commute) subsection\<^marker>\tag unimportant\ \A non-injective linear function maps into a hyperplane.\ lemma linear_surj_adj_imp_inj: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" "surj (adjoint f)" shows "inj f" proof - have "\x. y = adjoint f x" for y using assms by (simp add: surjD) then show "inj f" using assms unfolding inj_on_def image_def by (metis (no_types) adjoint_works euclidean_eqI) qed \ \\<^url>\https://mathonline.wikidot.com/injectivity-and-surjectivity-of-the-adjoint-of-a-linear-map\\ lemma surj_adjoint_iff_inj [simp]: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "surj (adjoint f) \ inj f" proof assume "surj (adjoint f)" then show "inj f" by (simp add: assms linear_surj_adj_imp_inj) next assume "inj f" have "f -` {0} = {0}" using assms \inj f\ linear_0 linear_injective_0 by fastforce moreover have "f -` {0} = range (adjoint f)\<^sup>\" by (intro ker_orthogonal_comp_adjoint assms) ultimately have "range (adjoint f)\<^sup>\\<^sup>\ = UNIV" by (metis orthogonal_comp_null) then show "surj (adjoint f)" using adjoint_linear \linear f\ by (subst (asm) orthogonal_comp_self) (simp add: adjoint_linear linear_subspace_image) qed lemma inj_adjoint_iff_surj [simp]: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "inj (adjoint f) \ surj f" proof assume "inj (adjoint f)" have "(adjoint f) -` {0} = {0}" by (metis \inj (adjoint f)\ adjoint_linear assms surj_adjoint_iff_inj ker_orthogonal_comp_adjoint orthogonal_comp_UNIV) then have "(range(f))\<^sup>\ = {0}" by (metis (no_types, hide_lams) adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint set_zero) then show "surj f" by (metis \inj (adjoint f)\ adjoint_adjoint adjoint_linear assms surj_adjoint_iff_inj) next assume "surj f" then have "range f = (adjoint f -` {0})\<^sup>\" by (simp add: adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint) then have "{0} = adjoint f -` {0}" using \surj f\ adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint by force then show "inj (adjoint f)" by (simp add: \surj f\ adjoint_adjoint adjoint_linear assms linear_surj_adj_imp_inj) qed lemma linear_singular_into_hyperplane: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" shows "\ inj f \ (\a. a \ 0 \ (\x. a \ f x = 0))" (is "_ = ?rhs") proof assume "\inj f" then show ?rhs using all_zero_iff by (metis (no_types, hide_lams) adjoint_clauses(2) adjoint_linear assms linear_injective_0 linear_injective_imp_surjective linear_surj_adj_imp_inj) next assume ?rhs then show "\inj f" by (metis assms linear_injective_isomorphism all_zero_iff) qed lemma linear_singular_image_hyperplane: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" "\inj f" obtains a where "a \ 0" "\S. f ` S \ {x. a \ x = 0}" using assms by (fastforce simp add: linear_singular_into_hyperplane) end