# HG changeset patch # User paulson # Date 1475496061 -3600 # Node ID 0de4736dad8b37279f34cdca6094e9de9c8cba4c # Parent 5649a993666d9fcc3cec798fe324289109e5d5c6 new theorems including the theory FurtherTopology diff -r 5649a993666d -r 0de4736dad8b src/HOL/Analysis/Analysis.thy --- a/src/HOL/Analysis/Analysis.thy Sun Oct 02 21:05:14 2016 +0200 +++ b/src/HOL/Analysis/Analysis.thy Mon Oct 03 13:01:01 2016 +0100 @@ -10,6 +10,7 @@ Bounded_Continuous_Function Weierstrass_Theorems Polytope + FurtherTopology Poly_Roots Conformal_Mappings Generalised_Binomial_Theorem diff -r 5649a993666d -r 0de4736dad8b src/HOL/Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Analysis/Brouwer_Fixpoint.thy Sun Oct 02 21:05:14 2016 +0200 +++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy Mon Oct 03 13:01:01 2016 +0100 @@ -1975,7 +1975,7 @@ text \So we get the no-retraction theorem.\ -lemma no_retraction_cball: +theorem no_retraction_cball: fixes a :: "'a::euclidean_space" assumes "e > 0" shows "\ (frontier (cball a e) retract_of (cball a e))" @@ -2001,6 +2001,26 @@ using x assms by auto qed +corollary contractible_sphere: + fixes a :: "'a::euclidean_space" + shows "contractible(sphere a r) \ r \ 0" +proof (cases "0 < r") + case True + then show ?thesis + unfolding contractible_def nullhomotopic_from_sphere_extension + using no_retraction_cball [OF True, of a] + by (auto simp: retract_of_def retraction_def) +next + case False + then show ?thesis + unfolding contractible_def nullhomotopic_from_sphere_extension + apply (simp add: not_less) + apply (rule_tac x=id in exI) + apply (auto simp: continuous_on_def) + apply (meson dist_not_less_zero le_less less_le_trans) + done +qed + subsection\Retractions\ lemma retraction: diff -r 5649a993666d -r 0de4736dad8b src/HOL/Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Analysis/Convex_Euclidean_Space.thy Sun Oct 02 21:05:14 2016 +0200 +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Mon Oct 03 13:01:01 2016 +0100 @@ -7495,6 +7495,11 @@ by (auto simp: closed_segment_commute) qed +lemma open_segment_eq_real_ivl: + fixes a b::real + shows "open_segment a b = (if a \ b then {a<..x. (v - u) * x + u) ` {0..1}" by (simp add: add.commute [of u] image_affinity_atLeastAtMost [where c=u] closed_segment_eq_real_ivl) @@ -11353,6 +11358,81 @@ 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) diff -r 5649a993666d -r 0de4736dad8b src/HOL/Analysis/FurtherTopology.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/FurtherTopology.thy Mon Oct 03 13:01:01 2016 +0100 @@ -0,0 +1,1891 @@ +section \Extending Continous Maps, etc..\ + +text\Ported from HOL Light (moretop.ml) by L C Paulson\ + +theory "FurtherTopology" + imports Equivalence_Lebesgue_Henstock_Integration Weierstrass_Theorems Polytope + +begin + +subsection\A map from a sphere to a higher dimensional sphere is nullhomotopic\ + +lemma spheremap_lemma1: + fixes f :: "'a::euclidean_space \ 'a::euclidean_space" + assumes "subspace S" "subspace T" and dimST: "dim S < dim T" + and "S \ T" + and diff_f: "f differentiable_on sphere 0 1 \ S" + shows "f ` (sphere 0 1 \ S) \ sphere 0 1 \ T" +proof + assume fim: "f ` (sphere 0 1 \ S) = sphere 0 1 \ T" + have inS: "\x. \x \ S; x \ 0\ \ (x /\<^sub>R norm x) \ S" + using subspace_mul \subspace S\ by blast + have subS01: "(\x. x /\<^sub>R norm x) ` (S - {0}) \ sphere 0 1 \ S" + using \subspace S\ subspace_mul by fastforce + then have diff_f': "f differentiable_on (\x. x /\<^sub>R norm x) ` (S - {0})" + by (rule differentiable_on_subset [OF diff_f]) + define g where "g \ \x. norm x *\<^sub>R f(inverse(norm x) *\<^sub>R x)" + have gdiff: "g differentiable_on S - {0}" + unfolding g_def + by (rule diff_f' derivative_intros differentiable_on_compose [where f=f] | force)+ + have geq: "g ` (S - {0}) = T - {0}" + proof + have "g ` (S - {0}) \ T" + apply (auto simp: g_def subspace_mul [OF \subspace T\]) + apply (metis (mono_tags, lifting) DiffI subS01 subspace_mul [OF \subspace T\] fim image_subset_iff inf_le2 singletonD) + done + moreover have "g ` (S - {0}) \ UNIV - {0}" + proof (clarsimp simp: g_def) + fix y + assume "y \ S" and f0: "f (y /\<^sub>R norm y) = 0" + then have "y \ 0 \ y /\<^sub>R norm y \ sphere 0 1 \ S" + by (auto simp: subspace_mul [OF \subspace S\]) + then show "y = 0" + by (metis fim f0 Int_iff image_iff mem_sphere_0 norm_eq_zero zero_neq_one) + qed + ultimately show "g ` (S - {0}) \ T - {0}" + by auto + next + have *: "sphere 0 1 \ T \ f ` (sphere 0 1 \ S)" + using fim by (simp add: image_subset_iff) + have "x \ (\x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})" + if "x \ T" "x \ 0" for x + proof - + have "x /\<^sub>R norm x \ T" + using \subspace T\ subspace_mul that by blast + then show ?thesis + using * [THEN subsetD, of "x /\<^sub>R norm x"] that apply clarsimp + apply (rule_tac x="norm x *\<^sub>R xa" in image_eqI, simp) + apply (metis norm_eq_zero right_inverse scaleR_one scaleR_scaleR) + using \subspace S\ subspace_mul apply force + done + qed + then have "T - {0} \ (\x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})" + by force + then show "T - {0} \ g ` (S - {0})" + by (simp add: g_def) + qed + define T' where "T' \ {y. \x \ T. orthogonal x y}" + have "subspace T'" + by (simp add: subspace_orthogonal_to_vectors T'_def) + have dim_eq: "dim T' + dim T = DIM('a)" + using dim_subspace_orthogonal_to_vectors [of T UNIV] \subspace T\ + by (simp add: dim_UNIV T'_def) + have "\v1 v2. v1 \ span T \ (\w \ span T. orthogonal v2 w) \ x = v1 + v2" for x + by (force intro: orthogonal_subspace_decomp_exists [of T x]) + then obtain p1 p2 where p1span: "p1 x \ span T" + and "\w. w \ span T \ orthogonal (p2 x) w" + and eq: "p1 x + p2 x = x" for x + by metis + then have p1: "\z. p1 z \ T" and ortho: "\w. w \ T \ orthogonal (p2 x) w" for x + using span_eq \subspace T\ by blast+ + then have p2: "\z. p2 z \ T'" + by (simp add: T'_def orthogonal_commute) + have p12_eq: "\x y. \x \ T; y \ T'\ \ p1(x + y) = x \ p2(x + y) = y" + proof (rule orthogonal_subspace_decomp_unique [OF eq p1span, where T=T']) + show "\x y. \x \ T; y \ T'\ \ p2 (x + y) \ span T'" + using span_eq p2 \subspace T'\ by blast + show "\a b. \a \ T; b \ T'\ \ orthogonal a b" + using T'_def by blast + qed (auto simp: span_superset) + then have "\c x. p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \ p2 (c *\<^sub>R x) = c *\<^sub>R p2 x" + by (metis eq \subspace T\ \subspace T'\ p1 p2 scaleR_add_right subspace_mul) + moreover have lin_add: "\x y. p1 (x + y) = p1 x + p1 y \ p2 (x + y) = p2 x + p2 y" + proof (rule orthogonal_subspace_decomp_unique [OF _ p1span, where T=T']) + show "\x y. p1 (x + y) + p2 (x + y) = p1 x + p1 y + (p2 x + p2 y)" + by (simp add: add.assoc add.left_commute eq) + show "\a b. \a \ T; b \ T'\ \ orthogonal a b" + using T'_def by blast + qed (auto simp: p1span p2 span_superset subspace_add) + ultimately have "linear p1" "linear p2" + by unfold_locales auto + have "(\z. g (p1 z)) differentiable_on {x + y |x y. x \ S - {0} \ y \ T'}" + apply (rule differentiable_on_compose [where f=g]) + apply (rule linear_imp_differentiable_on [OF \linear p1\]) + apply (rule differentiable_on_subset [OF gdiff]) + using p12_eq \S \ T\ apply auto + done + then have diff: "(\x. g (p1 x) + p2 x) differentiable_on {x + y |x y. x \ S - {0} \ y \ T'}" + by (intro derivative_intros linear_imp_differentiable_on [OF \linear p2\]) + have "dim {x + y |x y. x \ S - {0} \ y \ T'} \ dim {x + y |x y. x \ S \ y \ T'}" + by (blast intro: dim_subset) + also have "... = dim S + dim T' - dim (S \ T')" + using dim_sums_Int [OF \subspace S\ \subspace T'\] + by (simp add: algebra_simps) + also have "... < DIM('a)" + using dimST dim_eq by auto + finally have neg: "negligible {x + y |x y. x \ S - {0} \ y \ T'}" + by (rule negligible_lowdim) + have "negligible ((\x. g (p1 x) + p2 x) ` {x + y |x y. x \ S - {0} \ y \ T'})" + by (rule negligible_differentiable_image_negligible [OF order_refl neg diff]) + then have "negligible {x + y |x y. x \ g ` (S - {0}) \ y \ T'}" + proof (rule negligible_subset) + have "\t' \ T'; s \ S; s \ 0\ + \ g s + t' \ (\x. g (p1 x) + p2 x) ` + {x + t' |x t'. x \ S \ x \ 0 \ t' \ T'}" for t' s + apply (rule_tac x="s + t'" in image_eqI) + using \S \ T\ p12_eq by auto + then show "{x + y |x y. x \ g ` (S - {0}) \ y \ T'} + \ (\x. g (p1 x) + p2 x) ` {x + y |x y. x \ S - {0} \ y \ T'}" + by auto + qed + moreover have "- T' \ {x + y |x y. x \ g ` (S - {0}) \ y \ T'}" + proof clarsimp + fix z assume "z \ T'" + show "\x y. z = x + y \ x \ g ` (S - {0}) \ y \ T'" + apply (rule_tac x="p1 z" in exI) + apply (rule_tac x="p2 z" in exI) + apply (simp add: p1 eq p2 geq) + by (metis \z \ T'\ add.left_neutral eq p2) + qed + ultimately have "negligible (-T')" + using negligible_subset by blast + moreover have "negligible T'" + using negligible_lowdim + by (metis add.commute assms(3) diff_add_inverse2 diff_self_eq_0 dim_eq le_add1 le_antisym linordered_semidom_class.add_diff_inverse not_less0) + ultimately have "negligible (-T' \ T')" + by (metis negligible_Un_eq) + then show False + using negligible_Un_eq non_negligible_UNIV by simp +qed + + +lemma spheremap_lemma2: + fixes f :: "'a::euclidean_space \ 'a::euclidean_space" + assumes ST: "subspace S" "subspace T" "dim S < dim T" + and "S \ T" + and contf: "continuous_on (sphere 0 1 \ S) f" + and fim: "f ` (sphere 0 1 \ S) \ sphere 0 1 \ T" + shows "\c. homotopic_with (\x. True) (sphere 0 1 \ S) (sphere 0 1 \ T) f (\x. c)" +proof - + have [simp]: "\x. \norm x = 1; x \ S\ \ norm (f x) = 1" + using fim by (simp add: image_subset_iff) + have "compact (sphere 0 1 \ S)" + by (simp add: \subspace S\ closed_subspace compact_Int_closed) + then obtain g where pfg: "polynomial_function g" and gim: "g ` (sphere 0 1 \ S) \ T" + and g12: "\x. x \ sphere 0 1 \ S \ norm(f x - g x) < 1/2" + apply (rule Stone_Weierstrass_polynomial_function_subspace [OF _ contf _ \subspace T\, of "1/2"]) + using fim apply auto + done + have gnz: "g x \ 0" if "x \ sphere 0 1 \ S" for x + proof - + have "norm (f x) = 1" + using fim that by (simp add: image_subset_iff) + then show ?thesis + using g12 [OF that] by auto + qed + have diffg: "g differentiable_on sphere 0 1 \ S" + by (metis pfg differentiable_on_polynomial_function) + define h where "h \ \x. inverse(norm(g x)) *\<^sub>R g x" + have h: "x \ sphere 0 1 \ S \ h x \ sphere 0 1 \ T" for x + unfolding h_def + using gnz [of x] + by (auto simp: subspace_mul [OF \subspace T\] subsetD [OF gim]) + have diffh: "h differentiable_on sphere 0 1 \ S" + unfolding h_def + apply (intro derivative_intros diffg differentiable_on_compose [OF diffg]) + using gnz apply auto + done + have homfg: "homotopic_with (\z. True) (sphere 0 1 \ S) (T - {0}) f g" + proof (rule homotopic_with_linear [OF contf]) + show "continuous_on (sphere 0 1 \ S) g" + using pfg by (simp add: differentiable_imp_continuous_on diffg) + next + have non0fg: "0 \ closed_segment (f x) (g x)" if "norm x = 1" "x \ S" for x + proof - + have "f x \ sphere 0 1" + using fim that by (simp add: image_subset_iff) + moreover have "norm(f x - g x) < 1/2" + apply (rule g12) + using that by force + ultimately show ?thesis + by (auto simp: norm_minus_commute dest: segment_bound) + qed + show "\x. x \ sphere 0 1 \ S \ closed_segment (f x) (g x) \ T - {0}" + apply (simp add: subset_Diff_insert non0fg) + apply (simp add: segment_convex_hull) + apply (rule hull_minimal) + using fim image_eqI gim apply force + apply (rule subspace_imp_convex [OF \subspace T\]) + done + qed + obtain d where d: "d \ (sphere 0 1 \ T) - h ` (sphere 0 1 \ S)" + using h spheremap_lemma1 [OF ST \S \ T\ diffh] by force + then have non0hd: "0 \ closed_segment (h x) (- d)" if "norm x = 1" "x \ S" for x + using midpoint_between [of 0 "h x" "-d"] that h [of x] + by (auto simp: between_mem_segment midpoint_def) + have conth: "continuous_on (sphere 0 1 \ S) h" + using differentiable_imp_continuous_on diffh by blast + have hom_hd: "homotopic_with (\z. True) (sphere 0 1 \ S) (T - {0}) h (\x. -d)" + apply (rule homotopic_with_linear [OF conth continuous_on_const]) + apply (simp add: subset_Diff_insert non0hd) + apply (simp add: segment_convex_hull) + apply (rule hull_minimal) + using h d apply (force simp: subspace_neg [OF \subspace T\]) + apply (rule subspace_imp_convex [OF \subspace T\]) + done + have conT0: "continuous_on (T - {0}) (\y. inverse(norm y) *\<^sub>R y)" + by (intro continuous_intros) auto + have sub0T: "(\y. y /\<^sub>R norm y) ` (T - {0}) \ sphere 0 1 \ T" + by (fastforce simp: assms(2) subspace_mul) + obtain c where homhc: "homotopic_with (\z. True) (sphere 0 1 \ S) (sphere 0 1 \ T) h (\x. c)" + apply (rule_tac c="-d" in that) + apply (rule homotopic_with_eq) + apply (rule homotopic_compose_continuous_left [OF hom_hd conT0 sub0T]) + using d apply (auto simp: h_def) + done + show ?thesis + apply (rule_tac x=c in exI) + apply (rule homotopic_with_trans [OF _ homhc]) + apply (rule homotopic_with_eq) + apply (rule homotopic_compose_continuous_left [OF homfg conT0 sub0T]) + apply (auto simp: h_def) + done +qed + + +lemma spheremap_lemma3: + assumes "bounded S" "convex S" "subspace U" and affSU: "aff_dim S \ dim U" + obtains T where "subspace T" "T \ U" "S \ {} \ aff_dim T = aff_dim S" + "(rel_frontier S) homeomorphic (sphere 0 1 \ T)" +proof (cases "S = {}") + case True + with \subspace U\ subspace_0 show ?thesis + by (rule_tac T = "{0}" in that) auto +next + case False + then obtain a where "a \ S" + by auto + then have affS: "aff_dim S = int (dim ((\x. -a+x) ` S))" + by (metis hull_inc aff_dim_eq_dim) + with affSU have "dim ((\x. -a+x) ` S) \ dim U" + by linarith + with choose_subspace_of_subspace + obtain T where "subspace T" "T \ span U" and dimT: "dim T = dim ((\x. -a+x) ` S)" . + show ?thesis + proof (rule that [OF \subspace T\]) + show "T \ U" + using span_eq \subspace U\ \T \ span U\ by blast + show "aff_dim T = aff_dim S" + using dimT \subspace T\ affS aff_dim_subspace by fastforce + show "rel_frontier S homeomorphic sphere 0 1 \ T" + proof - + have "aff_dim (ball 0 1 \ T) = aff_dim (T)" + by (metis IntI interior_ball \subspace T\ aff_dim_convex_Int_nonempty_interior centre_in_ball empty_iff inf_commute subspace_0 subspace_imp_convex zero_less_one) + then have affS_eq: "aff_dim S = aff_dim (ball 0 1 \ T)" + using \aff_dim T = aff_dim S\ by simp + have "rel_frontier S homeomorphic rel_frontier(ball 0 1 \ T)" + apply (rule homeomorphic_rel_frontiers_convex_bounded_sets [OF \convex S\ \bounded S\]) + apply (simp add: \subspace T\ convex_Int subspace_imp_convex) + apply (simp add: bounded_Int) + apply (rule affS_eq) + done + also have "... = frontier (ball 0 1) \ T" + apply (rule convex_affine_rel_frontier_Int [OF convex_ball]) + apply (simp add: \subspace T\ subspace_imp_affine) + using \subspace T\ subspace_0 by force + also have "... = sphere 0 1 \ T" + by auto + finally show ?thesis . + qed + qed +qed + + +proposition inessential_spheremap_lowdim_gen: + fixes f :: "'M::euclidean_space \ 'a::euclidean_space" + assumes "convex S" "bounded S" "convex T" "bounded T" + and affST: "aff_dim S < aff_dim T" + and contf: "continuous_on (rel_frontier S) f" + and fim: "f ` (rel_frontier S) \ rel_frontier T" + obtains c where "homotopic_with (\z. True) (rel_frontier S) (rel_frontier T) f (\x. c)" +proof (cases "S = {}") + case True + then show ?thesis + by (simp add: that) +next + case False + then show ?thesis + proof (cases "T = {}") + case True + then show ?thesis + using fim that by auto + next + case False + obtain T':: "'a set" + where "subspace T'" and affT': "aff_dim T' = aff_dim T" + and homT: "rel_frontier T homeomorphic sphere 0 1 \ T'" + apply (rule spheremap_lemma3 [OF \bounded T\ \convex T\ subspace_UNIV, where 'b='a]) + apply (simp add: dim_UNIV aff_dim_le_DIM) + using \T \ {}\ by blast + with homeomorphic_imp_homotopy_eqv + have relT: "sphere 0 1 \ T' homotopy_eqv rel_frontier T" + using homotopy_eqv_sym by blast + have "aff_dim S \ int (dim T')" + using affT' \subspace T'\ affST aff_dim_subspace by force + with spheremap_lemma3 [OF \bounded S\ \convex S\ \subspace T'\] \S \ {}\ + obtain S':: "'a set" where "subspace S'" "S' \ T'" + and affS': "aff_dim S' = aff_dim S" + and homT: "rel_frontier S homeomorphic sphere 0 1 \ S'" + by metis + with homeomorphic_imp_homotopy_eqv + have relS: "sphere 0 1 \ S' homotopy_eqv rel_frontier S" + using homotopy_eqv_sym by blast + have dimST': "dim S' < dim T'" + by (metis \S' \ T'\ \subspace S'\ \subspace T'\ affS' affST affT' less_irrefl not_le subspace_dim_equal) + have "\c. homotopic_with (\z. True) (rel_frontier S) (rel_frontier T) f (\x. c)" + apply (rule homotopy_eqv_homotopic_triviality_null_imp [OF relT contf fim]) + apply (rule homotopy_eqv_cohomotopic_triviality_null[OF relS, THEN iffD1, rule_format]) + apply (metis dimST' \subspace S'\ \subspace T'\ \S' \ T'\ spheremap_lemma2, blast) + done + with that show ?thesis by blast + qed +qed + +lemma inessential_spheremap_lowdim: + fixes f :: "'M::euclidean_space \ 'a::euclidean_space" + assumes + "DIM('M) < DIM('a)" and f: "continuous_on (sphere a r) f" "f ` (sphere a r) \ (sphere b s)" + obtains c where "homotopic_with (\z. True) (sphere a r) (sphere b s) f (\x. c)" +proof (cases "s \ 0") + case True then show ?thesis + by (meson nullhomotopic_into_contractible f contractible_sphere that) +next + case False + show ?thesis + proof (cases "r \ 0") + case True then show ?thesis + by (meson f nullhomotopic_from_contractible contractible_sphere that) + next + case False + with \~ s \ 0\ have "r > 0" "s > 0" by auto + show ?thesis + apply (rule inessential_spheremap_lowdim_gen [of "cball a r" "cball b s" f]) + using \0 < r\ \0 < s\ assms(1) + apply (simp_all add: f aff_dim_cball) + using that by blast + qed +qed + + + +subsection\ Some technical lemmas about extending maps from cell complexes.\ + +lemma extending_maps_Union_aux: + assumes fin: "finite \" + and "\S. S \ \ \ closed S" + and "\S T. \S \ \; T \ \; S \ T\ \ S \ T \ K" + and "\S. S \ \ \ \g. continuous_on S g \ g ` S \ T \ (\x \ S \ K. g x = h x)" + shows "\g. continuous_on (\\) g \ g ` (\\) \ T \ (\x \ \\ \ K. g x = h x)" +using assms +proof (induction \) + case empty show ?case by simp +next + case (insert S \) + then obtain f where contf: "continuous_on (S) f" and fim: "f ` S \ T" and feq: "\x \ S \ K. f x = h x" + by (meson insertI1) + obtain g where contg: "continuous_on (\\) g" and gim: "g ` \\ \ T" and geq: "\x \ \\ \ K. g x = h x" + using insert by auto + have fg: "f x = g x" if "x \ T" "T \ \" "x \ S" for x T + proof - + have "T \ S \ K \ S = T" + using that by (metis (no_types) insert.prems(2) insertCI) + then show ?thesis + using UnionI feq geq \S \ \\ subsetD that by fastforce + qed + show ?case + apply (rule_tac x="\x. if x \ S then f x else g x" in exI, simp) + apply (intro conjI continuous_on_cases) + apply (simp_all add: insert closed_Union contf contg) + using fim gim feq geq + apply (force simp: insert closed_Union contf contg inf_commute intro: fg)+ + done +qed + +lemma extending_maps_Union: + assumes fin: "finite \" + and "\S. S \ \ \ \g. continuous_on S g \ g ` S \ T \ (\x \ S \ K. g x = h x)" + and "\S. S \ \ \ closed S" + and K: "\X Y. \X \ \; Y \ \; ~ X \ Y; ~ Y \ X\ \ X \ Y \ K" + shows "\g. continuous_on (\\) g \ g ` (\\) \ T \ (\x \ \\ \ K. g x = h x)" +apply (simp add: Union_maximal_sets [OF fin, symmetric]) +apply (rule extending_maps_Union_aux) +apply (simp_all add: Union_maximal_sets [OF fin] assms) +by (metis K psubsetI) + + +lemma extend_map_lemma: + assumes "finite \" "\ \ \" "convex T" "bounded T" + and poly: "\X. X \ \ \ polytope X" + and aff: "\X. X \ \ - \ \ aff_dim X < aff_dim T" + and face: "\S T. \S \ \; T \ \\ \ (S \ T) face_of S \ (S \ T) face_of T" + and contf: "continuous_on (\\) f" and fim: "f ` (\\) \ rel_frontier T" + obtains g where "continuous_on (\\) g" "g ` (\\) \ rel_frontier T" "\x. x \ \\ \ g x = f x" +proof (cases "\ - \ = {}") + case True + then have "\\ \ \\" + by (simp add: Union_mono) + then show ?thesis + apply (rule_tac g=f in that) + using contf continuous_on_subset apply blast + using fim apply blast + by simp +next + case False + then have "0 \ aff_dim T" + by (metis aff aff_dim_empty aff_dim_geq aff_dim_negative_iff all_not_in_conv not_less) + then obtain i::nat where i: "int i = aff_dim T" + by (metis nonneg_eq_int) + have Union_empty_eq: "\{D. D = {} \ P D} = {}" for P :: "'a set \ bool" + by auto + have extendf: "\g. continuous_on (\(\ \ {D. \C \ \. D face_of C \ aff_dim D < i})) g \ + g ` (\ (\ \ {D. \C \ \. D face_of C \ aff_dim D < i})) \ rel_frontier T \ + (\x \ \\. g x = f x)" + if "i \ aff_dim T" for i::nat + using that + proof (induction i) + case 0 then show ?case + apply (simp add: Union_empty_eq) + apply (rule_tac x=f in exI) + apply (intro conjI) + using contf continuous_on_subset apply blast + using fim apply blast + by simp + next + case (Suc p) + with \bounded T\ have "rel_frontier T \ {}" + by (auto simp: rel_frontier_eq_empty affine_bounded_eq_lowdim [of T]) + then obtain t where t: "t \ rel_frontier T" by auto + have ple: "int p \ aff_dim T" using Suc.prems by force + obtain h where conth: "continuous_on (\(\ \ {D. \C \ \. D face_of C \ aff_dim D < p})) h" + and him: "h ` (\ (\ \ {D. \C \ \. D face_of C \ aff_dim D < p})) + \ rel_frontier T" + and heq: "\x. x \ \\ \ h x = f x" + using Suc.IH [OF ple] by auto + let ?Faces = "{D. \C \ \. D face_of C \ aff_dim D \ p}" + have extendh: "\g. continuous_on D g \ + g ` D \ rel_frontier T \ + (\x \ D \ \(\ \ {D. \C \ \. D face_of C \ aff_dim D < p}). g x = h x)" + if D: "D \ \ \ ?Faces" for D + proof (cases "D \ \(\ \ {D. \C \ \. D face_of C \ aff_dim D < p})") + case True + then show ?thesis + apply (rule_tac x=h in exI) + apply (intro conjI) + apply (blast intro: continuous_on_subset [OF conth]) + using him apply blast + by simp + next + case False + note notDsub = False + show ?thesis + proof (cases "\a. D = {a}") + case True + then obtain a where "D = {a}" by auto + with notDsub t show ?thesis + by (rule_tac x="\x. t" in exI) simp + next + case False + have "D \ {}" using notDsub by auto + have Dnotin: "D \ \ \ {D. \C \ \. D face_of C \ aff_dim D < p}" + using notDsub by auto + then have "D \ \" by simp + have "D \ ?Faces - {D. \C \ \. D face_of C \ aff_dim D < p}" + using Dnotin that by auto + then obtain C where "C \ \" "D face_of C" and affD: "aff_dim D = int p" + by auto + then have "bounded D" + using face_of_polytope_polytope poly polytope_imp_bounded by blast + then have [simp]: "\ affine D" + using affine_bounded_eq_trivial False \D \ {}\ \bounded D\ by blast + have "{F. F facet_of D} \ {E. E face_of C \ aff_dim E < int p}" + apply clarify + apply (metis \D face_of C\ affD eq_iff face_of_trans facet_of_def zle_diff1_eq) + done + moreover have "polyhedron D" + using \C \ \\ \D face_of C\ face_of_polytope_polytope poly polytope_imp_polyhedron by auto + ultimately have relf_sub: "rel_frontier D \ \ {E. E face_of C \ aff_dim E < p}" + by (simp add: rel_frontier_of_polyhedron Union_mono) + then have him_relf: "h ` rel_frontier D \ rel_frontier T" + using \C \ \\ him by blast + have "convex D" + by (simp add: \polyhedron D\ polyhedron_imp_convex) + have affD_lessT: "aff_dim D < aff_dim T" + using Suc.prems affD by linarith + have contDh: "continuous_on (rel_frontier D) h" + using \C \ \\ relf_sub by (blast intro: continuous_on_subset [OF conth]) + then have *: "(\c. homotopic_with (\x. True) (rel_frontier D) (rel_frontier T) h (\x. c)) = + (\g. continuous_on UNIV g \ range g \ rel_frontier T \ + (\x\rel_frontier D. g x = h x))" + apply (rule nullhomotopic_into_rel_frontier_extension [OF closed_rel_frontier]) + apply (simp_all add: assms rel_frontier_eq_empty him_relf) + done + have "(\c. homotopic_with (\x. True) (rel_frontier D) + (rel_frontier T) h (\x. c))" + by (metis inessential_spheremap_lowdim_gen + [OF \convex D\ \bounded D\ \convex T\ \bounded T\ affD_lessT contDh him_relf]) + then obtain g where contg: "continuous_on UNIV g" + and gim: "range g \ rel_frontier T" + and gh: "\x. x \ rel_frontier D \ g x = h x" + by (metis *) + have "D \ E \ rel_frontier D" + if "E \ \ \ {D. Bex \ (op face_of D) \ aff_dim D < int p}" for E + proof (rule face_of_subset_rel_frontier) + show "D \ E face_of D" + using that \C \ \\ \D face_of C\ face + apply auto + apply (meson face_of_Int_subface \\ \ \\ face_of_refl_eq poly polytope_imp_convex subsetD) + using face_of_Int_subface apply blast + done + show "D \ E \ D" + using that notDsub by auto + qed + then show ?thesis + apply (rule_tac x=g in exI) + apply (intro conjI ballI) + using continuous_on_subset contg apply blast + using gim apply blast + using gh by fastforce + qed + qed + have intle: "i < 1 + int j \ i \ int j" for i j + by auto + have "finite \" + using \finite \\ \\ \ \\ rev_finite_subset by blast + then have fin: "finite (\ \ ?Faces)" + apply simp + apply (rule_tac B = "\{{D. D face_of C}| C. C \ \}" in finite_subset) + by (auto simp: \finite \\ finite_polytope_faces poly) + have clo: "closed S" if "S \ \ \ ?Faces" for S + using that \\ \ \\ face_of_polytope_polytope poly polytope_imp_closed by blast + have K: "X \ Y \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < int p})" + if "X \ \ \ ?Faces" "Y \ \ \ ?Faces" "~ Y \ X" for X Y + proof - + have ff: "X \ Y face_of X \ X \ Y face_of Y" + if XY: "X face_of D" "Y face_of E" and DE: "D \ \" "E \ \" for D E + apply (rule face_of_Int_subface [OF _ _ XY]) + apply (auto simp: face DE) + done + show ?thesis + using that + apply auto + apply (drule_tac x="X \ Y" in spec, safe) + using ff face_of_imp_convex [of X] face_of_imp_convex [of Y] + apply (fastforce dest: face_of_aff_dim_lt) + by (meson face_of_trans ff) + qed + obtain g where "continuous_on (\(\ \ ?Faces)) g" + "g ` \(\ \ ?Faces) \ rel_frontier T" + "(\x \ \(\ \ ?Faces) \ + \(\ \ {D. \C\\. D face_of C \ aff_dim D < p}). g x = h x)" + apply (rule exE [OF extending_maps_Union [OF fin extendh clo K]], blast+) + done + then show ?case + apply (simp add: intle local.heq [symmetric], blast) + done + qed + have eq: "\(\ \ {D. \C \ \. D face_of C \ aff_dim D < i}) = \\" + proof + show "\(\ \ {D. \C\\. D face_of C \ aff_dim D < int i}) \ \\" + apply (rule Union_subsetI) + using \\ \ \\ face_of_imp_subset apply force + done + show "\\ \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < i})" + apply (rule Union_mono) + using face apply (fastforce simp: aff i) + done + qed + have "int i \ aff_dim T" by (simp add: i) + then show ?thesis + using extendf [of i] unfolding eq by (metis that) +qed + +lemma extend_map_lemma_cofinite0: + assumes "finite \" + and "pairwise (\S T. S \ T \ K) \" + and "\S. S \ \ \ \a g. a \ U \ continuous_on (S - {a}) g \ g ` (S - {a}) \ T \ (\x \ S \ K. g x = h x)" + and "\S. S \ \ \ closed S" + shows "\C g. finite C \ disjnt C U \ card C \ card \ \ + continuous_on (\\ - C) g \ g ` (\\ - C) \ T + \ (\x \ (\\ - C) \ K. g x = h x)" + using assms +proof induction + case empty then show ?case + by force +next + case (insert X \) + then have "closed X" and clo: "\X. X \ \ \ closed X" + and \: "\S. S \ \ \ \a g. a \ U \ continuous_on (S - {a}) g \ g ` (S - {a}) \ T \ (\x \ S \ K. g x = h x)" + and pwX: "\Y. Y \ \ \ Y \ X \ X \ Y \ K \ Y \ X \ K" + and pwF: "pairwise (\ S T. S \ T \ K) \" + by (simp_all add: pairwise_insert) + obtain C g where C: "finite C" "disjnt C U" "card C \ card \" + and contg: "continuous_on (\\ - C) g" + and gim: "g ` (\\ - C) \ T" + and gh: "\x. x \ (\\ - C) \ K \ g x = h x" + using insert.IH [OF pwF \ clo] by auto + obtain a f where "a \ U" + and contf: "continuous_on (X - {a}) f" + and fim: "f ` (X - {a}) \ T" + and fh: "(\x \ X \ K. f x = h x)" + using insert.prems by (meson insertI1) + show ?case + proof (intro exI conjI) + show "finite (insert a C)" + by (simp add: C) + show "disjnt (insert a C) U" + using C \a \ U\ by simp + show "card (insert a C) \ card (insert X \)" + by (simp add: C card_insert_if insert.hyps le_SucI) + have "closed (\\)" + using clo insert.hyps by blast + have "continuous_on (X - insert a C \ (\\ - insert a C)) (\x. if x \ X then f x else g x)" + apply (rule continuous_on_cases_local) + apply (simp_all add: closedin_closed) + using \closed X\ apply blast + using \closed (\\)\ apply blast + using contf apply (force simp: elim: continuous_on_subset) + using contg apply (force simp: elim: continuous_on_subset) + using fh gh insert.hyps pwX by fastforce + then show "continuous_on (\insert X \ - insert a C) (\a. if a \ X then f a else g a)" + by (blast intro: continuous_on_subset) + show "\x\(\insert X \ - insert a C) \ K. (if x \ X then f x else g x) = h x" + using gh by (auto simp: fh) + show "(\a. if a \ X then f a else g a) ` (\insert X \ - insert a C) \ T" + using fim gim by auto force + qed +qed + + +lemma extend_map_lemma_cofinite1: +assumes "finite \" + and \: "\X. X \ \ \ \a g. a \ U \ continuous_on (X - {a}) g \ g ` (X - {a}) \ T \ (\x \ X \ K. g x = h x)" + and clo: "\X. X \ \ \ closed X" + and K: "\X Y. \X \ \; Y \ \; ~(X \ Y); ~(Y \ X)\ \ X \ Y \ K" + obtains C g where "finite C" "disjnt C U" "card C \ card \" "continuous_on (\\ - C) g" + "g ` (\\ - C) \ T" + "\x. x \ (\\ - C) \ K \ g x = h x" +proof - + let ?\ = "{X \ \. \Y\\. \ X \ Y}" + have [simp]: "\?\ = \\" + by (simp add: Union_maximal_sets assms) + have fin: "finite ?\" + by (force intro: finite_subset [OF _ \finite \\]) + have pw: "pairwise (\ S T. S \ T \ K) ?\" + by (simp add: pairwise_def) (metis K psubsetI) + have "card {X \ \. \Y\\. \ X \ Y} \ card \" + by (simp add: \finite \\ card_mono) + moreover + obtain C g where "finite C \ disjnt C U \ card C \ card ?\ \ + continuous_on (\?\ - C) g \ g ` (\?\ - C) \ T + \ (\x \ (\?\ - C) \ K. g x = h x)" + apply (rule exE [OF extend_map_lemma_cofinite0 [OF fin pw, of U T h]]) + apply (fastforce intro!: clo \)+ + done + ultimately show ?thesis + by (rule_tac C=C and g=g in that) auto +qed + + +lemma extend_map_lemma_cofinite: + assumes "finite \" "\ \ \" and T: "convex T" "bounded T" + and poly: "\X. X \ \ \ polytope X" + and contf: "continuous_on (\\) f" and fim: "f ` (\\) \ rel_frontier T" + and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X \ (X \ Y) face_of Y" + and aff: "\X. X \ \ - \ \ aff_dim X \ aff_dim T" + obtains C g where + "finite C" "disjnt C (\\)" "card C \ card \" "continuous_on (\\ - C) g" + "g ` (\ \ - C) \ rel_frontier T" "\x. x \ \\ \ g x = f x" +proof - + define \ where "\ \ \ \ {D. \C \ \ - \. D face_of C \ aff_dim D < aff_dim T}" + have "finite \" + using assms finite_subset by blast + moreover have "finite (\{{D. D face_of C} |C. C \ \})" + apply (rule finite_Union) + apply (simp add: \finite \\) + using finite_polytope_faces poly by auto + ultimately have "finite \" + apply (simp add: \_def) + apply (rule finite_subset [of _ "\ {{D. D face_of C} | C. C \ \}"], auto) + done + have *: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X \ X \ Y face_of Y" + unfolding \_def + apply (elim UnE bexE CollectE DiffE) + using subsetD [OF \\ \ \\] apply (simp_all add: face) + apply (meson subsetD [OF \\ \ \\] face face_of_Int_subface face_of_imp_subset face_of_refl poly polytope_imp_convex)+ + done + obtain h where conth: "continuous_on (\\) h" and him: "h ` (\\) \ rel_frontier T" + and hf: "\x. x \ \\ \ h x = f x" + using \finite \\ + unfolding \_def + apply (rule extend_map_lemma [OF _ Un_upper1 T _ _ _ contf fim]) + using \\ \ \\ face_of_polytope_polytope poly apply fastforce + using * apply (auto simp: \_def) + done + have "bounded (\\)" + using \finite \\ \\ \ \\ poly polytope_imp_bounded by blast + then have "\\ \ UNIV" + by auto + then obtain a where a: "a \ \\" + by blast + have \: "\a g. a \ \\ \ continuous_on (D - {a}) g \ + g ` (D - {a}) \ rel_frontier T \ (\x \ D \ \\. g x = h x)" + if "D \ \" for D + proof (cases "D \ \\") + case True + then show ?thesis + apply (rule_tac x=a in exI) + apply (rule_tac x=h in exI) + using him apply (blast intro!: \a \ \\\ continuous_on_subset [OF conth]) + + done + next + case False + note D_not_subset = False + show ?thesis + proof (cases "D \ \") + case True + with D_not_subset show ?thesis + by (auto simp: \_def) + next + case False + then have affD: "aff_dim D \ aff_dim T" + by (simp add: \D \ \\ aff) + show ?thesis + proof (cases "rel_interior D = {}") + case True + with \D \ \\ poly a show ?thesis + by (force simp: rel_interior_eq_empty polytope_imp_convex) + next + case False + then obtain b where brelD: "b \ rel_interior D" + by blast + have "polyhedron D" + by (simp add: poly polytope_imp_polyhedron that) + have "rel_frontier D retract_of affine hull D - {b}" + by (simp add: rel_frontier_retract_of_punctured_affine_hull poly polytope_imp_bounded polytope_imp_convex that brelD) + then obtain r where relfD: "rel_frontier D \ affine hull D - {b}" + and contr: "continuous_on (affine hull D - {b}) r" + and rim: "r ` (affine hull D - {b}) \ rel_frontier D" + and rid: "\x. x \ rel_frontier D \ r x = x" + by (auto simp: retract_of_def retraction_def) + show ?thesis + proof (intro exI conjI ballI) + show "b \ \\" + proof clarify + fix E + assume "b \ E" "E \ \" + then have "E \ D face_of E \ E \ D face_of D" + using \\ \ \\ face that by auto + with face_of_subset_rel_frontier \E \ \\ \b \ E\ brelD rel_interior_subset [of D] + D_not_subset rel_frontier_def \_def + show False + by blast + qed + have "r ` (D - {b}) \ r ` (affine hull D - {b})" + by (simp add: Diff_mono hull_subset image_mono) + also have "... \ rel_frontier D" + by (rule rim) + also have "... \ \{E. E face_of D \ aff_dim E < aff_dim T}" + using affD + by (force simp: rel_frontier_of_polyhedron [OF \polyhedron D\] facet_of_def) + also have "... \ \(\)" + using D_not_subset \_def that by fastforce + finally have rsub: "r ` (D - {b}) \ \(\)" . + show "continuous_on (D - {b}) (h \ r)" + apply (intro conjI \b \ \\\ continuous_on_compose) + apply (rule continuous_on_subset [OF contr]) + apply (simp add: Diff_mono hull_subset) + apply (rule continuous_on_subset [OF conth rsub]) + done + show "(h \ r) ` (D - {b}) \ rel_frontier T" + using brelD him rsub by fastforce + show "(h \ r) x = h x" if x: "x \ D \ \\" for x + proof - + consider A where "x \ D" "A \ \" "x \ A" + | A B where "x \ D" "A face_of B" "B \ \" "B \ \" "aff_dim A < aff_dim T" "x \ A" + using x by (auto simp: \_def) + then have xrel: "x \ rel_frontier D" + proof cases + case 1 show ?thesis + proof (rule face_of_subset_rel_frontier [THEN subsetD]) + show "D \ A face_of D" + using \A \ \\ \\ \ \\ face \D \ \\ by blast + show "D \ A \ D" + using \A \ \\ D_not_subset \_def by blast + qed (auto simp: 1) + next + case 2 show ?thesis + proof (rule face_of_subset_rel_frontier [THEN subsetD]) + show "D \ A face_of D" + apply (rule face_of_Int_subface [of D B _ A, THEN conjunct1]) + apply (simp_all add: 2 \D \ \\ face) + apply (simp add: \polyhedron D\ polyhedron_imp_convex face_of_refl) + done + show "D \ A \ D" + using "2" D_not_subset \_def by blast + qed (auto simp: 2) + qed + show ?thesis + by (simp add: rid xrel) + qed + qed + qed + qed + qed + have clo: "\S. S \ \ \ closed S" + by (simp add: poly polytope_imp_closed) + obtain C g where "finite C" "disjnt C (\\)" "card C \ card \" "continuous_on (\\ - C) g" + "g ` (\\ - C) \ rel_frontier T" + and gh: "\x. x \ (\\ - C) \ \\ \ g x = h x" + proof (rule extend_map_lemma_cofinite1 [OF \finite \\ \ clo]) + show "X \ Y \ \\" if XY: "X \ \" "Y \ \" and "\ X \ Y" "\ Y \ X" for X Y + proof (cases "X \ \") + case True + then show ?thesis + by (auto simp: \_def) + next + case False + have "X \ Y \ X" + using \\ X \ Y\ by blast + with XY + show ?thesis + by (clarsimp simp: \_def) + (metis Diff_iff Int_iff aff antisym_conv face face_of_aff_dim_lt face_of_refl + not_le poly polytope_imp_convex) + qed + qed (blast)+ + with \\ \ \\ show ?thesis + apply (rule_tac C=C and g=g in that) + apply (auto simp: disjnt_def hf [symmetric] \_def intro!: gh) + done +qed + +text\The next two proofs are similar\ +theorem extend_map_cell_complex_to_sphere: + assumes "finite \" and S: "S \ \\" "closed S" and T: "convex T" "bounded T" + and poly: "\X. X \ \ \ polytope X" + and aff: "\X. X \ \ \ aff_dim X < aff_dim T" + and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X \ (X \ Y) face_of Y" + and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier T" + obtains g where "continuous_on (\\) g" + "g ` (\\) \ rel_frontier T" "\x. x \ S \ g x = f x" +proof - + obtain V g where "S \ V" "open V" "continuous_on V g" and gim: "g ` V \ rel_frontier T" and gf: "\x. x \ S \ g x = f x" + using neighbourhood_extension_into_ANR [OF contf fim _ \closed S\] ANR_rel_frontier_convex T by blast + have "compact S" + by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact) + then obtain d where "d > 0" and d: "\x y. \x \ S; y \ - V\ \ d \ dist x y" + using separate_compact_closed [of S "-V"] \open V\ \S \ V\ by force + obtain \ where "finite \" "\\ = \\" + and diaG: "\X. X \ \ \ diameter X < d" + and polyG: "\X. X \ \ \ polytope X" + and affG: "\X. X \ \ \ aff_dim X \ aff_dim T - 1" + and faceG: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X \ X \ Y face_of Y" + proof (rule cell_complex_subdivision_exists [OF \d>0\ \finite \\ poly _ face]) + show "\X. X \ \ \ aff_dim X \ aff_dim T - 1" + by (simp add: aff) + qed auto + obtain h where conth: "continuous_on (\\) h" and him: "h ` \\ \ rel_frontier T" and hg: "\x. x \ \(\ \ Pow V) \ h x = g x" + proof (rule extend_map_lemma [of \ "\ \ Pow V" T g]) + show "continuous_on (\(\ \ Pow V)) g" + by (metis Union_Int_subset Union_Pow_eq \continuous_on V g\ continuous_on_subset le_inf_iff) + qed (use \finite \\ T polyG affG faceG gim in fastforce)+ + show ?thesis + proof + show "continuous_on (\\) h" + using \\\ = \\\ conth by auto + show "h ` \\ \ rel_frontier T" + using \\\ = \\\ him by auto + show "h x = f x" if "x \ S" for x + proof - + have "x \ \\" + using \\\ = \\\ \S \ \\\ that by auto + then obtain X where "x \ X" "X \ \" by blast + then have "diameter X < d" "bounded X" + by (auto simp: diaG \X \ \\ polyG polytope_imp_bounded) + then have "X \ V" using d [OF \x \ S\] diameter_bounded_bound [OF \bounded X\ \x \ X\] + by fastforce + have "h x = g x" + apply (rule hg) + using \X \ \\ \X \ V\ \x \ X\ by blast + also have "... = f x" + by (simp add: gf that) + finally show "h x = f x" . + qed + qed +qed + + +theorem extend_map_cell_complex_to_sphere_cofinite: + assumes "finite \" and S: "S \ \\" "closed S" and T: "convex T" "bounded T" + and poly: "\X. X \ \ \ polytope X" + and aff: "\X. X \ \ \ aff_dim X \ aff_dim T" + and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X \ (X \ Y) face_of Y" + and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier T" + obtains C g where "finite C" "disjnt C S" "continuous_on (\\ - C) g" + "g ` (\\ - C) \ rel_frontier T" "\x. x \ S \ g x = f x" +proof - + obtain V g where "S \ V" "open V" "continuous_on V g" and gim: "g ` V \ rel_frontier T" and gf: "\x. x \ S \ g x = f x" + using neighbourhood_extension_into_ANR [OF contf fim _ \closed S\] ANR_rel_frontier_convex T by blast + have "compact S" + by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact) + then obtain d where "d > 0" and d: "\x y. \x \ S; y \ - V\ \ d \ dist x y" + using separate_compact_closed [of S "-V"] \open V\ \S \ V\ by force + obtain \ where "finite \" "\\ = \\" + and diaG: "\X. X \ \ \ diameter X < d" + and polyG: "\X. X \ \ \ polytope X" + and affG: "\X. X \ \ \ aff_dim X \ aff_dim T" + and faceG: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X \ X \ Y face_of Y" + by (rule cell_complex_subdivision_exists [OF \d>0\ \finite \\ poly aff face]) auto + obtain C h where "finite C" and dis: "disjnt C (\(\ \ Pow V))" + and card: "card C \ card \" and conth: "continuous_on (\\ - C) h" + and him: "h ` (\\ - C) \ rel_frontier T" + and hg: "\x. x \ \(\ \ Pow V) \ h x = g x" + proof (rule extend_map_lemma_cofinite [of \ "\ \ Pow V" T g]) + show "continuous_on (\(\ \ Pow V)) g" + by (metis Union_Int_subset Union_Pow_eq \continuous_on V g\ continuous_on_subset le_inf_iff) + show "g ` \(\ \ Pow V) \ rel_frontier T" + using gim by force + qed (auto intro: \finite \\ T polyG affG dest: faceG) + have Ssub: "S \ \(\ \ Pow V)" + proof + fix x + assume "x \ S" + then have "x \ \\" + using \\\ = \\\ \S \ \\\ by auto + then obtain X where "x \ X" "X \ \" by blast + then have "diameter X < d" "bounded X" + by (auto simp: diaG \X \ \\ polyG polytope_imp_bounded) + then have "X \ V" using d [OF \x \ S\] diameter_bounded_bound [OF \bounded X\ \x \ X\] + by fastforce + then show "x \ \(\ \ Pow V)" + using \X \ \\ \x \ X\ by blast + qed + show ?thesis + proof + show "continuous_on (\\-C) h" + using \\\ = \\\ conth by auto + show "h ` (\\ - C) \ rel_frontier T" + using \\\ = \\\ him by auto + show "h x = f x" if "x \ S" for x + proof - + have "h x = g x" + apply (rule hg) + using Ssub that by blast + also have "... = f x" + by (simp add: gf that) + finally show "h x = f x" . + qed + show "disjnt C S" + using dis Ssub by (meson disjnt_iff subset_eq) + qed (intro \finite C\) +qed + + + +subsection\ Special cases and corollaries involving spheres.\ + +lemma disjnt_Diff1: "X \ Y' \ disjnt (X - Y) (X' - Y')" + by (auto simp: disjnt_def) + +proposition extend_map_affine_to_sphere_cofinite_simple: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "compact S" "convex U" "bounded U" + and aff: "aff_dim T \ aff_dim U" + and "S \ T" and contf: "continuous_on S f" + and fim: "f ` S \ rel_frontier U" + obtains K g where "finite K" "K \ T" "disjnt K S" "continuous_on (T - K) g" + "g ` (T - K) \ rel_frontier U" + "\x. x \ S \ g x = f x" +proof - + have "\K g. finite K \ disjnt K S \ continuous_on (T - K) g \ + g ` (T - K) \ rel_frontier U \ (\x \ S. g x = f x)" + if "affine T" "S \ T" and aff: "aff_dim T \ aff_dim U" for T + proof (cases "S = {}") + case True + show ?thesis + proof (cases "rel_frontier U = {}") + case True + with \bounded U\ have "aff_dim U \ 0" + using affine_bounded_eq_lowdim rel_frontier_eq_empty by auto + with aff have "aff_dim T \ 0" by auto + then obtain a where "T \ {a}" + using \affine T\ affine_bounded_eq_lowdim affine_bounded_eq_trivial by auto + then show ?thesis + using \S = {}\ fim + by (metis Diff_cancel contf disjnt_empty2 finite.emptyI finite_insert finite_subset) + next + case False + then obtain a where "a \ rel_frontier U" + by auto + then show ?thesis + using continuous_on_const [of _ a] \S = {}\ by force + qed + next + case False + have "bounded S" + by (simp add: \compact S\ compact_imp_bounded) + then obtain b where b: "S \ cbox (-b) b" + using bounded_subset_cbox_symmetric by blast + define bbox where "bbox \ cbox (-(b+One)) (b+One)" + have "cbox (-b) b \ bbox" + by (auto simp: bbox_def algebra_simps intro!: subset_box_imp) + with b \S \ T\ have "S \ bbox \ T" + by auto + then have Ssub: "S \ \{bbox \ T}" + by auto + then have "aff_dim (bbox \ T) \ aff_dim U" + by (metis aff aff_dim_subset inf_commute inf_le1 order_trans) + obtain K g where K: "finite K" "disjnt K S" + and contg: "continuous_on (\{bbox \ T} - K) g" + and gim: "g ` (\{bbox \ T} - K) \ rel_frontier U" + and gf: "\x. x \ S \ g x = f x" + proof (rule extend_map_cell_complex_to_sphere_cofinite + [OF _ Ssub _ \convex U\ \bounded U\ _ _ _ contf fim]) + show "closed S" + using \compact S\ compact_eq_bounded_closed by auto + show poly: "\X. X \ {bbox \ T} \ polytope X" + by (simp add: polytope_Int_polyhedron bbox_def polytope_interval affine_imp_polyhedron \affine T\) + show "\X Y. \X \ {bbox \ T}; Y \ {bbox \ T}\ \ X \ Y face_of X \ X \ Y face_of Y" + by (simp add:poly face_of_refl polytope_imp_convex) + show "\X. X \ {bbox \ T} \ aff_dim X \ aff_dim U" + by (simp add: \aff_dim (bbox \ T) \ aff_dim U\) + qed auto + define fro where "fro \ \d. frontier(cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))" + obtain d where d12: "1/2 \ d" "d \ 1" and dd: "disjnt K (fro d)" + proof (rule disjoint_family_elem_disjnt [OF _ \finite K\]) + show "infinite {1/2..1::real}" + by (simp add: infinite_Icc) + have dis1: "disjnt (fro x) (fro y)" if "x b + d *\<^sub>R One" + have cbsub: "cbox (-b) b \ box (-c) c" "cbox (-b) b \ cbox (-c) c" "cbox (-c) c \ bbox" + using d12 by (auto simp: algebra_simps subset_box_imp c_def bbox_def) + have clo_cbT: "closed (cbox (- c) c \ T)" + by (simp add: affine_closed closed_Int closed_cbox \affine T\) + have cpT_ne: "cbox (- c) c \ T \ {}" + using \S \ {}\ b cbsub(2) \S \ T\ by fastforce + have "closest_point (cbox (- c) c \ T) x \ K" if "x \ T" "x \ K" for x + proof (cases "x \ cbox (-c) c") + case True with that show ?thesis + by (simp add: closest_point_self) + next + case False + have int_ne: "interior (cbox (-c) c) \ T \ {}" + using \S \ {}\ \S \ T\ b \cbox (- b) b \ box (- c) c\ by force + have "convex T" + by (meson \affine T\ affine_imp_convex) + then have "x \ affine hull (cbox (- c) c \ T)" + by (metis Int_commute Int_iff \S \ {}\ \S \ T\ cbsub(1) \x \ T\ affine_hull_convex_Int_nonempty_interior all_not_in_conv b hull_inc inf.orderE interior_cbox) + then have "x \ affine hull (cbox (- c) c \ T) - rel_interior (cbox (- c) c \ T)" + by (meson DiffI False Int_iff rel_interior_subset subsetCE) + then have "closest_point (cbox (- c) c \ T) x \ rel_frontier (cbox (- c) c \ T)" + by (rule closest_point_in_rel_frontier [OF clo_cbT cpT_ne]) + moreover have "(rel_frontier (cbox (- c) c \ T)) \ fro d" + apply (subst convex_affine_rel_frontier_Int [OF _ \affine T\ int_ne]) + apply (auto simp: fro_def c_def) + done + ultimately show ?thesis + using dd by (force simp: disjnt_def) + qed + then have cpt_subset: "closest_point (cbox (- c) c \ T) ` (T - K) \ \{bbox \ T} - K" + using closest_point_in_set [OF clo_cbT cpT_ne] cbsub(3) by force + show ?thesis + proof (intro conjI ballI exI) + have "continuous_on (T - K) (closest_point (cbox (- c) c \ T))" + apply (rule continuous_on_closest_point) + using \S \ {}\ cbsub(2) b that + by (auto simp: affine_imp_convex convex_Int affine_closed closed_Int closed_cbox \affine T\) + then show "continuous_on (T - K) (g \ closest_point (cbox (- c) c \ T))" + by (metis continuous_on_compose continuous_on_subset [OF contg cpt_subset]) + have "(g \ closest_point (cbox (- c) c \ T)) ` (T - K) \ g ` (\{bbox \ T} - K)" + by (metis image_comp image_mono cpt_subset) + also have "... \ rel_frontier U" + by (rule gim) + finally show "(g \ closest_point (cbox (- c) c \ T)) ` (T - K) \ rel_frontier U" . + show "(g \ closest_point (cbox (- c) c \ T)) x = f x" if "x \ S" for x + proof - + have "(g \ closest_point (cbox (- c) c \ T)) x = g x" + unfolding o_def + by (metis IntI \S \ T\ b cbsub(2) closest_point_self subset_eq that) + also have "... = f x" + by (simp add: that gf) + finally show ?thesis . + qed + qed (auto simp: K) + qed + then obtain K g where "finite K" "disjnt K S" + and contg: "continuous_on (affine hull T - K) g" + and gim: "g ` (affine hull T - K) \ rel_frontier U" + and gf: "\x. x \ S \ g x = f x" + by (metis aff affine_affine_hull aff_dim_affine_hull + order_trans [OF \S \ T\ hull_subset [of T affine]]) + then obtain K g where "finite K" "disjnt K S" + and contg: "continuous_on (T - K) g" + and gim: "g ` (T - K) \ rel_frontier U" + and gf: "\x. x \ S \ g x = f x" + by (rule_tac K=K and g=g in that) (auto simp: hull_inc elim: continuous_on_subset) + then show ?thesis + by (rule_tac K="K \ T" and g=g in that) (auto simp: disjnt_iff Diff_Int contg) +qed + +subsection\Extending maps to spheres\ + +(*Up to extend_map_affine_to_sphere_cofinite_gen*) + +lemma closedin_closed_subset: + "\closedin (subtopology euclidean U) V; T \ U; S = V \ T\ + \ closedin (subtopology euclidean T) S" + by (metis (no_types, lifting) Int_assoc Int_commute closedin_closed inf.orderE) + +lemma extend_map_affine_to_sphere1: + fixes f :: "'a::euclidean_space \ 'b::topological_space" + assumes "finite K" "affine U" and contf: "continuous_on (U - K) f" + and fim: "f ` (U - K) \ T" + and comps: "\C. \C \ components(U - S); C \ K \ {}\ \ C \ L \ {}" + and clo: "closedin (subtopology euclidean U) S" and K: "disjnt K S" "K \ U" + obtains g where "continuous_on (U - L) g" "g ` (U - L) \ T" "\x. x \ S \ g x = f x" +proof (cases "K = {}") + case True + then show ?thesis + by (metis Diff_empty Diff_subset contf fim continuous_on_subset image_subsetI rev_image_eqI subset_iff that) +next + case False + have "S \ U" + using clo closedin_limpt by blast + then have "(U - S) \ K \ {}" + by (metis Diff_triv False Int_Diff K disjnt_def inf.absorb_iff2 inf_commute) + then have "\(components (U - S)) \ K \ {}" + using Union_components by simp + then obtain C0 where C0: "C0 \ components (U - S)" "C0 \ K \ {}" + by blast + have "convex U" + by (simp add: affine_imp_convex \affine U\) + then have "locally connected U" + by (rule convex_imp_locally_connected) + have "\a g. a \ C \ a \ L \ continuous_on (S \ (C - {a})) g \ + g ` (S \ (C - {a})) \ T \ (\x \ S. g x = f x)" + if C: "C \ components (U - S)" and CK: "C \ K \ {}" for C + proof - + have "C \ U-S" "C \ L \ {}" + by (simp_all add: in_components_subset comps that) + then obtain a where a: "a \ C" "a \ L" by auto + have opeUC: "openin (subtopology euclidean U) C" + proof (rule openin_trans) + show "openin (subtopology euclidean (U-S)) C" + by (simp add: \locally connected U\ clo locally_diff_closed openin_components_locally_connected [OF _ C]) + show "openin (subtopology euclidean U) (U - S)" + by (simp add: clo openin_diff) + qed + then obtain d where "C \ U" "0 < d" and d: "cball a d \ U \ C" + using openin_contains_cball by (metis \a \ C\) + then have "ball a d \ U \ C" + by auto + obtain h k where homhk: "homeomorphism (S \ C) (S \ C) h k" + and subC: "{x. (~ (h x = x \ k x = x))} \ C" + and bou: "bounded {x. (~ (h x = x \ k x = x))}" + and hin: "\x. x \ C \ K \ h x \ ball a d \ U" + proof (rule homeomorphism_grouping_points_exists_gen [of C "ball a d \ U" "C \ K" "S \ C"]) + show "openin (subtopology euclidean C) (ball a d \ U)" + by (metis Topology_Euclidean_Space.open_ball \C \ U\ \ball a d \ U \ C\ inf.absorb_iff2 inf.orderE inf_assoc open_openin openin_subtopology) + show "openin (subtopology euclidean (affine hull C)) C" + by (metis \a \ C\ \openin (subtopology euclidean U) C\ affine_hull_eq affine_hull_openin all_not_in_conv \affine U\) + show "ball a d \ U \ {}" + using \0 < d\ \C \ U\ \a \ C\ by force + show "finite (C \ K)" + by (simp add: \finite K\) + show "S \ C \ affine hull C" + by (metis \C \ U\ \S \ U\ \a \ C\ opeUC affine_hull_eq affine_hull_openin all_not_in_conv assms(2) sup.bounded_iff) + show "connected C" + by (metis C in_components_connected) + qed auto + have a_BU: "a \ ball a d \ U" + using \0 < d\ \C \ U\ \a \ C\ by auto + have "rel_frontier (cball a d \ U) retract_of (affine hull (cball a d \ U) - {a})" + apply (rule rel_frontier_retract_of_punctured_affine_hull) + apply (auto simp: \convex U\ convex_Int) + by (metis \affine U\ convex_cball empty_iff interior_cball a_BU rel_interior_convex_Int_affine) + moreover have "rel_frontier (cball a d \ U) = frontier (cball a d) \ U" + apply (rule convex_affine_rel_frontier_Int) + using a_BU by (force simp: \affine U\)+ + moreover have "affine hull (cball a d \ U) = U" + by (metis \convex U\ a_BU affine_hull_convex_Int_nonempty_interior affine_hull_eq \affine U\ equals0D inf.commute interior_cball) + ultimately have "frontier (cball a d) \ U retract_of (U - {a})" + by metis + then obtain r where contr: "continuous_on (U - {a}) r" + and rim: "r ` (U - {a}) \ sphere a d" "r ` (U - {a}) \ U" + and req: "\x. x \ sphere a d \ U \ r x = x" + using \affine U\ by (auto simp: retract_of_def retraction_def hull_same) + define j where "j \ \x. if x \ ball a d then r x else x" + have kj: "\x. x \ S \ k (j x) = x" + using \C \ U - S\ \S \ U\ \ball a d \ U \ C\ j_def subC by auto + have Uaeq: "U - {a} = (cball a d - {a}) \ U \ (U - ball a d)" + using \0 < d\ by auto + have jim: "j ` (S \ (C - {a})) \ (S \ C) - ball a d" + proof clarify + fix y assume "y \ S \ (C - {a})" + then have "y \ U - {a}" + using \C \ U - S\ \S \ U\ \a \ C\ by auto + then have "r y \ sphere a d" + using rim by auto + then show "j y \ S \ C - ball a d" + apply (simp add: j_def) + using \r y \ sphere a d\ \y \ U - {a}\ \y \ S \ (C - {a})\ d rim by fastforce + qed + have contj: "continuous_on (U - {a}) j" + unfolding j_def Uaeq + proof (intro continuous_on_cases_local continuous_on_id, simp_all add: req closedin_closed Uaeq [symmetric]) + show "\T. closed T \ (cball a d - {a}) \ U = (U - {a}) \ T" + apply (rule_tac x="(cball a d) \ U" in exI) + using affine_closed \affine U\ by blast + show "\T. closed T \ U - ball a d = (U - {a}) \ T" + apply (rule_tac x="U - ball a d" in exI) + using \0 < d\ by (force simp: affine_closed \affine U\ closed_Diff) + show "continuous_on ((cball a d - {a}) \ U) r" + by (force intro: continuous_on_subset [OF contr]) + qed + have fT: "x \ U - K \ f x \ T" for x + using fim by blast + show ?thesis + proof (intro conjI exI) + show "continuous_on (S \ (C - {a})) (f \ k \ j)" + proof (intro continuous_on_compose) + show "continuous_on (S \ (C - {a})) j" + apply (rule continuous_on_subset [OF contj]) + using \C \ U - S\ \S \ U\ \a \ C\ by force + show "continuous_on (j ` (S \ (C - {a}))) k" + apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]]) + using jim \C \ U - S\ \S \ U\ \ball a d \ U \ C\ j_def by fastforce + show "continuous_on (k ` j ` (S \ (C - {a}))) f" + proof (clarify intro!: continuous_on_subset [OF contf]) + fix y assume "y \ S \ (C - {a})" + have ky: "k y \ S \ C" + using homeomorphism_image2 [OF homhk] \y \ S \ (C - {a})\ by blast + have jy: "j y \ S \ C - ball a d" + using Un_iff \y \ S \ (C - {a})\ jim by auto + show "k (j y) \ U - K" + apply safe + using \C \ U\ \S \ U\ homeomorphism_image2 [OF homhk] jy apply blast + by (metis DiffD1 DiffD2 Int_iff Un_iff \disjnt K S\ disjnt_def empty_iff hin homeomorphism_apply2 homeomorphism_image2 homhk imageI jy) + qed + qed + have ST: "\x. x \ S \ (f \ k \ j) x \ T" + apply (simp add: kj) + apply (metis DiffI \S \ U\ \disjnt K S\ subsetD disjnt_iff fim image_subset_iff) + done + moreover have "(f \ k \ j) x \ T" if "x \ C" "x \ a" "x \ S" for x + proof - + have rx: "r x \ sphere a d" + using \C \ U\ rim that by fastforce + have jj: "j x \ S \ C - ball a d" + using jim that by blast + have "k (j x) = j x \ k (j x) \ C \ j x \ C" + by (metis Diff_iff Int_iff Un_iff \S \ U\ subsetD d j_def jj rx sphere_cball that(1)) + then have "k (j x) \ C" + using homeomorphism_apply2 [OF homhk, of "j x"] \C \ U\ \S \ U\ a rx + by (metis (mono_tags, lifting) Diff_iff subsetD jj mem_Collect_eq subC) + with jj \C \ U\ show ?thesis + apply safe + using ST j_def apply fastforce + apply (auto simp: not_less intro!: fT) + by (metis DiffD1 DiffD2 Int_iff hin homeomorphism_apply2 [OF homhk] jj) + qed + ultimately show "(f \ k \ j) ` (S \ (C - {a})) \ T" + by force + show "\x\S. (f \ k \ j) x = f x" using kj by simp + qed (auto simp: a) + qed + then obtain a h where + ah: "\C. \C \ components (U - S); C \ K \ {}\ + \ a C \ C \ a C \ L \ continuous_on (S \ (C - {a C})) (h C) \ + h C ` (S \ (C - {a C})) \ T \ (\x \ S. h C x = f x)" + using that by metis + define F where "F \ {C \ components (U - S). C \ K \ {}}" + define G where "G \ {C \ components (U - S). C \ K = {}}" + define UF where "UF \ (\C\F. C - {a C})" + have "C0 \ F" + by (auto simp: F_def C0) + have "finite F" + proof (subst finite_image_iff [of "\C. C \ K" F, symmetric]) + show "inj_on (\C. C \ K) F" + unfolding F_def inj_on_def + using components_nonoverlap by blast + show "finite ((\C. C \ K) ` F)" + unfolding F_def + by (rule finite_subset [of _ "Pow K"]) (auto simp: \finite K\) + qed + obtain g where contg: "continuous_on (S \ UF) g" + and gh: "\x i. \i \ F; x \ (S \ UF) \ (S \ (i - {a i}))\ + \ g x = h i x" + proof (rule pasting_lemma_exists_closed [OF \finite F\, of "S \ UF" "\C. S \ (C - {a C})" h]) + show "S \ UF \ (\C\F. S \ (C - {a C}))" + using \C0 \ F\ by (force simp: UF_def) + show "closedin (subtopology euclidean (S \ UF)) (S \ (C - {a C}))" + if "C \ F" for C + proof (rule closedin_closed_subset [of U "S \ C"]) + show "closedin (subtopology euclidean U) (S \ C)" + apply (rule closedin_Un_complement_component [OF \locally connected U\ clo]) + using F_def that by blast + next + have "x = a C'" if "C' \ F" "x \ C'" "x \ U" for x C' + proof - + have "\A. x \ \A \ C' \ A" + using \x \ C'\ by blast + with that show "x = a C'" + by (metis (lifting) DiffD1 F_def Union_components mem_Collect_eq) + qed + then show "S \ UF \ U" + using \S \ U\ by (force simp: UF_def) + next + show "S \ (C - {a C}) = (S \ C) \ (S \ UF)" + using F_def UF_def components_nonoverlap that by auto + qed + next + show "continuous_on (S \ (C' - {a C'})) (h C')" if "C' \ F" for C' + using ah F_def that by blast + show "\i j x. \i \ F; j \ F; + x \ (S \ UF) \ (S \ (i - {a i})) \ (S \ (j - {a j}))\ + \ h i x = h j x" + using components_eq by (fastforce simp: components_eq F_def ah) + qed blast + have SU': "S \ \G \ (S \ UF) \ U" + using \S \ U\ in_components_subset by (auto simp: F_def G_def UF_def) + have clo1: "closedin (subtopology euclidean (S \ \G \ (S \ UF))) (S \ \G)" + proof (rule closedin_closed_subset [OF _ SU']) + have *: "\C. C \ F \ openin (subtopology euclidean U) C" + unfolding F_def + by clarify (metis (no_types, lifting) \locally connected U\ clo closedin_def locally_diff_closed openin_components_locally_connected openin_trans topspace_euclidean_subtopology) + show "closedin (subtopology euclidean U) (U - UF)" + unfolding UF_def + by (force intro: openin_delete *) + show "S \ \G = (U - UF) \ (S \ \G \ (S \ UF))" + using \S \ U\ apply (auto simp: F_def G_def UF_def) + apply (metis Diff_iff UnionI Union_components) + apply (metis DiffD1 UnionI Union_components) + by (metis (no_types, lifting) IntI components_nonoverlap empty_iff) + qed + have clo2: "closedin (subtopology euclidean (S \ \G \ (S \ UF))) (S \ UF)" + proof (rule closedin_closed_subset [OF _ SU']) + show "closedin (subtopology euclidean U) (\C\F. S \ C)" + apply (rule closedin_Union) + apply (simp add: \finite F\) + using F_def \locally connected U\ clo closedin_Un_complement_component by blast + show "S \ UF = (\C\F. S \ C) \ (S \ \G \ (S \ UF))" + using \S \ U\ apply (auto simp: F_def G_def UF_def) + using C0 apply blast + by (metis components_nonoverlap disjnt_def disjnt_iff) + qed + have SUG: "S \ \G \ U - K" + using \S \ U\ K apply (auto simp: G_def disjnt_iff) + by (meson Diff_iff subsetD in_components_subset) + then have contf': "continuous_on (S \ \G) f" + by (rule continuous_on_subset [OF contf]) + have contg': "continuous_on (S \ UF) g" + apply (rule continuous_on_subset [OF contg]) + using \S \ U\ by (auto simp: F_def G_def) + have "\x. \S \ U; x \ S\ \ f x = g x" + by (subst gh) (auto simp: ah C0 intro: \C0 \ F\) + then have f_eq_g: "\x. x \ S \ UF \ x \ S \ \G \ f x = g x" + using \S \ U\ apply (auto simp: F_def G_def UF_def dest: in_components_subset) + using components_eq by blast + have cont: "continuous_on (S \ \G \ (S \ UF)) (\x. if x \ S \ \G then f x else g x)" + by (blast intro: continuous_on_cases_local [OF clo1 clo2 contf' contg' f_eq_g, of "\x. x \ S \ \G"]) + show ?thesis + proof + have UF: "\F - L \ UF" + unfolding F_def UF_def using ah by blast + have "U - S - L = \(components (U - S)) - L" + by simp + also have "... = \F \ \G - L" + unfolding F_def G_def by blast + also have "... \ UF \ \G" + using UF by blast + finally have "U - L \ S \ \G \ (S \ UF)" + by blast + then show "continuous_on (U - L) (\x. if x \ S \ \G then f x else g x)" + by (rule continuous_on_subset [OF cont]) + have "((U - L) \ {x. x \ S \ (\xa\G. x \ xa)}) \ ((U - L) \ (-S \ UF))" + using \U - L \ S \ \G \ (S \ UF)\ by auto + moreover have "g ` ((U - L) \ (-S \ UF)) \ T" + proof - + have "g x \ T" if "x \ U" "x \ L" "x \ S" "C \ F" "x \ C" "x \ a C" for x C + proof (subst gh) + show "x \ (S \ UF) \ (S \ (C - {a C}))" + using that by (auto simp: UF_def) + show "h C x \ T" + using ah that by (fastforce simp add: F_def) + qed (rule that) + then show ?thesis + by (force simp: UF_def) + qed + ultimately have "g ` ((U - L) \ {x. x \ S \ (\xa\G. x \ xa)}) \ T" + using image_mono order_trans by blast + moreover have "f ` ((U - L) \ (S \ \G)) \ T" + using fim SUG by blast + ultimately show "(\x. if x \ S \ \G then f x else g x) ` (U - L) \ T" + by force + show "\x. x \ S \ (if x \ S \ \G then f x else g x) = f x" + by (simp add: F_def G_def) + qed +qed + + +lemma extend_map_affine_to_sphere2: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "compact S" "convex U" "bounded U" "affine T" "S \ T" + and affTU: "aff_dim T \ aff_dim U" + and contf: "continuous_on S f" + and fim: "f ` S \ rel_frontier U" + and ovlap: "\C. C \ components(T - S) \ C \ L \ {}" + obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" + "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" + "\x. x \ S \ g x = f x" +proof - + obtain K g where K: "finite K" "K \ T" "disjnt K S" + and contg: "continuous_on (T - K) g" + and gim: "g ` (T - K) \ rel_frontier U" + and gf: "\x. x \ S \ g x = f x" + using assms extend_map_affine_to_sphere_cofinite_simple by metis + have "(\y C. C \ components (T - S) \ x \ C \ y \ C \ y \ L)" if "x \ K" for x + proof - + have "x \ T-S" + using \K \ T\ \disjnt K S\ disjnt_def that by fastforce + then obtain C where "C \ components(T - S)" "x \ C" + by (metis UnionE Union_components) + with ovlap [of C] show ?thesis + by blast + qed + then obtain \ where \: "\x. x \ K \ \C. C \ components (T - S) \ x \ C \ \ x \ C \ \ x \ L" + by metis + obtain h where conth: "continuous_on (T - \ ` K) h" + and him: "h ` (T - \ ` K) \ rel_frontier U" + and hg: "\x. x \ S \ h x = g x" + proof (rule extend_map_affine_to_sphere1 [OF \finite K\ \affine T\ contg gim, of S "\ ` K"]) + show cloTS: "closedin (subtopology euclidean T) S" + by (simp add: \compact S\ \S \ T\ closed_subset compact_imp_closed) + show "\C. \C \ components (T - S); C \ K \ {}\ \ C \ \ ` K \ {}" + using \ components_eq by blast + qed (use K in auto) + show ?thesis + proof + show *: "\ ` K \ L" + using \ by blast + show "finite (\ ` K)" + by (simp add: K) + show "\ ` K \ T" + by clarify (meson \ Diff_iff contra_subsetD in_components_subset) + show "continuous_on (T - \ ` K) h" + by (rule conth) + show "disjnt (\ ` K) S" + using K + apply (auto simp: disjnt_def) + by (metis \ DiffD2 UnionI Union_components) + qed (simp_all add: him hg gf) +qed + + +proposition extend_map_affine_to_sphere_cofinite_gen: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes SUT: "compact S" "convex U" "bounded U" "affine T" "S \ T" + and aff: "aff_dim T \ aff_dim U" + and contf: "continuous_on S f" + and fim: "f ` S \ rel_frontier U" + and dis: "\C. \C \ components(T - S); bounded C\ \ C \ L \ {}" + obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" + "g ` (T - K) \ rel_frontier U" + "\x. x \ S \ g x = f x" +proof (cases "S = {}") + case True + show ?thesis + proof (cases "rel_frontier U = {}") + case True + with aff have "aff_dim T \ 0" + apply (simp add: rel_frontier_eq_empty) + using affine_bounded_eq_lowdim \bounded U\ order_trans by auto + with aff_dim_geq [of T] consider "aff_dim T = -1" | "aff_dim T = 0" + by linarith + then show ?thesis + proof cases + assume "aff_dim T = -1" + then have "T = {}" + by (simp add: aff_dim_empty) + then show ?thesis + by (rule_tac K="{}" in that) auto + next + assume "aff_dim T = 0" + then obtain a where "T = {a}" + using aff_dim_eq_0 by blast + then have "a \ L" + using dis [of "{a}"] \S = {}\ by (auto simp: in_components_self) + with \S = {}\ \T = {a}\ show ?thesis + by (rule_tac K="{a}" and g=f in that) auto + qed + next + case False + then obtain y where "y \ rel_frontier U" + by auto + with \S = {}\ show ?thesis + by (rule_tac K="{}" and g="\x. y" in that) (auto simp: continuous_on_const) + qed +next + case False + have "bounded S" + by (simp add: assms compact_imp_bounded) + then obtain b where b: "S \ cbox (-b) b" + using bounded_subset_cbox_symmetric by blast + define LU where "LU \ L \ (\ {C \ components (T - S). ~bounded C} - cbox (-(b+One)) (b+One))" + obtain K g where "finite K" "K \ LU" "K \ T" "disjnt K S" + and contg: "continuous_on (T - K) g" + and gim: "g ` (T - K) \ rel_frontier U" + and gf: "\x. x \ S \ g x = f x" + proof (rule extend_map_affine_to_sphere2 [OF SUT aff contf fim]) + show "C \ LU \ {}" if "C \ components (T - S)" for C + proof (cases "bounded C") + case True + with dis that show ?thesis + unfolding LU_def by fastforce + next + case False + then have "\ bounded (\{C \ components (T - S). \ bounded C})" + by (metis (no_types, lifting) Sup_upper bounded_subset mem_Collect_eq that) + then show ?thesis + apply (clarsimp simp: LU_def Int_Un_distrib Diff_Int_distrib Int_UN_distrib) + by (metis (no_types, lifting) False Sup_upper bounded_cbox bounded_subset inf.orderE mem_Collect_eq that) + qed + qed blast + have *: False if "x \ cbox (- b - m *\<^sub>R One) (b + m *\<^sub>R One)" + "x \ box (- b - n *\<^sub>R One) (b + n *\<^sub>R One)" + "0 \ m" "m < n" "n \ 1" for m n x + using that by (auto simp: mem_box algebra_simps) + have "disjoint_family_on (\d. frontier (cbox (- b - d *\<^sub>R One) (b + d *\<^sub>R One))) {1 / 2..1}" + by (auto simp: disjoint_family_on_def neq_iff frontier_def dest: *) + then obtain d where d12: "1/2 \ d" "d \ 1" + and ddis: "disjnt K (frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One)))" + using disjoint_family_elem_disjnt [of "{1/2..1::real}" K "\d. frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))"] + by (auto simp: \finite K\) + define c where "c \ b + d *\<^sub>R One" + have cbsub: "cbox (-b) b \ box (-c) c" + "cbox (-b) b \ cbox (-c) c" + "cbox (-c) c \ cbox (-(b+One)) (b+One)" + using d12 by (simp_all add: subset_box c_def inner_diff_left inner_left_distrib) + have clo_cT: "closed (cbox (- c) c \ T)" + using affine_closed \affine T\ by blast + have cT_ne: "cbox (- c) c \ T \ {}" + using \S \ {}\ \S \ T\ b cbsub by fastforce + have S_sub_cc: "S \ cbox (- c) c" + using \cbox (- b) b \ cbox (- c) c\ b by auto + show ?thesis + proof + show "finite (K \ cbox (-(b+One)) (b+One))" + using \finite K\ by blast + show "K \ cbox (- (b + One)) (b + One) \ L" + using \K \ LU\ by (auto simp: LU_def) + show "K \ cbox (- (b + One)) (b + One) \ T" + using \K \ T\ by auto + show "disjnt (K \ cbox (- (b + One)) (b + One)) S" + using \disjnt K S\ by (simp add: disjnt_def disjoint_eq_subset_Compl inf.coboundedI1) + have cloTK: "closest_point (cbox (- c) c \ T) x \ T - K" + if "x \ T" and Knot: "x \ K \ x \ cbox (- b - One) (b + One)" for x + proof (cases "x \ cbox (- c) c") + case True + with \x \ T\ show ?thesis + using cbsub(3) Knot by (force simp: closest_point_self) + next + case False + have clo_in_rf: "closest_point (cbox (- c) c \ T) x \ rel_frontier (cbox (- c) c \ T)" + proof (intro closest_point_in_rel_frontier [OF clo_cT cT_ne] DiffI notI) + have "T \ interior (cbox (- c) c) \ {}" + using \S \ {}\ \S \ T\ b cbsub(1) by fastforce + then show "x \ affine hull (cbox (- c) c \ T)" + by (simp add: Int_commute affine_hull_affine_Int_nonempty_interior \affine T\ hull_inc that(1)) + next + show "False" if "x \ rel_interior (cbox (- c) c \ T)" + proof - + have "interior (cbox (- c) c) \ T \ {}" + using \S \ {}\ \S \ T\ b cbsub(1) by fastforce + then have "affine hull (T \ cbox (- c) c) = T" + using affine_hull_convex_Int_nonempty_interior [of T "cbox (- c) c"] + by (simp add: affine_imp_convex \affine T\ inf_commute) + then show ?thesis + by (meson subsetD le_inf_iff rel_interior_subset that False) + qed + qed + have "closest_point (cbox (- c) c \ T) x \ K" + proof + assume inK: "closest_point (cbox (- c) c \ T) x \ K" + have "\x. x \ K \ x \ frontier (cbox (- (b + d *\<^sub>R One)) (b + d *\<^sub>R One))" + by (metis ddis disjnt_iff) + then show False + by (metis DiffI Int_iff \affine T\ cT_ne c_def clo_cT clo_in_rf closest_point_in_set + convex_affine_rel_frontier_Int convex_box(1) empty_iff frontier_cbox inK interior_cbox) + qed + then show ?thesis + using cT_ne clo_cT closest_point_in_set by blast + qed + show "continuous_on (T - K \ cbox (- (b + One)) (b + One)) (g \ closest_point (cbox (-c) c \ T))" + apply (intro continuous_on_compose continuous_on_closest_point continuous_on_subset [OF contg]) + apply (simp_all add: clo_cT affine_imp_convex \affine T\ convex_Int cT_ne) + using cloTK by blast + have "g (closest_point (cbox (- c) c \ T) x) \ rel_frontier U" + if "x \ T" "x \ K \ x \ cbox (- b - One) (b + One)" for x + apply (rule gim [THEN subsetD]) + using that cloTK by blast + then show "(g \ closest_point (cbox (- c) c \ T)) ` (T - K \ cbox (- (b + One)) (b + One)) + \ rel_frontier U" + by force + show "\x. x \ S \ (g \ closest_point (cbox (- c) c \ T)) x = f x" + by simp (metis (mono_tags, lifting) IntI \S \ T\ cT_ne clo_cT closest_point_refl gf subsetD S_sub_cc) + qed +qed + + +corollary extend_map_affine_to_sphere_cofinite: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes SUT: "compact S" "affine T" "S \ T" + and aff: "aff_dim T \ DIM('b)" and "0 \ r" + and contf: "continuous_on S f" + and fim: "f ` S \ sphere a r" + and dis: "\C. \C \ components(T - S); bounded C\ \ C \ L \ {}" + obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" + "g ` (T - K) \ sphere a r" "\x. x \ S \ g x = f x" +proof (cases "r = 0") + case True + with fim show ?thesis + by (rule_tac K="{}" and g = "\x. a" in that) (auto simp: continuous_on_const) +next + case False + with assms have "0 < r" by auto + then have "aff_dim T \ aff_dim (cball a r)" + by (simp add: aff aff_dim_cball) + then show ?thesis + apply (rule extend_map_affine_to_sphere_cofinite_gen + [OF \compact S\ convex_cball bounded_cball \affine T\ \S \ T\ _ contf]) + using fim apply (auto simp: assms False that dest: dis) + done +qed + +corollary extend_map_UNIV_to_sphere_cofinite: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes aff: "DIM('a) \ DIM('b)" and "0 \ r" + and SUT: "compact S" + and contf: "continuous_on S f" + and fim: "f ` S \ sphere a r" + and dis: "\C. \C \ components(- S); bounded C\ \ C \ L \ {}" + obtains K g where "finite K" "K \ L" "disjnt K S" "continuous_on (- K) g" + "g ` (- K) \ sphere a r" "\x. x \ S \ g x = f x" +apply (rule extend_map_affine_to_sphere_cofinite + [OF \compact S\ affine_UNIV subset_UNIV _ \0 \ r\ contf fim dis]) + apply (auto simp: assms that Compl_eq_Diff_UNIV [symmetric]) +done + +corollary extend_map_UNIV_to_sphere_no_bounded_component: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes aff: "DIM('a) \ DIM('b)" and "0 \ r" + and SUT: "compact S" + and contf: "continuous_on S f" + and fim: "f ` S \ sphere a r" + and dis: "\C. C \ components(- S) \ \ bounded C" + obtains g where "continuous_on UNIV g" "g ` UNIV \ sphere a r" "\x. x \ S \ g x = f x" +apply (rule extend_map_UNIV_to_sphere_cofinite [OF aff \0 \ r\ \compact S\ contf fim, of "{}"]) + apply (auto simp: that dest: dis) +done + +theorem Borsuk_separation_theorem_gen: + fixes S :: "'a::euclidean_space set" + assumes "compact S" + shows "(\c \ components(- S). ~bounded c) \ + (\f. continuous_on S f \ f ` S \ sphere (0::'a) 1 + \ (\c. homotopic_with (\x. True) S (sphere 0 1) f (\x. c)))" + (is "?lhs = ?rhs") +proof + assume L [rule_format]: ?lhs + show ?rhs + proof clarify + fix f :: "'a \ 'a" + assume contf: "continuous_on S f" and fim: "f ` S \ sphere 0 1" + obtain g where contg: "continuous_on UNIV g" and gim: "range g \ sphere 0 1" + and gf: "\x. x \ S \ g x = f x" + by (rule extend_map_UNIV_to_sphere_no_bounded_component [OF _ _ \compact S\ contf fim L]) auto + then show "\c. homotopic_with (\x. True) S (sphere 0 1) f (\x. c)" + using nullhomotopic_from_contractible [OF contg gim] + by (metis assms compact_imp_closed contf empty_iff fim homotopic_with_equal nullhomotopic_into_sphere_extension) + qed +next + assume R [rule_format]: ?rhs + show ?lhs + unfolding components_def + proof clarify + fix a + assume "a \ S" and a: "bounded (connected_component_set (- S) a)" + have cont: "continuous_on S (\x. inverse(norm(x - a)) *\<^sub>R (x - a))" + apply (intro continuous_intros) + using \a \ S\ by auto + have im: "(\x. inverse(norm(x - a)) *\<^sub>R (x - a)) ` S \ sphere 0 1" + by clarsimp (metis \a \ S\ eq_iff_diff_eq_0 left_inverse norm_eq_zero) + show False + using R cont im Borsuk_map_essential_bounded_component [OF \compact S\ \a \ S\] a by blast + qed +qed + + +corollary Borsuk_separation_theorem: + fixes S :: "'a::euclidean_space set" + assumes "compact S" and 2: "2 \ DIM('a)" + shows "connected(- S) \ + (\f. continuous_on S f \ f ` S \ sphere (0::'a) 1 + \ (\c. homotopic_with (\x. True) S (sphere 0 1) f (\x. c)))" + (is "?lhs = ?rhs") +proof + assume L: ?lhs + show ?rhs + proof (cases "S = {}") + case True + then show ?thesis by auto + next + case False + then have "(\c\components (- S). \ bounded c)" + by (metis L assms(1) bounded_empty cobounded_imp_unbounded compact_imp_bounded in_components_maximal order_refl) + then show ?thesis + by (simp add: Borsuk_separation_theorem_gen [OF \compact S\]) + qed +next + assume R: ?rhs + then show ?lhs + apply (simp add: Borsuk_separation_theorem_gen [OF \compact S\, symmetric]) + apply (auto simp: components_def connected_iff_eq_connected_component_set) + using connected_component_in apply fastforce + using cobounded_unique_unbounded_component [OF _ 2, of "-S"] \compact S\ compact_eq_bounded_closed by fastforce +qed + + +lemma homotopy_eqv_separation: + fixes S :: "'a::euclidean_space set" and T :: "'a set" + assumes "S homotopy_eqv T" and "compact S" and "compact T" + shows "connected(- S) \ connected(- T)" +proof - + consider "DIM('a) = 1" | "2 \ DIM('a)" + by (metis DIM_ge_Suc0 One_nat_def Suc_1 dual_order.antisym not_less_eq_eq) + then show ?thesis + proof cases + case 1 + then show ?thesis + using bounded_connected_Compl_1 compact_imp_bounded homotopy_eqv_empty1 homotopy_eqv_empty2 assms by metis + next + case 2 + with assms show ?thesis + by (simp add: Borsuk_separation_theorem homotopy_eqv_cohomotopic_triviality_null) + qed +qed + +lemma Jordan_Brouwer_separation: + fixes S :: "'a::euclidean_space set" and a::'a + assumes hom: "S homeomorphic sphere a r" and "0 < r" + shows "\ connected(- S)" +proof - + have "- sphere a r \ ball a r \ {}" + using \0 < r\ by (simp add: Int_absorb1 subset_eq) + moreover + have eq: "- sphere a r - ball a r = - cball a r" + by auto + have "- cball a r \ {}" + proof - + have "frontier (cball a r) \ {}" + using \0 < r\ by auto + then show ?thesis + by (metis frontier_complement frontier_empty) + qed + with eq have "- sphere a r - ball a r \ {}" + by auto + moreover + have "connected (- S) = connected (- sphere a r)" + proof (rule homotopy_eqv_separation) + show "S homotopy_eqv sphere a r" + using hom homeomorphic_imp_homotopy_eqv by blast + show "compact (sphere a r)" + by simp + then show " compact S" + using hom homeomorphic_compactness by blast + qed + ultimately show ?thesis + using connected_Int_frontier [of "- sphere a r" "ball a r"] by (auto simp: \0 < r\) +qed + + +lemma Jordan_Brouwer_frontier: + fixes S :: "'a::euclidean_space set" and a::'a + assumes S: "S homeomorphic sphere a r" and T: "T \ components(- S)" and 2: "2 \ DIM('a)" + shows "frontier T = S" +proof (cases r rule: linorder_cases) + assume "r < 0" + with S T show ?thesis by auto +next + assume "r = 0" + with S T card_eq_SucD obtain b where "S = {b}" + by (auto simp: homeomorphic_finite [of "{a}" S]) + have "components (- {b}) = { -{b}}" + using T \S = {b}\ by (auto simp: components_eq_sing_iff connected_punctured_universe 2) + with T show ?thesis + by (metis \S = {b}\ cball_trivial frontier_cball frontier_complement singletonD sphere_trivial) +next + assume "r > 0" + have "compact S" + using homeomorphic_compactness compact_sphere S by blast + show ?thesis + proof (rule frontier_minimal_separating_closed) + show "closed S" + using \compact S\ compact_eq_bounded_closed by blast + show "\ connected (- S)" + using Jordan_Brouwer_separation S \0 < r\ by blast + obtain f g where hom: "homeomorphism S (sphere a r) f g" + using S by (auto simp: homeomorphic_def) + show "connected (- T)" if "closed T" "T \ S" for T + proof - + have "f ` T \ sphere a r" + using \T \ S\ hom homeomorphism_image1 by blast + moreover have "f ` T \ sphere a r" + using \T \ S\ hom + by (metis homeomorphism_image2 homeomorphism_of_subsets order_refl psubsetE) + ultimately have "f ` T \ sphere a r" by blast + then have "connected (- f ` T)" + by (rule psubset_sphere_Compl_connected [OF _ \0 < r\ 2]) + moreover have "compact T" + using \compact S\ bounded_subset compact_eq_bounded_closed that by blast + moreover then have "compact (f ` T)" + by (meson compact_continuous_image continuous_on_subset hom homeomorphism_def psubsetE \T \ S\) + moreover have "T homotopy_eqv f ` T" + by (meson \f ` T \ sphere a r\ dual_order.strict_implies_order hom homeomorphic_def homeomorphic_imp_homotopy_eqv homeomorphism_of_subsets \T \ S\) + ultimately show ?thesis + using homotopy_eqv_separation [of T "f`T"] by blast + qed + qed (rule T) +qed + +lemma Jordan_Brouwer_nonseparation: + fixes S :: "'a::euclidean_space set" and a::'a + assumes S: "S homeomorphic sphere a r" and "T \ S" and 2: "2 \ DIM('a)" + shows "connected(- T)" +proof - + have *: "connected(C \ (S - T))" if "C \ components(- S)" for C + proof (rule connected_intermediate_closure) + show "connected C" + using in_components_connected that by auto + have "S = frontier C" + using "2" Jordan_Brouwer_frontier S that by blast + with closure_subset show "C \ (S - T) \ closure C" + by (auto simp: frontier_def) + qed auto + have "components(- S) \ {}" + by (metis S bounded_empty cobounded_imp_unbounded compact_eq_bounded_closed compact_sphere + components_eq_empty homeomorphic_compactness) + then have "- T = (\C \ components(- S). C \ (S - T))" + using Union_components [of "-S"] \T \ S\ by auto + then show ?thesis + apply (rule ssubst) + apply (rule connected_Union) + using \T \ S\ apply (auto simp: *) + done +qed + +end diff -r 5649a993666d -r 0de4736dad8b src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy Sun Oct 02 21:05:14 2016 +0200 +++ b/src/HOL/Analysis/Path_Connected.thy Mon Oct 03 13:01:01 2016 +0100 @@ -1883,6 +1883,10 @@ finally show ?thesis . qed +corollary connected_punctured_universe: + "2 \ DIM('N::euclidean_space) \ connected(- {a::'N})" + by (simp add: path_connected_punctured_universe path_connected_imp_connected) + lemma path_connected_sphere: assumes "2 \ DIM('a::euclidean_space)" shows "path_connected {x::'a. norm (x - a) = r}" @@ -2104,6 +2108,32 @@ thus ?case by (metis Diff_insert) qed +lemma psubset_sphere_Compl_connected: + fixes S :: "'a::euclidean_space set" + assumes S: "S \ sphere a r" and "0 < r" and 2: "2 \ DIM('a)" + shows "connected(- S)" +proof - + have "S \ sphere a r" + using S by blast + obtain b where "dist a b = r" and "b \ S" + using S mem_sphere by blast + have CS: "- S = {x. dist a x \ r \ (x \ S)} \ {x. r \ dist a x \ (x \ S)}" + by (auto simp: ) + have "{x. dist a x \ r \ x \ S} \ {x. r \ dist a x \ x \ S} \ {}" + using \b \ S\ \dist a b = r\ by blast + moreover have "connected {x. dist a x \ r \ x \ S}" + apply (rule connected_intermediate_closure [of "ball a r"]) + using assms by auto + moreover + have "connected {x. r \ dist a x \ x \ S}" + apply (rule connected_intermediate_closure [of "- cball a r"]) + using assms apply (auto intro: connected_complement_bounded_convex) + apply (metis ComplI interior_cball interior_closure mem_ball not_less) + done + ultimately show ?thesis + by (simp add: CS connected_Un) +qed + subsection\Relations between components and path components\ lemma open_connected_component: @@ -2505,9 +2535,9 @@ { fix y assume y1: "y \ closure (connected_component_set S x)" and y2: "y \ interior (connected_component_set S x)" - have 1: "y \ closure S" + have "y \ closure S" using y1 closure_mono connected_component_subset by blast - have "z \ interior (connected_component_set S x)" + moreover have "z \ interior (connected_component_set S x)" if "0 < e" "ball y e \ interior S" "dist y z < e" for e z proof - have "ball y e \ connected_component_set S y" @@ -2516,12 +2546,12 @@ done then show ?thesis using y1 apply (simp add: closure_approachable open_contains_ball_eq [OF open_interior]) - by (metis (no_types, hide_lams) connected_component_eq_eq connected_component_in subsetD - dist_commute mem_Collect_eq mem_ball mem_interior \0 < e\ y2) + by (metis connected_component_eq dist_commute mem_Collect_eq mem_ball mem_interior subsetD \0 < e\ y2) qed - then have 2: "y \ interior S" + then have "y \ interior S" using y2 by (force simp: open_contains_ball_eq [OF open_interior]) - note 1 2 + ultimately have "y \ frontier S" + by (auto simp: frontier_def) } then show ?thesis by (auto simp: frontier_def) qed @@ -2565,6 +2595,49 @@ by (rule order_trans [OF frontier_Union_subset_closure]) (auto simp: closure_subset_eq) +lemma frontier_of_components_subset: + fixes S :: "'a::real_normed_vector set" + shows "C \ components S \ frontier C \ frontier S" + by (metis Path_Connected.frontier_of_connected_component_subset components_iff) + +lemma frontier_of_components_closed_complement: + fixes S :: "'a::real_normed_vector set" + shows "\closed S; C \ components (- S)\ \ frontier C \ S" + using frontier_complement frontier_of_components_subset frontier_subset_eq by blast + +lemma frontier_minimal_separating_closed: + fixes S :: "'a::real_normed_vector set" + assumes "closed S" + and nconn: "~ connected(- S)" + and C: "C \ components (- S)" + and conn: "\T. \closed T; T \ S\ \ connected(- T)" + shows "frontier C = S" +proof (rule ccontr) + assume "frontier C \ S" + then have "frontier C \ S" + using frontier_of_components_closed_complement [OF \closed S\ C] by blast + then have "connected(- (frontier C))" + by (simp add: conn) + have "\ connected(- (frontier C))" + unfolding connected_def not_not + proof (intro exI conjI) + show "open C" + using C \closed S\ open_components by blast + show "open (- closure C)" + by blast + show "C \ - closure C \ - frontier C = {}" + using closure_subset by blast + show "C \ - frontier C \ {}" + using C \open C\ components_eq frontier_disjoint_eq by fastforce + show "- frontier C \ C \ - closure C" + by (simp add: \open C\ closed_Compl frontier_closures) + then show "- closure C \ - frontier C \ {}" + by (metis (no_types, lifting) C Compl_subset_Compl_iff \frontier C \ S\ compl_sup frontier_closures in_components_subset psubsetE sup.absorb_iff2 sup.boundedE sup_bot.right_neutral sup_inf_absorb) + qed + then show False + using \connected (- frontier C)\ by blast +qed + lemma connected_component_UNIV [simp]: fixes x :: "'a::real_normed_vector" shows "connected_component_set UNIV x = UNIV" @@ -6140,6 +6213,51 @@ apply (metis assms homotopy_eqv_cohomotopic_triviality_null_imp) by (metis assms homotopy_eqv_cohomotopic_triviality_null_imp homotopy_eqv_sym) +lemma homotopy_eqv_homotopic_triviality_null_imp: + fixes S :: "'a::real_normed_vector set" + and T :: "'b::real_normed_vector set" + and U :: "'c::real_normed_vector set" + assumes "S homotopy_eqv T" + and f: "continuous_on U f" "f ` U \ T" + and homSU: "\f. \continuous_on U f; f ` U \ S\ + \ \c. homotopic_with (\x. True) U S f (\x. c)" + shows "\c. homotopic_with (\x. True) U T f (\x. c)" +proof - + obtain h k where h: "continuous_on S h" "h ` S \ T" + and k: "continuous_on T k" "k ` T \ S" + and hom: "homotopic_with (\x. True) S S (k \ h) id" + "homotopic_with (\x. True) T T (h \ k) id" + using assms by (auto simp: homotopy_eqv_def) + obtain c::'a where "homotopic_with (\x. True) U S (k \ f) (\x. c)" + apply (rule exE [OF homSU [of "k \ f"]]) + apply (intro continuous_on_compose h) + using k f apply (force elim!: continuous_on_subset)+ + done + then have "homotopic_with (\x. True) U T (h \ (k \ f)) (h \ (\x. c))" + apply (rule homotopic_with_compose_continuous_left [where Y=S]) + using h by auto + moreover have "homotopic_with (\x. True) U T (id \ f) ((h \ k) \ f)" + apply (rule homotopic_with_compose_continuous_right [where X=T]) + apply (simp add: hom homotopic_with_symD) + using f apply auto + done + ultimately show ?thesis + using homotopic_with_trans by (fastforce simp add: o_def) +qed + +lemma homotopy_eqv_homotopic_triviality_null: + fixes S :: "'a::real_normed_vector set" + and T :: "'b::real_normed_vector set" + and U :: "'c::real_normed_vector set" + assumes "S homotopy_eqv T" + shows "(\f. continuous_on U f \ f ` U \ S + \ (\c. homotopic_with (\x. True) U S f (\x. c))) \ + (\f. continuous_on U f \ f ` U \ T + \ (\c. homotopic_with (\x. True) U T f (\x. c)))" +apply (rule iffI) +apply (metis assms homotopy_eqv_homotopic_triviality_null_imp) +by (metis assms homotopy_eqv_homotopic_triviality_null_imp homotopy_eqv_sym) + lemma homotopy_eqv_contractible_sets: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set"