# HG changeset patch # User paulson # Date 1464098913 -3600 # Node ID 4ae5da02d6278ddfac8ef602a9d742cab92c72fe # Parent e5cb3440af74fe00f067c95013b02e44daf4729c New theory for Homeomorphisms diff -r e5cb3440af74 -r 4ae5da02d627 src/HOL/Multivariate_Analysis/Homeomorphism.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Homeomorphism.thy Tue May 24 15:08:33 2016 +0100 @@ -0,0 +1,1106 @@ +(* Title: HOL/Multivariate_Analysis/Homeomorphism.thy + Author: LC Paulson (ported from HOL Light) +*) + +section \Homeomorphism Theorems\ + +theory Homeomorphism +imports Path_Connected +begin + +subsection \Homeomorphism of all convex compact sets with nonempty interior\ + +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] divide_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: divide_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 + +proposition + fixes S :: "'a::euclidean_space set" + assumes "compact S" and 0: "0 \ rel_interior S" + and star: "\x. x \ S \ open_segment 0 x \ rel_interior S" + shows starlike_compact_projective1_0: + "S - rel_interior S homeomorphic sphere 0 1 \ affine hull S" + (is "?SMINUS homeomorphic ?SPHER") + and starlike_compact_projective2_0: + "S homeomorphic cball 0 1 \ affine hull S" + (is "S homeomorphic ?CBALL") +proof - + have starI: "(u *\<^sub>R x) \ rel_interior S" if "x \ S" "0 \ u" "u < 1" for x u + proof (cases "x=0 \ u=0") + case True with 0 show ?thesis by force + next + case False with that show ?thesis + by (auto simp: in_segment intro: star [THEN subsetD]) + qed + have "0 \ S" using assms rel_interior_subset by auto + define proj where "proj \ \x::'a. x /\<^sub>R norm x" + have eqI: "x = y" if "proj x = proj y" "norm x = norm y" for x y + using that by (force simp: proj_def) + then have iff_eq: "\x y. (proj x = proj y \ norm x = norm y) \ x = y" + by blast + have projI: "x \ affine hull S \ proj x \ affine hull S" for x + by (metis \0 \ S\ affine_hull_span_0 hull_inc span_mul proj_def) + have nproj1 [simp]: "x \ 0 \ norm(proj x) = 1" for x + by (simp add: proj_def) + have proj0_iff [simp]: "proj x = 0 \ x = 0" for x + by (simp add: proj_def) + have cont_proj: "continuous_on (UNIV - {0}) proj" + unfolding proj_def by (rule continuous_intros | force)+ + have proj_spherI: "\x. \x \ affine hull S; x \ 0\ \ proj x \ ?SPHER" + by (simp add: projI) + have "bounded S" "closed S" + using \compact S\ compact_eq_bounded_closed by blast+ + have inj_on_proj: "inj_on proj (S - rel_interior S)" + proof + fix x y + assume x: "x \ S - rel_interior S" + and y: "y \ S - rel_interior S" and eq: "proj x = proj y" + then have xynot: "x \ 0" "y \ 0" "x \ S" "y \ S" "x \ rel_interior S" "y \ rel_interior S" + using 0 by auto + consider "norm x = norm y" | "norm x < norm y" | "norm x > norm y" by linarith + then show "x = y" + proof cases + assume "norm x = norm y" + with iff_eq eq show "x = y" by blast + next + assume *: "norm x < norm y" + have "x /\<^sub>R norm x = norm x *\<^sub>R (x /\<^sub>R norm x) /\<^sub>R norm (norm x *\<^sub>R (x /\<^sub>R norm x))" + by force + then have "proj ((norm x / norm y) *\<^sub>R y) = proj x" + by (metis (no_types) divide_inverse local.proj_def eq scaleR_scaleR) + then have [simp]: "(norm x / norm y) *\<^sub>R y = x" + by (rule eqI) (simp add: \y \ 0\) + have no: "0 \ norm x / norm y" "norm x / norm y < 1" + using * by (auto simp: divide_simps) + then show "x = y" + using starI [OF \y \ S\ no] xynot by auto + next + assume *: "norm x > norm y" + have "y /\<^sub>R norm y = norm y *\<^sub>R (y /\<^sub>R norm y) /\<^sub>R norm (norm y *\<^sub>R (y /\<^sub>R norm y))" + by force + then have "proj ((norm y / norm x) *\<^sub>R x) = proj y" + by (metis (no_types) divide_inverse local.proj_def eq scaleR_scaleR) + then have [simp]: "(norm y / norm x) *\<^sub>R x = y" + by (rule eqI) (simp add: \x \ 0\) + have no: "0 \ norm y / norm x" "norm y / norm x < 1" + using * by (auto simp: divide_simps) + then show "x = y" + using starI [OF \x \ S\ no] xynot by auto + qed + qed + have "\surf. homeomorphism (S - rel_interior S) ?SPHER proj surf" + proof (rule homeomorphism_compact) + show "compact (S - rel_interior S)" + using \compact S\ compact_rel_boundary by blast + show "continuous_on (S - rel_interior S) proj" + using 0 by (blast intro: continuous_on_subset [OF cont_proj]) + show "proj ` (S - rel_interior S) = ?SPHER" + proof + show "proj ` (S - rel_interior S) \ ?SPHER" + using 0 by (force simp: hull_inc projI intro: nproj1) + show "?SPHER \ proj ` (S - rel_interior S)" + proof (clarsimp simp: proj_def) + fix x + assume "x \ affine hull S" and nox: "norm x = 1" + then have "x \ 0" by auto + obtain d where "0 < d" and dx: "(d *\<^sub>R x) \ rel_frontier S" + and ri: "\e. \0 \ e; e < d\ \ (e *\<^sub>R x) \ rel_interior S" + using ray_to_rel_frontier [OF \bounded S\ 0] \x \ affine hull S\ \x \ 0\ by auto + show "x \ (\x. x /\<^sub>R norm x) ` (S - rel_interior S)" + apply (rule_tac x="d *\<^sub>R x" in image_eqI) + using \0 < d\ + using dx \closed S\ apply (auto simp: rel_frontier_def divide_simps nox) + done + qed + qed + qed (rule inj_on_proj) + then obtain surf where surf: "homeomorphism (S - rel_interior S) ?SPHER proj surf" + by blast + then have cont_surf: "continuous_on (proj ` (S - rel_interior S)) surf" + by (auto simp: homeomorphism_def) + have surf_nz: "\x. x \ ?SPHER \ surf x \ 0" + by (metis "0" DiffE homeomorphism_def imageI surf) + have cont_nosp: "continuous_on (?SPHER) (\x. norm x *\<^sub>R ((surf o proj) x))" + apply (rule continuous_intros)+ + apply (rule continuous_on_subset [OF cont_proj], force) + apply (rule continuous_on_subset [OF cont_surf]) + apply (force simp: homeomorphism_image1 [OF surf] dest: proj_spherI) + done + have surfpS: "\x. \norm x = 1; x \ affine hull S\ \ surf (proj x) \ S" + by (metis (full_types) DiffE \0 \ S\ homeomorphism_def image_eqI norm_zero proj_spherI real_vector.scale_zero_left scaleR_one surf) + have *: "\y. norm y = 1 \ y \ affine hull S \ x = surf (proj y)" + if "x \ S" "x \ rel_interior S" for x + proof - + have "proj x \ ?SPHER" + by (metis (full_types) "0" hull_inc proj_spherI that) + moreover have "surf (proj x) = x" + by (metis Diff_iff homeomorphism_def surf that) + ultimately show ?thesis + by (metis \\x. x \ ?SPHER \ surf x \ 0\ hull_inc inverse_1 local.proj_def norm_sgn projI scaleR_one sgn_div_norm that(1)) + qed + have surfp_notin: "\x. \norm x = 1; x \ affine hull S\ \ surf (proj x) \ rel_interior S" + by (metis (full_types) DiffE one_neq_zero homeomorphism_def image_eqI norm_zero proj_spherI surf) + have no_sp_im: "(\x. norm x *\<^sub>R surf (proj x)) ` (?SPHER) = S - rel_interior S" + by (auto simp: surfpS image_def Bex_def surfp_notin *) + have inj_spher: "inj_on (\x. norm x *\<^sub>R surf (proj x)) ?SPHER" + proof + fix x y + assume xy: "x \ ?SPHER" "y \ ?SPHER" + and eq: " norm x *\<^sub>R surf (proj x) = norm y *\<^sub>R surf (proj y)" + then have "norm x = 1" "norm y = 1" "x \ affine hull S" "y \ affine hull S" + using 0 by auto + with eq show "x = y" + by (simp add: proj_def) (metis surf xy homeomorphism_def) + qed + have co01: "compact ?SPHER" + by (simp add: closed_affine_hull compact_Int_closed) + show "?SMINUS homeomorphic ?SPHER" + apply (subst homeomorphic_sym) + apply (rule homeomorphic_compact [OF co01 cont_nosp [unfolded o_def] no_sp_im inj_spher]) + done + have proj_scaleR: "\a x. 0 < a \ proj (a *\<^sub>R x) = proj x" + by (simp add: proj_def) + have cont_sp0: "continuous_on (affine hull S - {0}) (surf o proj)" + apply (rule continuous_on_compose [OF continuous_on_subset [OF cont_proj]], force) + apply (rule continuous_on_subset [OF cont_surf]) + using homeomorphism_image1 proj_spherI surf by fastforce + obtain B where "B>0" and B: "\x. x \ S \ norm x \ B" + by (metis compact_imp_bounded \compact S\ bounded_pos_less less_eq_real_def) + have cont_nosp: "continuous (at x within ?CBALL) (\x. norm x *\<^sub>R surf (proj x))" + if "norm x \ 1" "x \ affine hull S" for x + proof (cases "x=0") + case True + show ?thesis using True + apply (simp add: continuous_within) + apply (rule lim_null_scaleR_bounded [where B=B]) + apply (simp_all add: tendsto_norm_zero eventually_at) + apply (rule_tac x=B in exI) + using B surfpS proj_def projI apply (auto simp: \B > 0\) + done + next + case False + then have "\\<^sub>F x in at x. (x \ affine hull S - {0}) = (x \ affine hull S)" + apply (simp add: eventually_at) + apply (rule_tac x="norm x" in exI) + apply (auto simp: False) + done + with cont_sp0 have *: "continuous (at x within affine hull S) (\x. surf (proj x))" + apply (simp add: continuous_on_eq_continuous_within) + apply (drule_tac x=x in bspec, force simp: False that) + apply (simp add: continuous_within Lim_transform_within_set) + done + show ?thesis + apply (rule continuous_within_subset [where s = "affine hull S", OF _ Int_lower2]) + apply (rule continuous_intros *)+ + done + qed + have cont_nosp2: "continuous_on ?CBALL (\x. norm x *\<^sub>R ((surf o proj) x))" + by (simp add: continuous_on_eq_continuous_within cont_nosp) + have "norm y *\<^sub>R surf (proj y) \ S" if "y \ cball 0 1" and yaff: "y \ affine hull S" for y + proof (cases "y=0") + case True then show ?thesis + by (simp add: \0 \ S\) + next + case False + then have "norm y *\<^sub>R surf (proj y) = norm y *\<^sub>R surf (proj (y /\<^sub>R norm y))" + by (simp add: proj_def) + have "norm y \ 1" using that by simp + have "surf (proj (y /\<^sub>R norm y)) \ S" + apply (rule surfpS) + using proj_def projI yaff + by (auto simp: False) + then have "surf (proj y) \ S" + by (simp add: False proj_def) + then show "norm y *\<^sub>R surf (proj y) \ S" + by (metis dual_order.antisym le_less_linear norm_ge_zero rel_interior_subset scaleR_one + starI subset_eq \norm y \ 1\) + qed + moreover have "x \ (\x. norm x *\<^sub>R surf (proj x)) ` (?CBALL)" if "x \ S" for x + proof (cases "x=0") + case True with that hull_inc show ?thesis by fastforce + next + case False + then have psp: "proj (surf (proj x)) = proj x" + by (metis homeomorphism_def hull_inc proj_spherI surf that) + have nxx: "norm x *\<^sub>R proj x = x" + by (simp add: False local.proj_def) + have affineI: "(1 / norm (surf (proj x))) *\<^sub>R x \ affine hull S" + by (metis \0 \ S\ affine_hull_span_0 hull_inc span_clauses(4) that) + have sproj_nz: "surf (proj x) \ 0" + by (metis False proj0_iff psp) + then have "proj x = proj (proj x)" + by (metis False nxx proj_scaleR zero_less_norm_iff) + moreover have scaleproj: "\a r. r *\<^sub>R proj a = (r / norm a) *\<^sub>R a" + by (simp add: divide_inverse local.proj_def) + ultimately have "(norm (surf (proj x)) / norm x) *\<^sub>R x \ rel_interior S" + by (metis (no_types) sproj_nz divide_self_if hull_inc norm_eq_zero nproj1 projI psp scaleR_one surfp_notin that) + then have "(norm (surf (proj x)) / norm x) \ 1" + using starI [OF that] by (meson starI [OF that] le_less_linear norm_ge_zero zero_le_divide_iff) + then have nole: "norm x \ norm (surf (proj x))" + by (simp add: le_divide_eq_1) + show ?thesis + apply (rule_tac x="inverse(norm(surf (proj x))) *\<^sub>R x" in image_eqI) + apply (metis (no_types, hide_lams) mult.commute scaleproj abs_inverse abs_norm_cancel divide_inverse norm_scaleR nxx positive_imp_inverse_positive proj_scaleR psp sproj_nz zero_less_norm_iff) + apply (auto simp: divide_simps nole affineI) + done + qed + ultimately have im_cball: "(\x. norm x *\<^sub>R surf (proj x)) ` ?CBALL = S" + by blast + have inj_cball: "inj_on (\x. norm x *\<^sub>R surf (proj x)) ?CBALL" + proof + fix x y + assume "x \ ?CBALL" "y \ ?CBALL" + and eq: "norm x *\<^sub>R surf (proj x) = norm y *\<^sub>R surf (proj y)" + then have x: "x \ affine hull S" and y: "y \ affine hull S" + using 0 by auto + show "x = y" + proof (cases "x=0 \ y=0") + case True then show "x = y" using eq proj_spherI surf_nz x y by force + next + case False + with x y have speq: "surf (proj x) = surf (proj y)" + by (metis eq homeomorphism_apply2 proj_scaleR proj_spherI surf zero_less_norm_iff) + then have "norm x = norm y" + by (metis \x \ affine hull S\ \y \ affine hull S\ eq proj_spherI real_vector.scale_cancel_right surf_nz) + moreover have "proj x = proj y" + by (metis (no_types) False speq homeomorphism_apply2 proj_spherI surf x y) + ultimately show "x = y" + using eq eqI by blast + qed + qed + have co01: "compact ?CBALL" + by (simp add: closed_affine_hull compact_Int_closed) + show "S homeomorphic ?CBALL" + apply (subst homeomorphic_sym) + apply (rule homeomorphic_compact [OF co01 cont_nosp2 [unfolded o_def] im_cball inj_cball]) + done +qed + +corollary + fixes S :: "'a::euclidean_space set" + assumes "compact S" and a: "a \ rel_interior S" + and star: "\x. x \ S \ open_segment a x \ rel_interior S" + shows starlike_compact_projective1: + "S - rel_interior S homeomorphic sphere a 1 \ affine hull S" + and starlike_compact_projective2: + "S homeomorphic cball a 1 \ affine hull S" +proof - + have 1: "compact (op+ (-a) ` S)" by (meson assms compact_translation) + have 2: "0 \ rel_interior (op+ (-a) ` S)" + by (simp add: a rel_interior_translation) + have 3: "open_segment 0 x \ rel_interior (op+ (-a) ` S)" if "x \ (op+ (-a) ` S)" for x + proof - + have "x+a \ S" using that by auto + then have "open_segment a (x+a) \ rel_interior S" by (metis star) + then show ?thesis using open_segment_translation + using rel_interior_translation by fastforce + qed + have "S - rel_interior S homeomorphic (op+ (-a) ` S) - rel_interior (op+ (-a) ` S)" + by (metis rel_interior_translation translation_diff homeomorphic_translation) + also have "... homeomorphic sphere 0 1 \ affine hull (op+ (-a) ` S)" + by (rule starlike_compact_projective1_0 [OF 1 2 3]) + also have "... = op+ (-a) ` (sphere a 1 \ affine hull S)" + by (metis affine_hull_translation left_minus sphere_translation translation_Int) + also have "... homeomorphic sphere a 1 \ affine hull S" + using homeomorphic_translation homeomorphic_sym by blast + finally show "S - rel_interior S homeomorphic sphere a 1 \ affine hull S" . + + have "S homeomorphic (op+ (-a) ` S)" + by (metis homeomorphic_translation) + also have "... homeomorphic cball 0 1 \ affine hull (op+ (-a) ` S)" + by (rule starlike_compact_projective2_0 [OF 1 2 3]) + also have "... = op+ (-a) ` (cball a 1 \ affine hull S)" + by (metis affine_hull_translation left_minus cball_translation translation_Int) + also have "... homeomorphic cball a 1 \ affine hull S" + using homeomorphic_translation homeomorphic_sym by blast + finally show "S homeomorphic cball a 1 \ affine hull S" . +qed + +corollary starlike_compact_projective_special: + assumes "compact S" + and cb01: "cball (0::'a::euclidean_space) 1 \ S" + and scale: "\x u. \x \ S; 0 \ u; u < 1\ \ u *\<^sub>R x \ S - frontier S" + shows "S homeomorphic (cball (0::'a::euclidean_space) 1)" +proof - + have "ball 0 1 \ interior S" + using cb01 interior_cball interior_mono by blast + then have 0: "0 \ rel_interior S" + by (meson centre_in_ball subsetD interior_subset_rel_interior le_numeral_extra(2) not_le) + have [simp]: "affine hull S = UNIV" + using \ball 0 1 \ interior S\ by (auto intro!: affine_hull_nonempty_interior) + have star: "open_segment 0 x \ rel_interior S" if "x \ S" for x + proof + fix p assume "p \ open_segment 0 x" + then obtain u where "x \ 0" and u: "0 \ u" "u < 1" and p: "u *\<^sub>R x = p" + by (auto simp: in_segment) + then show "p \ rel_interior S" + using scale [OF that u] closure_subset frontier_def interior_subset_rel_interior by fastforce + qed + show ?thesis + using starlike_compact_projective2_0 [OF \compact S\ 0 star] by simp +qed + +lemma homeomorphic_convex_lemma: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "convex S" "compact S" "convex T" "compact T" + and affeq: "aff_dim S = aff_dim T" + shows "(S - rel_interior S) homeomorphic (T - rel_interior T) \ + S homeomorphic T" +proof (cases "rel_interior S = {} \ rel_interior T = {}") + case True + then show ?thesis + by (metis Diff_empty affeq \convex S\ \convex T\ aff_dim_empty homeomorphic_empty rel_interior_eq_empty aff_dim_empty) +next + case False + then obtain a b where a: "a \ rel_interior S" and b: "b \ rel_interior T" by auto + have starS: "\x. x \ S \ open_segment a x \ rel_interior S" + using rel_interior_closure_convex_segment + a \convex S\ closure_subset subsetCE by blast + have starT: "\x. x \ T \ open_segment b x \ rel_interior T" + using rel_interior_closure_convex_segment + b \convex T\ closure_subset subsetCE by blast + let ?aS = "op+ (-a) ` S" and ?bT = "op+ (-b) ` T" + have 0: "0 \ affine hull ?aS" "0 \ affine hull ?bT" + by (metis a b subsetD hull_inc image_eqI left_minus rel_interior_subset)+ + have subs: "subspace (span ?aS)" "subspace (span ?bT)" + by (rule subspace_span)+ + moreover + have "dim (span (op + (- a) ` S)) = dim (span (op + (- b) ` T))" + by (metis 0 aff_dim_translation_eq aff_dim_zero affeq dim_span nat_int) + ultimately obtain f g where "linear f" "linear g" + and fim: "f ` span ?aS = span ?bT" + and gim: "g ` span ?bT = span ?aS" + and fno: "\x. x \ span ?aS \ norm(f x) = norm x" + and gno: "\x. x \ span ?bT \ norm(g x) = norm x" + and gf: "\x. x \ span ?aS \ g(f x) = x" + and fg: "\x. x \ span ?bT \ f(g x) = x" + by (rule isometries_subspaces) blast + have [simp]: "continuous_on A f" for A + using \linear f\ linear_conv_bounded_linear linear_continuous_on by blast + have [simp]: "continuous_on B g" for B + using \linear g\ linear_conv_bounded_linear linear_continuous_on by blast + have eqspanS: "affine hull ?aS = span ?aS" + by (metis a affine_hull_span_0 subsetD hull_inc image_eqI left_minus rel_interior_subset) + have eqspanT: "affine hull ?bT = span ?bT" + by (metis b affine_hull_span_0 subsetD hull_inc image_eqI left_minus rel_interior_subset) + have "S homeomorphic cball a 1 \ affine hull S" + by (rule starlike_compact_projective2 [OF \compact S\ a starS]) + also have "... homeomorphic op+ (-a) ` (cball a 1 \ affine hull S)" + by (metis homeomorphic_translation) + also have "... = cball 0 1 \ op+ (-a) ` (affine hull S)" + by (auto simp: dist_norm) + also have "... = cball 0 1 \ span ?aS" + using eqspanS affine_hull_translation by blast + also have "... homeomorphic cball 0 1 \ span ?bT" + proof (rule homeomorphicI [where f=f and g=g]) + show fim1: "f ` (cball 0 1 \ span ?aS) = cball 0 1 \ span ?bT" + apply (rule subset_antisym) + using fim fno apply (force simp:, clarify) + by (metis IntI fg gim gno image_eqI mem_cball_0) + show "g ` (cball 0 1 \ span ?bT) = cball 0 1 \ span ?aS" + apply (rule subset_antisym) + using gim gno apply (force simp:, clarify) + by (metis IntI fim1 gf image_eqI) + qed (auto simp: fg gf) + also have "... = cball 0 1 \ op+ (-b) ` (affine hull T)" + using eqspanT affine_hull_translation by blast + also have "... = op+ (-b) ` (cball b 1 \ affine hull T)" + by (auto simp: dist_norm) + also have "... homeomorphic (cball b 1 \ affine hull T)" + by (metis homeomorphic_translation homeomorphic_sym) + also have "... homeomorphic T" + by (metis starlike_compact_projective2 [OF \compact T\ b starT] homeomorphic_sym) + finally have 1: "S homeomorphic T" . + + have "S - rel_interior S homeomorphic sphere a 1 \ affine hull S" + by (rule starlike_compact_projective1 [OF \compact S\ a starS]) + also have "... homeomorphic op+ (-a) ` (sphere a 1 \ affine hull S)" + by (metis homeomorphic_translation) + also have "... = sphere 0 1 \ op+ (-a) ` (affine hull S)" + by (auto simp: dist_norm) + also have "... = sphere 0 1 \ span ?aS" + using eqspanS affine_hull_translation by blast + also have "... homeomorphic sphere 0 1 \ span ?bT" + proof (rule homeomorphicI [where f=f and g=g]) + show fim1: "f ` (sphere 0 1 \ span ?aS) = sphere 0 1 \ span ?bT" + apply (rule subset_antisym) + using fim fno apply (force simp:, clarify) + by (metis IntI fg gim gno image_eqI mem_sphere_0) + show "g ` (sphere 0 1 \ span ?bT) = sphere 0 1 \ span ?aS" + apply (rule subset_antisym) + using gim gno apply (force simp:, clarify) + by (metis IntI fim1 gf image_eqI) + qed (auto simp: fg gf) + also have "... = sphere 0 1 \ op+ (-b) ` (affine hull T)" + using eqspanT affine_hull_translation by blast + also have "... = op+ (-b) ` (sphere b 1 \ affine hull T)" + by (auto simp: dist_norm) + also have "... homeomorphic (sphere b 1 \ affine hull T)" + by (metis homeomorphic_translation homeomorphic_sym) + also have "... homeomorphic T - rel_interior T" + by (metis starlike_compact_projective1 [OF \compact T\ b starT] homeomorphic_sym) + finally have 2: "S - rel_interior S homeomorphic T - rel_interior T" . + show ?thesis + using 1 2 by blast +qed + +lemma homeomorphic_convex_compact_sets: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "convex S" "compact S" "convex T" "compact T" + and affeq: "aff_dim S = aff_dim T" + shows "S homeomorphic T" +using homeomorphic_convex_lemma [OF assms] assms +by (auto simp: rel_frontier_def) + +lemma homeomorphic_rel_frontiers_convex_bounded_sets: + fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" + assumes "convex S" "bounded S" "convex T" "bounded T" + and affeq: "aff_dim S = aff_dim T" + shows "rel_frontier S homeomorphic rel_frontier T" +using assms homeomorphic_convex_lemma [of "closure S" "closure T"] +by (simp add: rel_frontier_def convex_rel_interior_closure) + + +subsection\Homeomorphisms between punctured spheres and affine sets\ +text\Including the famous stereoscopic projection of the 3-D sphere to the complex plane\ + +text\The special case with centre 0 and radius 1\ +lemma homeomorphic_punctured_affine_sphere_affine_01: + assumes "b \ sphere 0 1" "affine T" "0 \ T" "b \ T" "affine p" + and affT: "aff_dim T = aff_dim p + 1" + shows "(sphere 0 1 \ T) - {b} homeomorphic p" +proof - + have [simp]: "norm b = 1" "b\b = 1" + using assms by (auto simp: norm_eq_1) + have [simp]: "T \ {v. b\v = 0} \ {}" + using \0 \ T\ by auto + have [simp]: "\ T \ {v. b\v = 0}" + using \norm b = 1\ \b \ T\ by auto + define f where "f \ \x. 2 *\<^sub>R b + (2 / (1 - b\x)) *\<^sub>R (x - b)" + define g where "g \ \y. b + (4 / (norm y ^ 2 + 4)) *\<^sub>R (y - 2 *\<^sub>R b)" + have [simp]: "\x. \x \ T; b\x = 0\ \ f (g x) = x" + unfolding f_def g_def by (simp add: algebra_simps divide_simps add_nonneg_eq_0_iff) + have no: "\x. \norm x = 1; b\x \ 1\ \ (norm (f x))\<^sup>2 = 4 * (1 + b\x) / (1 - b\x)" + apply (simp add: dot_square_norm [symmetric]) + apply (simp add: f_def vector_add_divide_simps divide_simps norm_eq_1) + apply (simp add: algebra_simps inner_commute) + done + have [simp]: "\u::real. 8 + u * (u * 8) = u * 16 \ u=1" + by algebra + have [simp]: "\x. \norm x = 1; b \ x \ 1\ \ g (f x) = x" + unfolding g_def no by (auto simp: f_def divide_simps) + have [simp]: "\x. \x \ T; b \ x = 0\ \ norm (g x) = 1" + unfolding g_def + apply (rule power2_eq_imp_eq) + apply (simp_all add: dot_square_norm [symmetric] divide_simps vector_add_divide_simps) + apply (simp add: algebra_simps inner_commute) + done + have [simp]: "\x. \x \ T; b \ x = 0\ \ b \ g x \ 1" + unfolding g_def + apply (simp_all add: dot_square_norm [symmetric] divide_simps vector_add_divide_simps add_nonneg_eq_0_iff) + apply (auto simp: algebra_simps) + done + have "subspace T" + by (simp add: assms subspace_affine) + have [simp]: "\x. \x \ T; b \ x = 0\ \ g x \ T" + unfolding g_def + by (blast intro: \subspace T\ \b \ T\ subspace_add subspace_mul subspace_diff) + have "f ` {x. norm x = 1 \ b\x \ 1} \ {x. b\x = 0}" + unfolding f_def using \norm b = 1\ norm_eq_1 + by (force simp: field_simps inner_add_right inner_diff_right) + moreover have "f ` T \ T" + unfolding f_def using assms + apply (auto simp: field_simps inner_add_right inner_diff_right) + by (metis add_0 diff_zero mem_affine_3_minus) + moreover have "{x. b\x = 0} \ T \ f ` ({x. norm x = 1 \ b\x \ 1} \ T)" + apply clarify + apply (rule_tac x = "g x" in image_eqI, auto) + done + ultimately have imf: "f ` ({x. norm x = 1 \ b\x \ 1} \ T) = {x. b\x = 0} \ T" + by blast + have no4: "\y. b\y = 0 \ norm ((y\y + 4) *\<^sub>R b + 4 *\<^sub>R (y - 2 *\<^sub>R b)) = y\y + 4" + apply (rule power2_eq_imp_eq) + apply (simp_all add: dot_square_norm [symmetric]) + apply (auto simp: power2_eq_square algebra_simps inner_commute) + done + have [simp]: "\x. \norm x = 1; b \ x \ 1\ \ b \ f x = 0" + by (simp add: f_def algebra_simps divide_simps) + have [simp]: "\x. \x \ T; norm x = 1; b \ x \ 1\ \ f x \ T" + unfolding f_def + by (blast intro: \subspace T\ \b \ T\ subspace_add subspace_mul subspace_diff) + have "g ` {x. b\x = 0} \ {x. norm x = 1 \ b\x \ 1}" + unfolding g_def + apply (clarsimp simp: no4 vector_add_divide_simps divide_simps add_nonneg_eq_0_iff dot_square_norm [symmetric]) + apply (auto simp: algebra_simps) + done + moreover have "g ` T \ T" + unfolding g_def + by (blast intro: \subspace T\ \b \ T\ subspace_add subspace_mul subspace_diff) + moreover have "{x. norm x = 1 \ b\x \ 1} \ T \ g ` ({x. b\x = 0} \ T)" + apply clarify + apply (rule_tac x = "f x" in image_eqI, auto) + done + ultimately have img: "g ` ({x. b\x = 0} \ T) = {x. norm x = 1 \ b\x \ 1} \ T" + by blast + have aff: "affine ({x. b\x = 0} \ T)" + by (blast intro: affine_hyperplane assms) + have contf: "continuous_on ({x. norm x = 1 \ b\x \ 1} \ T) f" + unfolding f_def by (rule continuous_intros | force)+ + have contg: "continuous_on ({x. b\x = 0} \ T) g" + unfolding g_def by (rule continuous_intros | force simp: add_nonneg_eq_0_iff)+ + have "(sphere 0 1 \ T) - {b} = {x. norm x = 1 \ (b\x \ 1)} \ T" + using \norm b = 1\ by (auto simp: norm_eq_1) (metis vector_eq \b\b = 1\) + also have "... homeomorphic {x. b\x = 0} \ T" + by (rule homeomorphicI [OF imf img contf contg]) auto + also have "... homeomorphic p" + apply (rule homeomorphic_affine_sets [OF aff \affine p\]) + apply (simp add: Int_commute aff_dim_affine_Int_hyperplane [OF \affine T\] affT) + done + finally show ?thesis . +qed + +theorem homeomorphic_punctured_affine_sphere_affine: + fixes a :: "'a :: euclidean_space" + assumes "0 < r" "b \ sphere a r" "affine T" "a \ T" "b \ T" "affine p" + and aff: "aff_dim T = aff_dim p + 1" + shows "((sphere a r \ T) - {b}) homeomorphic p" +proof - + have "a \ b" using assms by auto + then have inj: "inj (\x::'a. x /\<^sub>R norm (a - b))" + by (simp add: inj_on_def) + have "((sphere a r \ T) - {b}) homeomorphic + op+ (-a) ` ((sphere a r \ T) - {b})" + by (rule homeomorphic_translation) + also have "... homeomorphic op *\<^sub>R (inverse r) ` op + (- a) ` (sphere a r \ T - {b})" + by (metis \0 < r\ homeomorphic_scaling inverse_inverse_eq inverse_zero less_irrefl) + also have "... = sphere 0 1 \ (op *\<^sub>R (inverse r) ` op + (- a) ` T) - {(b - a) /\<^sub>R r}" + using assms by (auto simp: dist_norm norm_minus_commute divide_simps) + also have "... homeomorphic p" + apply (rule homeomorphic_punctured_affine_sphere_affine_01) + using assms + apply (auto simp: dist_norm norm_minus_commute affine_scaling affine_translation [symmetric] aff_dim_translation_eq inj) + done + finally show ?thesis . +qed + +proposition homeomorphic_punctured_sphere_affine_gen: + fixes a :: "'a :: euclidean_space" + assumes "convex S" "bounded S" and a: "a \ rel_frontier S" + and "affine T" and affS: "aff_dim S = aff_dim T + 1" + shows "rel_frontier S - {a} homeomorphic T" +proof - + have "S \ {}" using assms by auto + obtain U :: "'a set" where "affine U" and affdS: "aff_dim U = aff_dim S" + using choose_affine_subset [OF affine_UNIV aff_dim_geq] + by (meson aff_dim_affine_hull affine_affine_hull) + have "convex U" + by (simp add: affine_imp_convex \affine U\) + have "U \ {}" + by (metis \S \ {}\ \aff_dim U = aff_dim S\ aff_dim_empty) + then obtain z where "z \ U" + by auto + then have bne: "ball z 1 \ U \ {}" by force + have [simp]: "aff_dim(ball z 1 \ U) = aff_dim U" + using aff_dim_convex_Int_open [OF \convex U\ open_ball] bne + by (fastforce simp add: Int_commute) + have "rel_frontier S homeomorphic rel_frontier (ball z 1 \ U)" + apply (rule homeomorphic_rel_frontiers_convex_bounded_sets) + apply (auto simp: \affine U\ affine_imp_convex convex_Int affdS assms) + done + also have "... = sphere z 1 \ U" + using convex_affine_rel_frontier_Int [of "ball z 1" U] + by (simp add: \affine U\ bne) + finally obtain h k where him: "h ` rel_frontier S = sphere z 1 \ U" + and kim: "k ` (sphere z 1 \ U) = rel_frontier S" + and hcon: "continuous_on (rel_frontier S) h" + and kcon: "continuous_on (sphere z 1 \ U) k" + and kh: "\x. x \ rel_frontier S \ k(h(x)) = x" + and hk: "\y. y \ sphere z 1 \ U \ h(k(y)) = y" + unfolding homeomorphic_def homeomorphism_def by auto + have "rel_frontier S - {a} homeomorphic (sphere z 1 \ U) - {h a}" + proof (rule homeomorphicI [where f=h and g=k]) + show h: "h ` (rel_frontier S - {a}) = sphere z 1 \ U - {h a}" + using him a kh by auto metis + show "k ` (sphere z 1 \ U - {h a}) = rel_frontier S - {a}" + by (force simp: h [symmetric] image_comp o_def kh) + qed (auto intro: continuous_on_subset hcon kcon simp: kh hk) + also have "... homeomorphic T" + apply (rule homeomorphic_punctured_affine_sphere_affine) + using a him + by (auto simp: affS affdS \affine T\ \affine U\ \z \ U\) + finally show ?thesis . +qed + + +lemma homeomorphic_punctured_sphere_affine: + fixes a :: "'a :: euclidean_space" + assumes "0 < r" and b: "b \ sphere a r" + and "affine T" and affS: "aff_dim T + 1 = DIM('a)" + shows "(sphere a r - {b}) homeomorphic T" +using homeomorphic_punctured_sphere_affine_gen [of "cball a r" b T] + assms aff_dim_cball by force + +corollary homeomorphic_punctured_sphere_hyperplane: + fixes a :: "'a :: euclidean_space" + assumes "0 < r" and b: "b \ sphere a r" + and "c \ 0" + shows "(sphere a r - {b}) homeomorphic {x::'a. c \ x = d}" +apply (rule homeomorphic_punctured_sphere_affine) +using assms +apply (auto simp: affine_hyperplane of_nat_diff) +done + + +text\ When dealing with AR, ANR and ANR later, it's useful to know that every set + is homeomorphic to a closed subset of a convex set, and + if the set is locally compact we can take the convex set to be the universe.\ + +proposition homeomorphic_closedin_convex: + fixes S :: "'m::euclidean_space set" + assumes "aff_dim S < DIM('n)" + obtains U and T :: "'n::euclidean_space set" + where "convex U" "U \ {}" "closedin (subtopology euclidean U) T" + "S homeomorphic T" +proof (cases "S = {}") + case True then show ?thesis + by (rule_tac U=UNIV and T="{}" in that) auto +next + case False + then obtain a where "a \ S" by auto + obtain i::'n where i: "i \ Basis" "i \ 0" + using SOME_Basis Basis_zero by force + have "0 \ affine hull (op + (- a) ` S)" + by (simp add: \a \ S\ hull_inc) + then have "dim (op + (- a) ` S) = aff_dim (op + (- a) ` S)" + by (simp add: aff_dim_zero) + also have "... < DIM('n)" + by (simp add: aff_dim_translation_eq assms) + finally have dd: "dim (op + (- a) ` S) < DIM('n)" + by linarith + obtain T where "subspace T" and Tsub: "T \ {x. i \ x = 0}" + and dimT: "dim T = dim (op + (- a) ` S)" + apply (rule choose_subspace_of_subspace [of "dim (op + (- a) ` S)" "{x::'n. i \ x = 0}"]) + apply (simp add: dim_hyperplane [OF \i \ 0\]) + apply (metis DIM_positive Suc_pred dd not_le not_less_eq_eq) + apply (metis span_eq subspace_hyperplane) + done + have "subspace (span (op + (- a) ` S))" + using subspace_span by blast + then obtain h k where "linear h" "linear k" + and heq: "h ` span (op + (- a) ` S) = T" + and keq:"k ` T = span (op + (- a) ` S)" + and hinv [simp]: "\x. x \ span (op + (- a) ` S) \ k(h x) = x" + and kinv [simp]: "\x. x \ T \ h(k x) = x" + apply (rule isometries_subspaces [OF _ \subspace T\]) + apply (auto simp: dimT) + done + have hcont: "continuous_on A h" and kcont: "continuous_on B k" for A B + using \linear h\ \linear k\ linear_continuous_on linear_conv_bounded_linear by blast+ + have ihhhh[simp]: "\x. x \ S \ i \ h (x - a) = 0" + using Tsub [THEN subsetD] heq span_inc by fastforce + have "sphere 0 1 - {i} homeomorphic {x. i \ x = 0}" + apply (rule homeomorphic_punctured_sphere_affine) + using i + apply (auto simp: affine_hyperplane) + by (metis DIM_positive Suc_eq_plus1 add.left_neutral diff_add_cancel not_le not_less_eq_eq of_nat_1 of_nat_diff) + then obtain f g where fg: "homeomorphism (sphere 0 1 - {i}) {x. i \ x = 0} f g" + by (force simp: homeomorphic_def) + have "h ` op + (- a) ` S \ T" + using heq span_clauses(1) span_linear_image by blast + then have "g ` h ` op + (- a) ` S \ g ` {x. i \ x = 0}" + using Tsub by (simp add: image_mono) + also have "... \ sphere 0 1 - {i}" + by (simp add: fg [unfolded homeomorphism_def]) + finally have gh_sub_sph: "(g \ h) ` op + (- a) ` S \ sphere 0 1 - {i}" + by (metis image_comp) + then have gh_sub_cb: "(g \ h) ` op + (- a) ` S \ cball 0 1" + by (metis Diff_subset order_trans sphere_cball) + have [simp]: "\u. u \ S \ norm (g (h (u - a))) = 1" + using gh_sub_sph [THEN subsetD] by (auto simp: o_def) + have ghcont: "continuous_on (op + (- a) ` S) (\x. g (h x))" + apply (rule continuous_on_compose2 [OF homeomorphism_cont2 [OF fg] hcont], force) + done + have kfcont: "continuous_on ((g \ h \ op + (- a)) ` S) (\x. k (f x))" + apply (rule continuous_on_compose2 [OF kcont]) + using homeomorphism_cont1 [OF fg] gh_sub_sph apply (force intro: continuous_on_subset, blast) + done + have "S homeomorphic op + (- a) ` S" + by (simp add: homeomorphic_translation) + also have Shom: "\ homeomorphic (g \ h) ` op + (- a) ` S" + apply (simp add: homeomorphic_def homeomorphism_def) + apply (rule_tac x="g \ h" in exI) + apply (rule_tac x="k \ f" in exI) + apply (auto simp: ghcont kfcont span_clauses(1) homeomorphism_apply2 [OF fg] image_comp) + apply (force simp: o_def homeomorphism_apply2 [OF fg] span_clauses(1)) + done + finally have Shom: "S homeomorphic (g \ h) ` op + (- a) ` S" . + show ?thesis + apply (rule_tac U = "ball 0 1 \ image (g o h) (op + (- a) ` S)" + and T = "image (g o h) (op + (- a) ` S)" + in that) + apply (rule convex_intermediate_ball [of 0 1], force) + using gh_sub_cb apply force + apply force + apply (simp add: closedin_closed) + apply (rule_tac x="sphere 0 1" in exI) + apply (auto simp: Shom) + done +qed + +subsection\Locally compact sets in an open set\ + +text\ Locally compact sets are closed in an open set and are homeomorphic + to an absolutely closed set if we have one more dimension to play with.\ + +lemma locally_compact_open_Int_closure: + fixes S :: "'a :: metric_space set" + assumes "locally compact S" + obtains T where "open T" "S = T \ closure S" +proof - + have "\x\S. \T v u. u = S \ T \ x \ u \ u \ v \ v \ S \ open T \ compact v" + by (metis assms locally_compact openin_open) + then obtain t v where + tv: "\x. x \ S + \ v x \ S \ open (t x) \ compact (v x) \ (\u. x \ u \ u \ v x \ u = S \ t x)" + by metis + then have o: "open (UNION S t)" + by blast + have "S = \ (v ` S)" + using tv by blast + also have "... = UNION S t \ closure S" + proof + show "UNION S v \ UNION S t \ closure S" + apply safe + apply (metis Int_iff subsetD UN_iff tv) + apply (simp add: closure_def rev_subsetD tv) + done + have "t x \ closure S \ v x" if "x \ S" for x + proof - + have "t x \ closure S \ closure (t x \ S)" + by (simp add: open_Int_closure_subset that tv) + also have "... \ v x" + by (metis Int_commute closure_minimal compact_imp_closed that tv) + finally show ?thesis . + qed + then show "UNION S t \ closure S \ UNION S v" + by blast + qed + finally have e: "S = UNION S t \ closure S" . + show ?thesis + by (rule that [OF o e]) +qed + + +lemma locally_compact_closedin_open: + fixes S :: "'a :: metric_space set" + assumes "locally compact S" + obtains T where "open T" "closedin (subtopology euclidean T) S" + by (metis locally_compact_open_Int_closure [OF assms] closed_closure closedin_closed_Int) + + +lemma locally_compact_homeomorphism_projection_closed: + assumes "locally compact S" + obtains T and f :: "'a \ 'a :: euclidean_space \ 'b :: euclidean_space" + where "closed T" "homeomorphism S T f fst" +proof (cases "closed S") + case True + then show ?thesis + apply (rule_tac T = "S \ {0}" and f = "\x. (x, 0)" in that) + apply (auto simp: closed_Times homeomorphism_def continuous_intros) + done +next + case False + obtain U where "open U" and US: "U \ closure S = S" + by (metis locally_compact_open_Int_closure [OF assms]) + with False have Ucomp: "-U \ {}" + using closure_eq by auto + have [simp]: "closure (- U) = -U" + by (simp add: \open U\ closed_Compl) + define f :: "'a \ 'a \ 'b" where "f \ \x. (x, One /\<^sub>R setdist {x} (- U))" + have "continuous_on U (\x. (x, One /\<^sub>R setdist {x} (- U)))" + by (auto simp: Ucomp continuous_intros continuous_on_setdist) + then have homU: "homeomorphism U (f`U) f fst" + by (auto simp: f_def homeomorphism_def image_iff continuous_intros) + have cloS: "closedin (subtopology euclidean U) S" + by (metis US closed_closure closedin_closed_Int) + have cont: "isCont ((\x. setdist {x} (- U)) o fst) z" for z :: "'a \ 'b" + by (rule isCont_o continuous_intros continuous_at_setdist)+ + have setdist1D: "setdist {a} (- U) *\<^sub>R b = One \ setdist {a} (- U) \ 0" for a::'a and b::'b + by force + have *: "r *\<^sub>R b = One \ b = (1 / r) *\<^sub>R One" for r and b::'b + by (metis One_non_0 nonzero_divide_eq_eq real_vector.scale_eq_0_iff real_vector.scale_scale scaleR_one) + have "f ` U = {z. (setdist {fst z} (- U) *\<^sub>R snd z) \ {One}}" + apply (auto simp: f_def field_simps Ucomp) + apply (rule_tac x=a in image_eqI) + apply (auto simp: * dest: setdist1D) + done + then have clfU: "closed (f ` U)" + apply (rule ssubst) + apply (rule continuous_closed_preimage_univ) + apply (auto intro: continuous_intros cont [unfolded o_def]) + done + have "closed (f ` S)" + apply (rule closedin_closed_trans [OF _ clfU]) + apply (rule homeomorphism_imp_closed_map [OF homU cloS]) + done + then show ?thesis + apply (rule that) + apply (rule homeomorphism_of_subsets [OF homU]) + using US apply auto + done +qed + +lemma locally_compact_closed_Int_open: + fixes S :: "'a :: euclidean_space set" + shows + "locally compact S \ (\U u. closed U \ open u \ S = U \ u)" +by (metis closed_closure closed_imp_locally_compact inf_commute locally_compact_Int locally_compact_open_Int_closure open_imp_locally_compact) + + +proposition locally_compact_homeomorphic_closed: + fixes S :: "'a::euclidean_space set" + assumes "locally compact S" and dimlt: "DIM('a) < DIM('b)" + obtains T :: "'b::euclidean_space set" where "closed T" "S homeomorphic T" +proof - + obtain U:: "('a*real)set" and h + where "closed U" and homU: "homeomorphism S U h fst" + using locally_compact_homeomorphism_projection_closed assms by metis + let ?BP = "Basis :: ('a*real) set" + have "DIM('a * real) \ DIM('b)" + by (simp add: Suc_leI dimlt) + then obtain basf :: "'a*real \ 'b" where surbf: "basf ` ?BP \ Basis" and injbf: "inj_on basf Basis" + by (metis finite_Basis card_le_inj) + define basg:: "'b \ 'a * real" where + "basg \ \i. inv_into Basis basf i" + have bgf[simp]: "basg (basf i) = i" if "i \ Basis" for i + using inv_into_f_f injbf that by (force simp: basg_def) + define f :: "'a*real \ 'b" where "f \ \u. \j\Basis. (u \ basg j) *\<^sub>R j" + have "linear f" + unfolding f_def + apply (intro linear_compose_setsum linearI ballI) + apply (auto simp: algebra_simps) + done + define g :: "'b \ 'a*real" where "g \ \z. (\i\Basis. (z \ basf i) *\<^sub>R i)" + have "linear g" + unfolding g_def + apply (intro linear_compose_setsum linearI ballI) + apply (auto simp: algebra_simps) + done + have *: "(\a \ Basis. a \ basf b * (x \ basg a)) = x \ b" if "b \ Basis" for x b + using surbf that by auto + have gf[simp]: "g (f x) = x" for x + apply (rule euclidean_eqI) + apply (simp add: f_def g_def inner_setsum_left scaleR_setsum_left algebra_simps) + apply (simp add: Groups_Big.setsum_right_distrib [symmetric] *) + done + then have "inj f" by (metis injI) + have gfU: "g ` f ` U = U" + by (rule set_eqI) (auto simp: image_def) + have "S homeomorphic U" + using homU homeomorphic_def by blast + also have "... homeomorphic f ` U" + apply (rule homeomorphicI [OF refl gfU]) + apply (meson \inj f\ \linear f\ homeomorphism_cont2 linear_homeomorphism_image) + using \linear g\ linear_continuous_on linear_conv_bounded_linear apply blast + apply auto + done + finally show ?thesis + apply (rule_tac T = "f ` U" in that) + apply (rule closed_injective_linear_image [OF \closed U\ \linear f\ \inj f\], assumption) + done +qed + + +lemma homeomorphic_convex_compact_lemma: + fixes s :: "'a::euclidean_space set" + assumes "convex s" + and "compact s" + and "cball 0 1 \ s" + shows "s homeomorphic (cball (0::'a) 1)" +proof (rule starlike_compact_projective_special[OF assms(2-3)]) + fix x u + assume "x \ s" and "0 \ u" and "u < (1::real)" + have "open (ball (u *\<^sub>R x) (1 - u))" + by (rule open_ball) + moreover have "u *\<^sub>R x \ ball (u *\<^sub>R x) (1 - u)" + unfolding centre_in_ball using \u < 1\ by simp + moreover have "ball (u *\<^sub>R x) (1 - u) \ s" + proof + fix y + assume "y \ ball (u *\<^sub>R x) (1 - u)" + then have "dist (u *\<^sub>R x) y < 1 - u" + unfolding mem_ball . + with \u < 1\ have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \ cball 0 1" + by (simp add: dist_norm inverse_eq_divide norm_minus_commute) + with assms(3) have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \ s" .. + with assms(1) have "(1 - u) *\<^sub>R ((y - u *\<^sub>R x) /\<^sub>R (1 - u)) + u *\<^sub>R x \ s" + using \x \ s\ \0 \ u\ \u < 1\ [THEN less_imp_le] by (rule convexD_alt) + then show "y \ s" using \u < 1\ + by simp + qed + ultimately have "u *\<^sub>R x \ interior s" .. + then show "u *\<^sub>R x \ s - frontier s" + using frontier_def and interior_subset by auto +qed + +proposition homeomorphic_convex_compact_cball: + fixes e :: real + and s :: "'a::euclidean_space set" + assumes "convex s" + and "compact s" + and "interior s \ {}" + and "e > 0" + shows "s homeomorphic (cball (b::'a) e)" +proof - + obtain a where "a \ interior s" + using assms(3) by auto + then obtain d where "d > 0" and d: "cball a d \ s" + unfolding mem_interior_cball by auto + let ?d = "inverse d" and ?n = "0::'a" + have "cball ?n 1 \ (\x. inverse d *\<^sub>R (x - a)) ` s" + apply rule + apply (rule_tac x="d *\<^sub>R x + a" in image_eqI) + defer + apply (rule d[unfolded subset_eq, rule_format]) + using \d > 0\ + unfolding mem_cball dist_norm + apply (auto simp add: mult_right_le_one_le) + done + then have "(\x. inverse d *\<^sub>R (x - a)) ` s homeomorphic cball ?n 1" + using homeomorphic_convex_compact_lemma[of "(\x. ?d *\<^sub>R -a + ?d *\<^sub>R x) ` s", + OF convex_affinity compact_affinity] + using assms(1,2) + by (auto simp add: scaleR_right_diff_distrib) + then show ?thesis + apply (rule_tac homeomorphic_trans[OF _ homeomorphic_balls(2)[of 1 _ ?n]]) + apply (rule homeomorphic_trans[OF homeomorphic_affinity[of "?d" s "?d *\<^sub>R -a"]]) + using \d>0\ \e>0\ + apply (auto simp add: scaleR_right_diff_distrib) + done +qed + +corollary homeomorphic_convex_compact: + fixes s :: "'a::euclidean_space set" + and t :: "'a set" + assumes "convex s" "compact s" "interior s \ {}" + and "convex t" "compact t" "interior t \ {}" + shows "s homeomorphic t" + using assms + by (meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym) + +end