src/HOL/Multivariate_Analysis/Homeomorphism.thy
changeset 63130 4ae5da02d627
child 63301 d3c87eb0bad2
equal deleted inserted replaced
63129:e5cb3440af74 63130:4ae5da02d627
       
     1 (*  Title:      HOL/Multivariate_Analysis/Homeomorphism.thy
       
     2     Author: LC Paulson (ported from HOL Light)
       
     3 *)
       
     4 
       
     5 section \<open>Homeomorphism Theorems\<close>
       
     6 
       
     7 theory Homeomorphism
       
     8 imports Path_Connected
       
     9 begin
       
    10 
       
    11 subsection \<open>Homeomorphism of all convex compact sets with nonempty interior\<close>
       
    12 
       
    13 proposition ray_to_rel_frontier:
       
    14   fixes a :: "'a::real_inner"
       
    15   assumes "bounded S"
       
    16       and a: "a \<in> rel_interior S"
       
    17       and aff: "(a + l) \<in> affine hull S"
       
    18       and "l \<noteq> 0"
       
    19   obtains d where "0 < d" "(a + d *\<^sub>R l) \<in> rel_frontier S"
       
    20            "\<And>e. \<lbrakk>0 \<le> e; e < d\<rbrakk> \<Longrightarrow> (a + e *\<^sub>R l) \<in> rel_interior S"
       
    21 proof -
       
    22   have aaff: "a \<in> affine hull S"
       
    23     by (meson a hull_subset rel_interior_subset rev_subsetD)
       
    24   let ?D = "{d. 0 < d \<and> a + d *\<^sub>R l \<notin> rel_interior S}"
       
    25   obtain B where "B > 0" and B: "S \<subseteq> ball a B"
       
    26     using bounded_subset_ballD [OF \<open>bounded S\<close>] by blast
       
    27   have "a + (B / norm l) *\<^sub>R l \<notin> ball a B"
       
    28     by (simp add: dist_norm \<open>l \<noteq> 0\<close>)
       
    29   with B have "a + (B / norm l) *\<^sub>R l \<notin> rel_interior S"
       
    30     using rel_interior_subset subsetCE by blast
       
    31   with \<open>B > 0\<close> \<open>l \<noteq> 0\<close> have nonMT: "?D \<noteq> {}"
       
    32     using divide_pos_pos zero_less_norm_iff by fastforce
       
    33   have bdd: "bdd_below ?D"
       
    34     by (metis (no_types, lifting) bdd_belowI le_less mem_Collect_eq)
       
    35   have relin_Ex: "\<And>x. x \<in> rel_interior S \<Longrightarrow>
       
    36                     \<exists>e>0. \<forall>x'\<in>affine hull S. dist x' x < e \<longrightarrow> x' \<in> rel_interior S"
       
    37     using openin_rel_interior [of S] by (simp add: openin_euclidean_subtopology_iff)
       
    38   define d where "d = Inf ?D"
       
    39   obtain \<epsilon> where "0 < \<epsilon>" and \<epsilon>: "\<And>\<eta>. \<lbrakk>0 \<le> \<eta>; \<eta> < \<epsilon>\<rbrakk> \<Longrightarrow> (a + \<eta> *\<^sub>R l) \<in> rel_interior S"
       
    40   proof -
       
    41     obtain e where "e>0"
       
    42             and e: "\<And>x'. x' \<in> affine hull S \<Longrightarrow> dist x' a < e \<Longrightarrow> x' \<in> rel_interior S"
       
    43       using relin_Ex a by blast
       
    44     show thesis
       
    45     proof (rule_tac \<epsilon> = "e / norm l" in that)
       
    46       show "0 < e / norm l" by (simp add: \<open>0 < e\<close> \<open>l \<noteq> 0\<close>)
       
    47     next
       
    48       show "a + \<eta> *\<^sub>R l \<in> rel_interior S" if "0 \<le> \<eta>" "\<eta> < e / norm l" for \<eta>
       
    49       proof (rule e)
       
    50         show "a + \<eta> *\<^sub>R l \<in> affine hull S"
       
    51           by (metis (no_types) add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff)
       
    52         show "dist (a + \<eta> *\<^sub>R l) a < e"
       
    53           using that by (simp add: \<open>l \<noteq> 0\<close> dist_norm pos_less_divide_eq)
       
    54       qed
       
    55     qed
       
    56   qed
       
    57   have inint: "\<And>e. \<lbrakk>0 \<le> e; e < d\<rbrakk> \<Longrightarrow> a + e *\<^sub>R l \<in> rel_interior S"
       
    58     unfolding d_def using cInf_lower [OF _ bdd]
       
    59     by (metis (no_types, lifting) a add.right_neutral le_less mem_Collect_eq not_less real_vector.scale_zero_left)
       
    60   have "\<epsilon> \<le> d"
       
    61     unfolding d_def
       
    62     apply (rule cInf_greatest [OF nonMT])
       
    63     using \<epsilon> dual_order.strict_implies_order le_less_linear by blast
       
    64   with \<open>0 < \<epsilon>\<close> have "0 < d" by simp
       
    65   have "a + d *\<^sub>R l \<notin> rel_interior S"
       
    66   proof
       
    67     assume adl: "a + d *\<^sub>R l \<in> rel_interior S"
       
    68     obtain e where "e > 0"
       
    69              and e: "\<And>x'. x' \<in> affine hull S \<Longrightarrow> dist x' (a + d *\<^sub>R l) < e \<Longrightarrow> x' \<in> rel_interior S"
       
    70       using relin_Ex adl by blast
       
    71     have "d + e / norm l \<le> Inf {d. 0 < d \<and> a + d *\<^sub>R l \<notin> rel_interior S}"
       
    72     proof (rule cInf_greatest [OF nonMT], clarsimp)
       
    73       fix x::real
       
    74       assume "0 < x" and nonrel: "a + x *\<^sub>R l \<notin> rel_interior S"
       
    75       show "d + e / norm l \<le> x"
       
    76       proof (cases "x < d")
       
    77         case True with inint nonrel \<open>0 < x\<close>
       
    78           show ?thesis by auto
       
    79       next
       
    80         case False
       
    81           then have dle: "x < d + e / norm l \<Longrightarrow> dist (a + x *\<^sub>R l) (a + d *\<^sub>R l) < e"
       
    82             by (simp add: field_simps \<open>l \<noteq> 0\<close>)
       
    83           have ain: "a + x *\<^sub>R l \<in> affine hull S"
       
    84             by (metis add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff)
       
    85           show ?thesis
       
    86             using e [OF ain] nonrel dle by force
       
    87       qed
       
    88     qed
       
    89     then show False
       
    90       using \<open>0 < e\<close> \<open>l \<noteq> 0\<close> by (simp add: d_def [symmetric] divide_simps)
       
    91   qed
       
    92   moreover have "a + d *\<^sub>R l \<in> closure S"
       
    93   proof (clarsimp simp: closure_approachable)
       
    94     fix \<eta>::real assume "0 < \<eta>"
       
    95     have 1: "a + (d - min d (\<eta> / 2 / norm l)) *\<^sub>R l \<in> S"
       
    96       apply (rule subsetD [OF rel_interior_subset inint])
       
    97       using \<open>l \<noteq> 0\<close> \<open>0 < d\<close> \<open>0 < \<eta>\<close> by auto
       
    98     have "norm l * min d (\<eta> / (norm l * 2)) \<le> norm l * (\<eta> / (norm l * 2))"
       
    99       by (metis min_def mult_left_mono norm_ge_zero order_refl)
       
   100     also have "... < \<eta>"
       
   101       using \<open>l \<noteq> 0\<close> \<open>0 < \<eta>\<close> by (simp add: divide_simps)
       
   102     finally have 2: "norm l * min d (\<eta> / (norm l * 2)) < \<eta>" .
       
   103     show "\<exists>y\<in>S. dist y (a + d *\<^sub>R l) < \<eta>"
       
   104       apply (rule_tac x="a + (d - min d (\<eta> / 2 / norm l)) *\<^sub>R l" in bexI)
       
   105       using 1 2 \<open>0 < d\<close> \<open>0 < \<eta>\<close> apply (auto simp: algebra_simps)
       
   106       done
       
   107   qed
       
   108   ultimately have infront: "a + d *\<^sub>R l \<in> rel_frontier S"
       
   109     by (simp add: rel_frontier_def)
       
   110   show ?thesis
       
   111     by (rule that [OF \<open>0 < d\<close> infront inint])
       
   112 qed
       
   113 
       
   114 corollary ray_to_frontier:
       
   115   fixes a :: "'a::euclidean_space"
       
   116   assumes "bounded S"
       
   117       and a: "a \<in> interior S"
       
   118       and "l \<noteq> 0"
       
   119   obtains d where "0 < d" "(a + d *\<^sub>R l) \<in> frontier S"
       
   120            "\<And>e. \<lbrakk>0 \<le> e; e < d\<rbrakk> \<Longrightarrow> (a + e *\<^sub>R l) \<in> interior S"
       
   121 proof -
       
   122   have "interior S = rel_interior S"
       
   123     using a rel_interior_nonempty_interior by auto
       
   124   then have "a \<in> rel_interior S"
       
   125     using a by simp
       
   126   then show ?thesis
       
   127     apply (rule ray_to_rel_frontier [OF \<open>bounded S\<close> _ _ \<open>l \<noteq> 0\<close>])
       
   128      using a affine_hull_nonempty_interior apply blast
       
   129     by (simp add: \<open>interior S = rel_interior S\<close> frontier_def rel_frontier_def that)
       
   130 qed
       
   131 
       
   132 proposition
       
   133   fixes S :: "'a::euclidean_space set"
       
   134   assumes "compact S" and 0: "0 \<in> rel_interior S"
       
   135       and star: "\<And>x. x \<in> S \<Longrightarrow> open_segment 0 x \<subseteq> rel_interior S"
       
   136     shows starlike_compact_projective1_0:
       
   137             "S - rel_interior S homeomorphic sphere 0 1 \<inter> affine hull S"
       
   138             (is "?SMINUS homeomorphic ?SPHER")
       
   139       and starlike_compact_projective2_0:
       
   140             "S homeomorphic cball 0 1 \<inter> affine hull S"
       
   141             (is "S homeomorphic ?CBALL")
       
   142 proof -
       
   143   have starI: "(u *\<^sub>R x) \<in> rel_interior S" if "x \<in> S" "0 \<le> u" "u < 1" for x u
       
   144   proof (cases "x=0 \<or> u=0")
       
   145     case True with 0 show ?thesis by force
       
   146   next
       
   147     case False with that show ?thesis
       
   148       by (auto simp: in_segment intro: star [THEN subsetD])
       
   149   qed
       
   150   have "0 \<in> S"  using assms rel_interior_subset by auto
       
   151   define proj where "proj \<equiv> \<lambda>x::'a. x /\<^sub>R norm x"
       
   152   have eqI: "x = y" if "proj x = proj y" "norm x = norm y" for x y
       
   153     using that  by (force simp: proj_def)
       
   154   then have iff_eq: "\<And>x y. (proj x = proj y \<and> norm x = norm y) \<longleftrightarrow> x = y"
       
   155     by blast
       
   156   have projI: "x \<in> affine hull S \<Longrightarrow> proj x \<in> affine hull S" for x
       
   157     by (metis \<open>0 \<in> S\<close> affine_hull_span_0 hull_inc span_mul proj_def)
       
   158   have nproj1 [simp]: "x \<noteq> 0 \<Longrightarrow> norm(proj x) = 1" for x
       
   159     by (simp add: proj_def)
       
   160   have proj0_iff [simp]: "proj x = 0 \<longleftrightarrow> x = 0" for x
       
   161     by (simp add: proj_def)
       
   162   have cont_proj: "continuous_on (UNIV - {0}) proj"
       
   163     unfolding proj_def by (rule continuous_intros | force)+
       
   164   have proj_spherI: "\<And>x. \<lbrakk>x \<in> affine hull S; x \<noteq> 0\<rbrakk> \<Longrightarrow> proj x \<in> ?SPHER"
       
   165     by (simp add: projI)
       
   166   have "bounded S" "closed S"
       
   167     using \<open>compact S\<close> compact_eq_bounded_closed by blast+
       
   168   have inj_on_proj: "inj_on proj (S - rel_interior S)"
       
   169   proof
       
   170     fix x y
       
   171     assume x: "x \<in> S - rel_interior S"
       
   172        and y: "y \<in> S - rel_interior S" and eq: "proj x = proj y"
       
   173     then have xynot: "x \<noteq> 0" "y \<noteq> 0" "x \<in> S" "y \<in> S" "x \<notin> rel_interior S" "y \<notin> rel_interior S"
       
   174       using 0 by auto
       
   175     consider "norm x = norm y" | "norm x < norm y" | "norm x > norm y" by linarith
       
   176     then show "x = y"
       
   177     proof cases
       
   178       assume "norm x = norm y"
       
   179         with iff_eq eq show "x = y" by blast
       
   180     next
       
   181       assume *: "norm x < norm y"
       
   182       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))"
       
   183         by force
       
   184       then have "proj ((norm x / norm y) *\<^sub>R y) = proj x"
       
   185         by (metis (no_types) divide_inverse local.proj_def eq scaleR_scaleR)
       
   186       then have [simp]: "(norm x / norm y) *\<^sub>R y = x"
       
   187         by (rule eqI) (simp add: \<open>y \<noteq> 0\<close>)
       
   188       have no: "0 \<le> norm x / norm y" "norm x / norm y < 1"
       
   189         using * by (auto simp: divide_simps)
       
   190       then show "x = y"
       
   191         using starI [OF \<open>y \<in> S\<close> no] xynot by auto
       
   192     next
       
   193       assume *: "norm x > norm y"
       
   194       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))"
       
   195         by force
       
   196       then have "proj ((norm y / norm x) *\<^sub>R x) = proj y"
       
   197         by (metis (no_types) divide_inverse local.proj_def eq scaleR_scaleR)
       
   198       then have [simp]: "(norm y / norm x) *\<^sub>R x = y"
       
   199         by (rule eqI) (simp add: \<open>x \<noteq> 0\<close>)
       
   200       have no: "0 \<le> norm y / norm x" "norm y / norm x < 1"
       
   201         using * by (auto simp: divide_simps)
       
   202       then show "x = y"
       
   203         using starI [OF \<open>x \<in> S\<close> no] xynot by auto
       
   204     qed
       
   205   qed
       
   206   have "\<exists>surf. homeomorphism (S - rel_interior S) ?SPHER proj surf"
       
   207   proof (rule homeomorphism_compact)
       
   208     show "compact (S - rel_interior S)"
       
   209        using \<open>compact S\<close> compact_rel_boundary by blast
       
   210     show "continuous_on (S - rel_interior S) proj"
       
   211       using 0 by (blast intro: continuous_on_subset [OF cont_proj])
       
   212     show "proj ` (S - rel_interior S) = ?SPHER"
       
   213     proof
       
   214       show "proj ` (S - rel_interior S) \<subseteq> ?SPHER"
       
   215         using 0 by (force simp: hull_inc projI intro: nproj1)
       
   216       show "?SPHER \<subseteq> proj ` (S - rel_interior S)"
       
   217       proof (clarsimp simp: proj_def)
       
   218         fix x
       
   219         assume "x \<in> affine hull S" and nox: "norm x = 1"
       
   220         then have "x \<noteq> 0" by auto
       
   221         obtain d where "0 < d" and dx: "(d *\<^sub>R x) \<in> rel_frontier S"
       
   222                    and ri: "\<And>e. \<lbrakk>0 \<le> e; e < d\<rbrakk> \<Longrightarrow> (e *\<^sub>R x) \<in> rel_interior S"
       
   223           using ray_to_rel_frontier [OF \<open>bounded S\<close> 0] \<open>x \<in> affine hull S\<close> \<open>x \<noteq> 0\<close> by auto
       
   224         show "x \<in> (\<lambda>x. x /\<^sub>R norm x) ` (S - rel_interior S)"
       
   225           apply (rule_tac x="d *\<^sub>R x" in image_eqI)
       
   226           using \<open>0 < d\<close>
       
   227           using dx \<open>closed S\<close> apply (auto simp: rel_frontier_def divide_simps nox)
       
   228           done
       
   229       qed
       
   230     qed
       
   231   qed (rule inj_on_proj)
       
   232   then obtain surf where surf: "homeomorphism (S - rel_interior S) ?SPHER proj surf"
       
   233     by blast
       
   234   then have cont_surf: "continuous_on (proj ` (S - rel_interior S)) surf"
       
   235     by (auto simp: homeomorphism_def)
       
   236   have surf_nz: "\<And>x. x \<in> ?SPHER \<Longrightarrow> surf x \<noteq> 0"
       
   237     by (metis "0" DiffE homeomorphism_def imageI surf)
       
   238   have cont_nosp: "continuous_on (?SPHER) (\<lambda>x. norm x *\<^sub>R ((surf o proj) x))"
       
   239     apply (rule continuous_intros)+
       
   240     apply (rule continuous_on_subset [OF cont_proj], force)
       
   241     apply (rule continuous_on_subset [OF cont_surf])
       
   242     apply (force simp: homeomorphism_image1 [OF surf] dest: proj_spherI)
       
   243     done
       
   244   have surfpS: "\<And>x. \<lbrakk>norm x = 1; x \<in> affine hull S\<rbrakk> \<Longrightarrow> surf (proj x) \<in> S"
       
   245     by (metis (full_types) DiffE \<open>0 \<in> S\<close> homeomorphism_def image_eqI norm_zero proj_spherI real_vector.scale_zero_left scaleR_one surf)
       
   246   have *: "\<exists>y. norm y = 1 \<and> y \<in> affine hull S \<and> x = surf (proj y)"
       
   247           if "x \<in> S" "x \<notin> rel_interior S" for x
       
   248   proof -
       
   249     have "proj x \<in> ?SPHER"
       
   250       by (metis (full_types) "0" hull_inc proj_spherI that)
       
   251     moreover have "surf (proj x) = x"
       
   252       by (metis Diff_iff homeomorphism_def surf that)
       
   253     ultimately show ?thesis
       
   254       by (metis \<open>\<And>x. x \<in> ?SPHER \<Longrightarrow> surf x \<noteq> 0\<close> hull_inc inverse_1 local.proj_def norm_sgn projI scaleR_one sgn_div_norm that(1))
       
   255   qed
       
   256   have surfp_notin: "\<And>x. \<lbrakk>norm x = 1; x \<in> affine hull S\<rbrakk> \<Longrightarrow> surf (proj x) \<notin> rel_interior S"
       
   257     by (metis (full_types) DiffE one_neq_zero homeomorphism_def image_eqI norm_zero proj_spherI surf)
       
   258   have no_sp_im: "(\<lambda>x. norm x *\<^sub>R surf (proj x)) ` (?SPHER) = S - rel_interior S"
       
   259     by (auto simp: surfpS image_def Bex_def surfp_notin *)
       
   260   have inj_spher: "inj_on (\<lambda>x. norm x *\<^sub>R surf (proj x)) ?SPHER"
       
   261   proof
       
   262     fix x y
       
   263     assume xy: "x \<in> ?SPHER" "y \<in> ?SPHER"
       
   264        and eq: " norm x *\<^sub>R surf (proj x) = norm y *\<^sub>R surf (proj y)"
       
   265     then have "norm x = 1" "norm y = 1" "x \<in> affine hull S" "y \<in> affine hull S"
       
   266       using 0 by auto
       
   267     with eq show "x = y"
       
   268       by (simp add: proj_def) (metis surf xy homeomorphism_def)
       
   269   qed
       
   270   have co01: "compact ?SPHER"
       
   271     by (simp add: closed_affine_hull compact_Int_closed)
       
   272   show "?SMINUS homeomorphic ?SPHER"
       
   273     apply (subst homeomorphic_sym)
       
   274     apply (rule homeomorphic_compact [OF co01 cont_nosp [unfolded o_def] no_sp_im inj_spher])
       
   275     done
       
   276   have proj_scaleR: "\<And>a x. 0 < a \<Longrightarrow> proj (a *\<^sub>R x) = proj x"
       
   277     by (simp add: proj_def)
       
   278   have cont_sp0: "continuous_on (affine hull S - {0}) (surf o proj)"
       
   279     apply (rule continuous_on_compose [OF continuous_on_subset [OF cont_proj]], force)
       
   280     apply (rule continuous_on_subset [OF cont_surf])
       
   281     using homeomorphism_image1 proj_spherI surf by fastforce
       
   282   obtain B where "B>0" and B: "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> B"
       
   283     by (metis compact_imp_bounded \<open>compact S\<close> bounded_pos_less less_eq_real_def)
       
   284   have cont_nosp: "continuous (at x within ?CBALL) (\<lambda>x. norm x *\<^sub>R surf (proj x))"
       
   285          if "norm x \<le> 1" "x \<in> affine hull S" for x
       
   286   proof (cases "x=0")
       
   287     case True
       
   288     show ?thesis using True
       
   289       apply (simp add: continuous_within)
       
   290       apply (rule lim_null_scaleR_bounded [where B=B])
       
   291       apply (simp_all add: tendsto_norm_zero eventually_at)
       
   292       apply (rule_tac x=B in exI)
       
   293       using B surfpS proj_def projI apply (auto simp: \<open>B > 0\<close>)
       
   294       done
       
   295   next
       
   296     case False
       
   297     then have "\<forall>\<^sub>F x in at x. (x \<in> affine hull S - {0}) = (x \<in> affine hull S)"
       
   298       apply (simp add: eventually_at)
       
   299       apply (rule_tac x="norm x" in exI)
       
   300       apply (auto simp: False)
       
   301       done
       
   302     with cont_sp0 have *: "continuous (at x within affine hull S) (\<lambda>x. surf (proj x))"
       
   303       apply (simp add: continuous_on_eq_continuous_within)
       
   304       apply (drule_tac x=x in bspec, force simp: False that)
       
   305       apply (simp add: continuous_within Lim_transform_within_set)
       
   306       done
       
   307     show ?thesis
       
   308       apply (rule continuous_within_subset [where s = "affine hull S", OF _ Int_lower2])
       
   309       apply (rule continuous_intros *)+
       
   310       done
       
   311   qed
       
   312   have cont_nosp2: "continuous_on ?CBALL (\<lambda>x. norm x *\<^sub>R ((surf o proj) x))"
       
   313     by (simp add: continuous_on_eq_continuous_within cont_nosp)
       
   314   have "norm y *\<^sub>R surf (proj y) \<in> S"  if "y \<in> cball 0 1" and yaff: "y \<in> affine hull S" for y
       
   315   proof (cases "y=0")
       
   316     case True then show ?thesis
       
   317       by (simp add: \<open>0 \<in> S\<close>)
       
   318   next
       
   319     case False
       
   320     then have "norm y *\<^sub>R surf (proj y) = norm y *\<^sub>R surf (proj (y /\<^sub>R norm y))"
       
   321       by (simp add: proj_def)
       
   322     have "norm y \<le> 1" using that by simp
       
   323     have "surf (proj (y /\<^sub>R norm y)) \<in> S"
       
   324       apply (rule surfpS)
       
   325       using proj_def projI yaff
       
   326       by (auto simp: False)
       
   327     then have "surf (proj y) \<in> S"
       
   328       by (simp add: False proj_def)
       
   329     then show "norm y *\<^sub>R surf (proj y) \<in> S"
       
   330       by (metis dual_order.antisym le_less_linear norm_ge_zero rel_interior_subset scaleR_one
       
   331                 starI subset_eq \<open>norm y \<le> 1\<close>)
       
   332   qed
       
   333   moreover have "x \<in> (\<lambda>x. norm x *\<^sub>R surf (proj x)) ` (?CBALL)" if "x \<in> S" for x
       
   334   proof (cases "x=0")
       
   335     case True with that hull_inc  show ?thesis by fastforce
       
   336   next
       
   337     case False
       
   338     then have psp: "proj (surf (proj x)) = proj x"
       
   339       by (metis homeomorphism_def hull_inc proj_spherI surf that)
       
   340     have nxx: "norm x *\<^sub>R proj x = x"
       
   341       by (simp add: False local.proj_def)
       
   342     have affineI: "(1 / norm (surf (proj x))) *\<^sub>R x \<in> affine hull S"
       
   343       by (metis \<open>0 \<in> S\<close> affine_hull_span_0 hull_inc span_clauses(4) that)
       
   344     have sproj_nz: "surf (proj x) \<noteq> 0"
       
   345       by (metis False proj0_iff psp)
       
   346     then have "proj x = proj (proj x)"
       
   347       by (metis False nxx proj_scaleR zero_less_norm_iff)
       
   348     moreover have scaleproj: "\<And>a r. r *\<^sub>R proj a = (r / norm a) *\<^sub>R a"
       
   349       by (simp add: divide_inverse local.proj_def)
       
   350     ultimately have "(norm (surf (proj x)) / norm x) *\<^sub>R x \<notin> rel_interior S"
       
   351       by (metis (no_types) sproj_nz divide_self_if hull_inc norm_eq_zero nproj1 projI psp scaleR_one surfp_notin that)
       
   352     then have "(norm (surf (proj x)) / norm x) \<ge> 1"
       
   353       using starI [OF that] by (meson starI [OF that] le_less_linear norm_ge_zero zero_le_divide_iff)
       
   354     then have nole: "norm x \<le> norm (surf (proj x))"
       
   355       by (simp add: le_divide_eq_1)
       
   356     show ?thesis
       
   357       apply (rule_tac x="inverse(norm(surf (proj x))) *\<^sub>R x" in image_eqI)
       
   358       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)
       
   359       apply (auto simp: divide_simps nole affineI)
       
   360       done
       
   361   qed
       
   362   ultimately have im_cball: "(\<lambda>x. norm x *\<^sub>R surf (proj x)) ` ?CBALL = S"
       
   363     by blast
       
   364   have inj_cball: "inj_on (\<lambda>x. norm x *\<^sub>R surf (proj x)) ?CBALL"
       
   365   proof
       
   366     fix x y
       
   367     assume "x \<in> ?CBALL" "y \<in> ?CBALL"
       
   368        and eq: "norm x *\<^sub>R surf (proj x) = norm y *\<^sub>R surf (proj y)"
       
   369     then have x: "x \<in> affine hull S" and y: "y \<in> affine hull S"
       
   370       using 0 by auto
       
   371     show "x = y"
       
   372     proof (cases "x=0 \<or> y=0")
       
   373       case True then show "x = y" using eq proj_spherI surf_nz x y by force
       
   374     next
       
   375       case False
       
   376       with x y have speq: "surf (proj x) = surf (proj y)"
       
   377         by (metis eq homeomorphism_apply2 proj_scaleR proj_spherI surf zero_less_norm_iff)
       
   378       then have "norm x = norm y"
       
   379         by (metis \<open>x \<in> affine hull S\<close> \<open>y \<in> affine hull S\<close> eq proj_spherI real_vector.scale_cancel_right surf_nz)
       
   380       moreover have "proj x = proj y"
       
   381         by (metis (no_types) False speq homeomorphism_apply2 proj_spherI surf x y)
       
   382       ultimately show "x = y"
       
   383         using eq eqI by blast
       
   384     qed
       
   385   qed
       
   386   have co01: "compact ?CBALL"
       
   387     by (simp add: closed_affine_hull compact_Int_closed)
       
   388   show "S homeomorphic ?CBALL"
       
   389     apply (subst homeomorphic_sym)
       
   390     apply (rule homeomorphic_compact [OF co01 cont_nosp2 [unfolded o_def] im_cball inj_cball])
       
   391     done
       
   392 qed
       
   393 
       
   394 corollary
       
   395   fixes S :: "'a::euclidean_space set"
       
   396   assumes "compact S" and a: "a \<in> rel_interior S"
       
   397       and star: "\<And>x. x \<in> S \<Longrightarrow> open_segment a x \<subseteq> rel_interior S"
       
   398     shows starlike_compact_projective1:
       
   399             "S - rel_interior S homeomorphic sphere a 1 \<inter> affine hull S"
       
   400       and starlike_compact_projective2:
       
   401             "S homeomorphic cball a 1 \<inter> affine hull S"
       
   402 proof -
       
   403   have 1: "compact (op+ (-a) ` S)" by (meson assms compact_translation)
       
   404   have 2: "0 \<in> rel_interior (op+ (-a) ` S)"
       
   405     by (simp add: a rel_interior_translation)
       
   406   have 3: "open_segment 0 x \<subseteq> rel_interior (op+ (-a) ` S)" if "x \<in> (op+ (-a) ` S)" for x
       
   407   proof -
       
   408     have "x+a \<in> S" using that by auto
       
   409     then have "open_segment a (x+a) \<subseteq> rel_interior S" by (metis star)
       
   410     then show ?thesis using open_segment_translation
       
   411       using rel_interior_translation by fastforce
       
   412   qed
       
   413   have "S - rel_interior S homeomorphic (op+ (-a) ` S) - rel_interior (op+ (-a) ` S)"
       
   414     by (metis rel_interior_translation translation_diff homeomorphic_translation)
       
   415   also have "... homeomorphic sphere 0 1 \<inter> affine hull (op+ (-a) ` S)"
       
   416     by (rule starlike_compact_projective1_0 [OF 1 2 3])
       
   417   also have "... = op+ (-a) ` (sphere a 1 \<inter> affine hull S)"
       
   418     by (metis affine_hull_translation left_minus sphere_translation translation_Int)
       
   419   also have "... homeomorphic sphere a 1 \<inter> affine hull S"
       
   420     using homeomorphic_translation homeomorphic_sym by blast
       
   421   finally show "S - rel_interior S homeomorphic sphere a 1 \<inter> affine hull S" .
       
   422 
       
   423   have "S homeomorphic (op+ (-a) ` S)"
       
   424     by (metis homeomorphic_translation)
       
   425   also have "... homeomorphic cball 0 1 \<inter> affine hull (op+ (-a) ` S)"
       
   426     by (rule starlike_compact_projective2_0 [OF 1 2 3])
       
   427   also have "... = op+ (-a) ` (cball a 1 \<inter> affine hull S)"
       
   428     by (metis affine_hull_translation left_minus cball_translation translation_Int)
       
   429   also have "... homeomorphic cball a 1 \<inter> affine hull S"
       
   430     using homeomorphic_translation homeomorphic_sym by blast
       
   431   finally show "S homeomorphic cball a 1 \<inter> affine hull S" .
       
   432 qed
       
   433 
       
   434 corollary starlike_compact_projective_special:
       
   435   assumes "compact S"
       
   436     and cb01: "cball (0::'a::euclidean_space) 1 \<subseteq> S"
       
   437     and scale: "\<And>x u. \<lbrakk>x \<in> S; 0 \<le> u; u < 1\<rbrakk> \<Longrightarrow> u *\<^sub>R x \<in> S - frontier S"
       
   438   shows "S homeomorphic (cball (0::'a::euclidean_space) 1)"
       
   439 proof -
       
   440   have "ball 0 1 \<subseteq> interior S"
       
   441     using cb01 interior_cball interior_mono by blast
       
   442   then have 0: "0 \<in> rel_interior S"
       
   443     by (meson centre_in_ball subsetD interior_subset_rel_interior le_numeral_extra(2) not_le)
       
   444   have [simp]: "affine hull S = UNIV"
       
   445     using \<open>ball 0 1 \<subseteq> interior S\<close> by (auto intro!: affine_hull_nonempty_interior)
       
   446   have star: "open_segment 0 x \<subseteq> rel_interior S" if "x \<in> S" for x
       
   447   proof 
       
   448     fix p assume "p \<in> open_segment 0 x"
       
   449     then obtain u where "x \<noteq> 0" and u: "0 \<le> u" "u < 1" and p: "u *\<^sub>R x = p"
       
   450       by (auto simp: in_segment) 
       
   451     then show "p \<in> rel_interior S"
       
   452       using scale [OF that u] closure_subset frontier_def interior_subset_rel_interior by fastforce
       
   453   qed
       
   454   show ?thesis
       
   455     using starlike_compact_projective2_0 [OF \<open>compact S\<close> 0 star] by simp
       
   456 qed
       
   457 
       
   458 lemma homeomorphic_convex_lemma:
       
   459   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
   460   assumes "convex S" "compact S" "convex T" "compact T"
       
   461       and affeq: "aff_dim S = aff_dim T"
       
   462     shows "(S - rel_interior S) homeomorphic (T - rel_interior T) \<and>
       
   463            S homeomorphic T"
       
   464 proof (cases "rel_interior S = {} \<or> rel_interior T = {}")
       
   465   case True
       
   466     then show ?thesis
       
   467       by (metis Diff_empty affeq \<open>convex S\<close> \<open>convex T\<close> aff_dim_empty homeomorphic_empty rel_interior_eq_empty aff_dim_empty)
       
   468 next
       
   469   case False
       
   470   then obtain a b where a: "a \<in> rel_interior S" and b: "b \<in> rel_interior T" by auto
       
   471   have starS: "\<And>x. x \<in> S \<Longrightarrow> open_segment a x \<subseteq> rel_interior S"
       
   472     using rel_interior_closure_convex_segment
       
   473           a \<open>convex S\<close> closure_subset subsetCE by blast
       
   474   have starT: "\<And>x. x \<in> T \<Longrightarrow> open_segment b x \<subseteq> rel_interior T"
       
   475     using rel_interior_closure_convex_segment
       
   476           b \<open>convex T\<close> closure_subset subsetCE by blast
       
   477   let ?aS = "op+ (-a) ` S" and ?bT = "op+ (-b) ` T"
       
   478   have 0: "0 \<in> affine hull ?aS" "0 \<in> affine hull ?bT"
       
   479     by (metis a b subsetD hull_inc image_eqI left_minus rel_interior_subset)+
       
   480   have subs: "subspace (span ?aS)" "subspace (span ?bT)"
       
   481     by (rule subspace_span)+
       
   482   moreover
       
   483   have "dim (span (op + (- a) ` S)) = dim (span (op + (- b) ` T))"
       
   484     by (metis 0 aff_dim_translation_eq aff_dim_zero affeq dim_span nat_int)
       
   485   ultimately obtain f g where "linear f" "linear g"
       
   486                 and fim: "f ` span ?aS = span ?bT"
       
   487                 and gim: "g ` span ?bT = span ?aS"
       
   488                 and fno: "\<And>x. x \<in> span ?aS \<Longrightarrow> norm(f x) = norm x"
       
   489                 and gno: "\<And>x. x \<in> span ?bT \<Longrightarrow> norm(g x) = norm x"
       
   490                 and gf: "\<And>x. x \<in> span ?aS \<Longrightarrow> g(f x) = x"
       
   491                 and fg: "\<And>x. x \<in> span ?bT \<Longrightarrow> f(g x) = x"
       
   492     by (rule isometries_subspaces) blast
       
   493   have [simp]: "continuous_on A f" for A
       
   494     using \<open>linear f\<close> linear_conv_bounded_linear linear_continuous_on by blast
       
   495   have [simp]: "continuous_on B g" for B
       
   496     using \<open>linear g\<close> linear_conv_bounded_linear linear_continuous_on by blast
       
   497   have eqspanS: "affine hull ?aS = span ?aS"
       
   498     by (metis a affine_hull_span_0 subsetD hull_inc image_eqI left_minus rel_interior_subset)
       
   499   have eqspanT: "affine hull ?bT = span ?bT"
       
   500     by (metis b affine_hull_span_0 subsetD hull_inc image_eqI left_minus rel_interior_subset)
       
   501   have "S homeomorphic cball a 1 \<inter> affine hull S"
       
   502     by (rule starlike_compact_projective2 [OF \<open>compact S\<close> a starS])
       
   503   also have "... homeomorphic op+ (-a) ` (cball a 1 \<inter> affine hull S)"
       
   504     by (metis homeomorphic_translation)
       
   505   also have "... = cball 0 1 \<inter> op+ (-a) ` (affine hull S)"
       
   506     by (auto simp: dist_norm)
       
   507   also have "... = cball 0 1 \<inter> span ?aS"
       
   508     using eqspanS affine_hull_translation by blast
       
   509   also have "... homeomorphic cball 0 1 \<inter> span ?bT"
       
   510     proof (rule homeomorphicI [where f=f and g=g])
       
   511       show fim1: "f ` (cball 0 1 \<inter> span ?aS) = cball 0 1 \<inter> span ?bT"
       
   512         apply (rule subset_antisym)
       
   513          using fim fno apply (force simp:, clarify)
       
   514         by (metis IntI fg gim gno image_eqI mem_cball_0)
       
   515       show "g ` (cball 0 1 \<inter> span ?bT) = cball 0 1 \<inter> span ?aS"
       
   516         apply (rule subset_antisym)
       
   517          using gim gno apply (force simp:, clarify)
       
   518         by (metis IntI fim1 gf image_eqI)
       
   519     qed (auto simp: fg gf)
       
   520   also have "... = cball 0 1 \<inter> op+ (-b) ` (affine hull T)"
       
   521     using eqspanT affine_hull_translation by blast
       
   522   also have "... = op+ (-b) ` (cball b 1 \<inter> affine hull T)"
       
   523     by (auto simp: dist_norm)
       
   524   also have "... homeomorphic (cball b 1 \<inter> affine hull T)"
       
   525     by (metis homeomorphic_translation homeomorphic_sym)
       
   526   also have "... homeomorphic T"
       
   527     by (metis starlike_compact_projective2 [OF \<open>compact T\<close> b starT] homeomorphic_sym)
       
   528   finally have 1: "S homeomorphic T" .
       
   529 
       
   530   have "S - rel_interior S homeomorphic sphere a 1 \<inter> affine hull S"
       
   531     by (rule starlike_compact_projective1 [OF \<open>compact S\<close> a starS])
       
   532   also have "... homeomorphic op+ (-a) ` (sphere a 1 \<inter> affine hull S)"
       
   533     by (metis homeomorphic_translation)
       
   534   also have "... = sphere 0 1 \<inter> op+ (-a) ` (affine hull S)"
       
   535     by (auto simp: dist_norm)
       
   536   also have "... = sphere 0 1 \<inter> span ?aS"
       
   537     using eqspanS affine_hull_translation by blast
       
   538   also have "... homeomorphic sphere 0 1 \<inter> span ?bT"
       
   539     proof (rule homeomorphicI [where f=f and g=g])
       
   540       show fim1: "f ` (sphere 0 1 \<inter> span ?aS) = sphere 0 1 \<inter> span ?bT"
       
   541         apply (rule subset_antisym)
       
   542         using fim fno apply (force simp:, clarify)
       
   543         by (metis IntI fg gim gno image_eqI mem_sphere_0)
       
   544       show "g ` (sphere 0 1 \<inter> span ?bT) = sphere 0 1 \<inter> span ?aS"
       
   545         apply (rule subset_antisym)
       
   546         using gim gno apply (force simp:, clarify)
       
   547         by (metis IntI fim1 gf image_eqI)
       
   548     qed (auto simp: fg gf)
       
   549   also have "... = sphere 0 1 \<inter> op+ (-b) ` (affine hull T)"
       
   550     using eqspanT affine_hull_translation by blast
       
   551   also have "... = op+ (-b) ` (sphere b 1 \<inter> affine hull T)"
       
   552     by (auto simp: dist_norm)
       
   553   also have "... homeomorphic (sphere b 1 \<inter> affine hull T)"
       
   554     by (metis homeomorphic_translation homeomorphic_sym)
       
   555   also have "... homeomorphic T - rel_interior T"
       
   556     by (metis starlike_compact_projective1 [OF \<open>compact T\<close> b starT] homeomorphic_sym)
       
   557   finally have 2: "S - rel_interior S homeomorphic T - rel_interior T" .
       
   558   show ?thesis
       
   559     using 1 2 by blast
       
   560 qed
       
   561 
       
   562 lemma homeomorphic_convex_compact_sets:
       
   563   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
   564   assumes "convex S" "compact S" "convex T" "compact T"
       
   565       and affeq: "aff_dim S = aff_dim T"
       
   566     shows "S homeomorphic T"
       
   567 using homeomorphic_convex_lemma [OF assms] assms
       
   568 by (auto simp: rel_frontier_def)
       
   569 
       
   570 lemma homeomorphic_rel_frontiers_convex_bounded_sets:
       
   571   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
   572   assumes "convex S" "bounded S" "convex T" "bounded T"
       
   573       and affeq: "aff_dim S = aff_dim T"
       
   574     shows  "rel_frontier S homeomorphic rel_frontier T"
       
   575 using assms homeomorphic_convex_lemma [of "closure S" "closure T"]
       
   576 by (simp add: rel_frontier_def convex_rel_interior_closure)
       
   577 
       
   578 
       
   579 subsection\<open>Homeomorphisms between punctured spheres and affine sets\<close>
       
   580 text\<open>Including the famous stereoscopic projection of the 3-D sphere to the complex plane\<close>
       
   581 
       
   582 text\<open>The special case with centre 0 and radius 1\<close>
       
   583 lemma homeomorphic_punctured_affine_sphere_affine_01:
       
   584   assumes "b \<in> sphere 0 1" "affine T" "0 \<in> T" "b \<in> T" "affine p"
       
   585       and affT: "aff_dim T = aff_dim p + 1"
       
   586     shows "(sphere 0 1 \<inter> T) - {b} homeomorphic p"
       
   587 proof -
       
   588   have [simp]: "norm b = 1" "b\<bullet>b = 1"
       
   589     using assms by (auto simp: norm_eq_1)
       
   590   have [simp]: "T \<inter> {v. b\<bullet>v = 0} \<noteq> {}"
       
   591     using \<open>0 \<in> T\<close> by auto
       
   592   have [simp]: "\<not> T \<subseteq> {v. b\<bullet>v = 0}"
       
   593     using \<open>norm b = 1\<close> \<open>b \<in> T\<close> by auto
       
   594   define f where "f \<equiv> \<lambda>x. 2 *\<^sub>R b + (2 / (1 - b\<bullet>x)) *\<^sub>R (x - b)"
       
   595   define g where "g \<equiv> \<lambda>y. b + (4 / (norm y ^ 2 + 4)) *\<^sub>R (y - 2 *\<^sub>R b)"
       
   596   have [simp]: "\<And>x. \<lbrakk>x \<in> T; b\<bullet>x = 0\<rbrakk> \<Longrightarrow> f (g x) = x"
       
   597     unfolding f_def g_def by (simp add: algebra_simps divide_simps add_nonneg_eq_0_iff)
       
   598   have no: "\<And>x. \<lbrakk>norm x = 1; b\<bullet>x \<noteq> 1\<rbrakk> \<Longrightarrow> (norm (f x))\<^sup>2 = 4 * (1 + b\<bullet>x) / (1 - b\<bullet>x)"
       
   599     apply (simp add: dot_square_norm [symmetric])
       
   600     apply (simp add: f_def vector_add_divide_simps divide_simps norm_eq_1)
       
   601     apply (simp add: algebra_simps inner_commute)
       
   602     done
       
   603   have [simp]: "\<And>u::real. 8 + u * (u * 8) = u * 16 \<longleftrightarrow> u=1"
       
   604     by algebra
       
   605   have [simp]: "\<And>x. \<lbrakk>norm x = 1; b \<bullet> x \<noteq> 1\<rbrakk> \<Longrightarrow> g (f x) = x"
       
   606     unfolding g_def no by (auto simp: f_def divide_simps)
       
   607   have [simp]: "\<And>x. \<lbrakk>x \<in> T; b \<bullet> x = 0\<rbrakk> \<Longrightarrow> norm (g x) = 1"
       
   608     unfolding g_def
       
   609     apply (rule power2_eq_imp_eq)
       
   610     apply (simp_all add: dot_square_norm [symmetric] divide_simps vector_add_divide_simps)
       
   611     apply (simp add: algebra_simps inner_commute)
       
   612     done
       
   613   have [simp]: "\<And>x. \<lbrakk>x \<in> T; b \<bullet> x = 0\<rbrakk> \<Longrightarrow> b \<bullet> g x \<noteq> 1"
       
   614     unfolding g_def
       
   615     apply (simp_all add: dot_square_norm [symmetric] divide_simps vector_add_divide_simps add_nonneg_eq_0_iff)
       
   616     apply (auto simp: algebra_simps)
       
   617     done
       
   618   have "subspace T"
       
   619     by (simp add: assms subspace_affine)
       
   620   have [simp]: "\<And>x. \<lbrakk>x \<in> T; b \<bullet> x = 0\<rbrakk> \<Longrightarrow> g x \<in> T"
       
   621     unfolding g_def
       
   622     by (blast intro: \<open>subspace T\<close> \<open>b \<in> T\<close> subspace_add subspace_mul subspace_diff)
       
   623   have "f ` {x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<subseteq> {x. b\<bullet>x = 0}"
       
   624     unfolding f_def using \<open>norm b = 1\<close> norm_eq_1
       
   625     by (force simp: field_simps inner_add_right inner_diff_right)
       
   626   moreover have "f ` T \<subseteq> T"
       
   627     unfolding f_def using assms
       
   628     apply (auto simp: field_simps inner_add_right inner_diff_right)
       
   629     by (metis add_0 diff_zero mem_affine_3_minus)
       
   630   moreover have "{x. b\<bullet>x = 0} \<inter> T \<subseteq> f ` ({x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T)"
       
   631     apply clarify
       
   632     apply (rule_tac x = "g x" in image_eqI, auto)
       
   633     done
       
   634   ultimately have imf: "f ` ({x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T) = {x. b\<bullet>x = 0} \<inter> T"
       
   635     by blast
       
   636   have no4: "\<And>y. b\<bullet>y = 0 \<Longrightarrow> norm ((y\<bullet>y + 4) *\<^sub>R b + 4 *\<^sub>R (y - 2 *\<^sub>R b)) = y\<bullet>y + 4"
       
   637     apply (rule power2_eq_imp_eq)
       
   638     apply (simp_all add: dot_square_norm [symmetric])
       
   639     apply (auto simp: power2_eq_square algebra_simps inner_commute)
       
   640     done
       
   641   have [simp]: "\<And>x. \<lbrakk>norm x = 1; b \<bullet> x \<noteq> 1\<rbrakk> \<Longrightarrow> b \<bullet> f x = 0"
       
   642     by (simp add: f_def algebra_simps divide_simps)
       
   643   have [simp]: "\<And>x. \<lbrakk>x \<in> T; norm x = 1; b \<bullet> x \<noteq> 1\<rbrakk> \<Longrightarrow> f x \<in> T"
       
   644     unfolding f_def
       
   645     by (blast intro: \<open>subspace T\<close> \<open>b \<in> T\<close> subspace_add subspace_mul subspace_diff)
       
   646   have "g ` {x. b\<bullet>x = 0} \<subseteq> {x. norm x = 1 \<and> b\<bullet>x \<noteq> 1}"
       
   647     unfolding g_def
       
   648     apply (clarsimp simp: no4 vector_add_divide_simps divide_simps add_nonneg_eq_0_iff dot_square_norm [symmetric])
       
   649     apply (auto simp: algebra_simps)
       
   650     done
       
   651   moreover have "g ` T \<subseteq> T"
       
   652     unfolding g_def
       
   653     by (blast intro: \<open>subspace T\<close> \<open>b \<in> T\<close> subspace_add subspace_mul subspace_diff)
       
   654   moreover have "{x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T \<subseteq> g ` ({x. b\<bullet>x = 0} \<inter> T)"
       
   655     apply clarify
       
   656     apply (rule_tac x = "f x" in image_eqI, auto)
       
   657     done
       
   658   ultimately have img: "g ` ({x. b\<bullet>x = 0} \<inter> T) = {x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T"
       
   659     by blast
       
   660   have aff: "affine ({x. b\<bullet>x = 0} \<inter> T)"
       
   661     by (blast intro: affine_hyperplane assms)
       
   662   have contf: "continuous_on ({x. norm x = 1 \<and> b\<bullet>x \<noteq> 1} \<inter> T) f"
       
   663     unfolding f_def by (rule continuous_intros | force)+
       
   664   have contg: "continuous_on ({x. b\<bullet>x = 0} \<inter> T) g"
       
   665     unfolding g_def by (rule continuous_intros | force simp: add_nonneg_eq_0_iff)+
       
   666   have "(sphere 0 1 \<inter> T) - {b} = {x. norm x = 1 \<and> (b\<bullet>x \<noteq> 1)} \<inter> T"
       
   667     using  \<open>norm b = 1\<close> by (auto simp: norm_eq_1) (metis vector_eq  \<open>b\<bullet>b = 1\<close>)
       
   668   also have "... homeomorphic {x. b\<bullet>x = 0} \<inter> T"
       
   669     by (rule homeomorphicI [OF imf img contf contg]) auto
       
   670   also have "... homeomorphic p"
       
   671     apply (rule homeomorphic_affine_sets [OF aff \<open>affine p\<close>])
       
   672     apply (simp add: Int_commute aff_dim_affine_Int_hyperplane [OF \<open>affine T\<close>] affT)
       
   673     done
       
   674   finally show ?thesis .
       
   675 qed
       
   676 
       
   677 theorem homeomorphic_punctured_affine_sphere_affine:
       
   678   fixes a :: "'a :: euclidean_space"
       
   679   assumes "0 < r" "b \<in> sphere a r" "affine T" "a \<in> T" "b \<in> T" "affine p"
       
   680       and aff: "aff_dim T = aff_dim p + 1"
       
   681     shows "((sphere a r \<inter> T) - {b}) homeomorphic p"
       
   682 proof -
       
   683   have "a \<noteq> b" using assms by auto
       
   684   then have inj: "inj (\<lambda>x::'a. x /\<^sub>R norm (a - b))"
       
   685     by (simp add: inj_on_def)
       
   686   have "((sphere a r \<inter> T) - {b}) homeomorphic
       
   687         op+ (-a) ` ((sphere a r \<inter> T) - {b})"
       
   688     by (rule homeomorphic_translation)
       
   689   also have "... homeomorphic op *\<^sub>R (inverse r) ` op + (- a) ` (sphere a r \<inter> T - {b})"
       
   690     by (metis \<open>0 < r\<close> homeomorphic_scaling inverse_inverse_eq inverse_zero less_irrefl)
       
   691   also have "... = sphere 0 1 \<inter> (op *\<^sub>R (inverse r) ` op + (- a) ` T) - {(b - a) /\<^sub>R r}"
       
   692     using assms by (auto simp: dist_norm norm_minus_commute divide_simps)
       
   693   also have "... homeomorphic p"
       
   694     apply (rule homeomorphic_punctured_affine_sphere_affine_01)
       
   695     using assms
       
   696     apply (auto simp: dist_norm norm_minus_commute affine_scaling affine_translation [symmetric] aff_dim_translation_eq inj)
       
   697     done
       
   698   finally show ?thesis .
       
   699 qed
       
   700 
       
   701 proposition homeomorphic_punctured_sphere_affine_gen:
       
   702   fixes a :: "'a :: euclidean_space"
       
   703   assumes "convex S" "bounded S" and a: "a \<in> rel_frontier S"
       
   704       and "affine T" and affS: "aff_dim S = aff_dim T + 1"
       
   705     shows "rel_frontier S - {a} homeomorphic T"
       
   706 proof -
       
   707   have "S \<noteq> {}" using assms by auto
       
   708   obtain U :: "'a set" where "affine U" and affdS: "aff_dim U = aff_dim S"
       
   709     using choose_affine_subset [OF affine_UNIV aff_dim_geq]
       
   710     by (meson aff_dim_affine_hull affine_affine_hull)
       
   711   have "convex U"
       
   712     by (simp add: affine_imp_convex \<open>affine U\<close>)
       
   713   have "U \<noteq> {}"
       
   714     by (metis \<open>S \<noteq> {}\<close> \<open>aff_dim U = aff_dim S\<close> aff_dim_empty)
       
   715   then obtain z where "z \<in> U"
       
   716     by auto
       
   717   then have bne: "ball z 1 \<inter> U \<noteq> {}" by force
       
   718   have [simp]: "aff_dim(ball z 1 \<inter> U) = aff_dim U"
       
   719     using aff_dim_convex_Int_open [OF \<open>convex U\<close> open_ball] bne
       
   720     by (fastforce simp add: Int_commute)
       
   721   have "rel_frontier S homeomorphic rel_frontier (ball z 1 \<inter> U)"
       
   722     apply (rule homeomorphic_rel_frontiers_convex_bounded_sets)
       
   723     apply (auto simp: \<open>affine U\<close> affine_imp_convex convex_Int affdS assms)
       
   724     done
       
   725   also have "... = sphere z 1 \<inter> U"
       
   726     using convex_affine_rel_frontier_Int [of "ball z 1" U]
       
   727     by (simp add: \<open>affine U\<close> bne)
       
   728   finally obtain h k where him: "h ` rel_frontier S = sphere z 1 \<inter> U"
       
   729                     and kim: "k ` (sphere z 1 \<inter> U) = rel_frontier S"
       
   730                     and hcon: "continuous_on (rel_frontier S) h"
       
   731                     and kcon: "continuous_on (sphere z 1 \<inter> U) k"
       
   732                     and kh:  "\<And>x. x \<in> rel_frontier S \<Longrightarrow> k(h(x)) = x"
       
   733                     and hk:  "\<And>y. y \<in> sphere z 1 \<inter> U \<Longrightarrow> h(k(y)) = y"
       
   734     unfolding homeomorphic_def homeomorphism_def by auto
       
   735   have "rel_frontier S - {a} homeomorphic (sphere z 1 \<inter> U) - {h a}"
       
   736   proof (rule homeomorphicI [where f=h and g=k])
       
   737     show h: "h ` (rel_frontier S - {a}) = sphere z 1 \<inter> U - {h a}"
       
   738       using him a kh by auto metis
       
   739     show "k ` (sphere z 1 \<inter> U - {h a}) = rel_frontier S - {a}"
       
   740       by (force simp: h [symmetric] image_comp o_def kh)
       
   741   qed (auto intro: continuous_on_subset hcon kcon simp: kh hk)
       
   742   also have "... homeomorphic T"
       
   743     apply (rule homeomorphic_punctured_affine_sphere_affine)
       
   744     using a him
       
   745     by (auto simp: affS affdS \<open>affine T\<close>  \<open>affine U\<close> \<open>z \<in> U\<close>)
       
   746   finally show ?thesis .
       
   747 qed
       
   748 
       
   749 
       
   750 lemma homeomorphic_punctured_sphere_affine:
       
   751   fixes a :: "'a :: euclidean_space"
       
   752   assumes "0 < r" and b: "b \<in> sphere a r"
       
   753       and "affine T" and affS: "aff_dim T + 1 = DIM('a)"
       
   754     shows "(sphere a r - {b}) homeomorphic T"
       
   755 using homeomorphic_punctured_sphere_affine_gen [of "cball a r" b T]
       
   756   assms aff_dim_cball by force
       
   757 
       
   758 corollary homeomorphic_punctured_sphere_hyperplane:
       
   759   fixes a :: "'a :: euclidean_space"
       
   760   assumes "0 < r" and b: "b \<in> sphere a r"
       
   761       and "c \<noteq> 0"
       
   762     shows "(sphere a r - {b}) homeomorphic {x::'a. c \<bullet> x = d}"
       
   763 apply (rule homeomorphic_punctured_sphere_affine)
       
   764 using assms
       
   765 apply (auto simp: affine_hyperplane of_nat_diff)
       
   766 done
       
   767 
       
   768 
       
   769 text\<open> When dealing with AR, ANR and ANR later, it's useful to know that every set
       
   770   is homeomorphic to a closed subset of a convex set, and
       
   771   if the set is locally compact we can take the convex set to be the universe.\<close>
       
   772 
       
   773 proposition homeomorphic_closedin_convex:
       
   774   fixes S :: "'m::euclidean_space set"
       
   775   assumes "aff_dim S < DIM('n)"
       
   776   obtains U and T :: "'n::euclidean_space set"
       
   777      where "convex U" "U \<noteq> {}" "closedin (subtopology euclidean U) T"
       
   778            "S homeomorphic T"
       
   779 proof (cases "S = {}")
       
   780   case True then show ?thesis
       
   781     by (rule_tac U=UNIV and T="{}" in that) auto
       
   782 next
       
   783   case False
       
   784   then obtain a where "a \<in> S" by auto
       
   785   obtain i::'n where i: "i \<in> Basis" "i \<noteq> 0"
       
   786     using SOME_Basis Basis_zero by force
       
   787   have "0 \<in> affine hull (op + (- a) ` S)"
       
   788     by (simp add: \<open>a \<in> S\<close> hull_inc)
       
   789   then have "dim (op + (- a) ` S) = aff_dim (op + (- a) ` S)"
       
   790     by (simp add: aff_dim_zero)
       
   791   also have "... < DIM('n)"
       
   792     by (simp add: aff_dim_translation_eq assms)
       
   793   finally have dd: "dim (op + (- a) ` S) < DIM('n)"
       
   794     by linarith
       
   795   obtain T where "subspace T" and Tsub: "T \<subseteq> {x. i \<bullet> x = 0}"
       
   796              and dimT: "dim T = dim (op + (- a) ` S)"
       
   797     apply (rule choose_subspace_of_subspace [of "dim (op + (- a) ` S)" "{x::'n. i \<bullet> x = 0}"])
       
   798      apply (simp add: dim_hyperplane [OF \<open>i \<noteq> 0\<close>])
       
   799      apply (metis DIM_positive Suc_pred dd not_le not_less_eq_eq)
       
   800     apply (metis span_eq subspace_hyperplane)
       
   801     done
       
   802   have "subspace (span (op + (- a) ` S))"
       
   803     using subspace_span by blast
       
   804   then obtain h k where "linear h" "linear k"
       
   805                and heq: "h ` span (op + (- a) ` S) = T"
       
   806                and keq:"k ` T = span (op + (- a) ` S)"
       
   807                and hinv [simp]:  "\<And>x. x \<in> span (op + (- a) ` S) \<Longrightarrow> k(h x) = x"
       
   808                and kinv [simp]:  "\<And>x. x \<in> T \<Longrightarrow> h(k x) = x"
       
   809     apply (rule isometries_subspaces [OF _ \<open>subspace T\<close>])
       
   810     apply (auto simp: dimT)
       
   811     done
       
   812   have hcont: "continuous_on A h" and kcont: "continuous_on B k" for A B
       
   813     using \<open>linear h\<close> \<open>linear k\<close> linear_continuous_on linear_conv_bounded_linear by blast+
       
   814   have ihhhh[simp]: "\<And>x. x \<in> S \<Longrightarrow> i \<bullet> h (x - a) = 0"
       
   815     using Tsub [THEN subsetD] heq span_inc by fastforce
       
   816   have "sphere 0 1 - {i} homeomorphic {x. i \<bullet> x = 0}"
       
   817     apply (rule homeomorphic_punctured_sphere_affine)
       
   818     using i
       
   819     apply (auto simp: affine_hyperplane)
       
   820     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)
       
   821   then obtain f g where fg: "homeomorphism (sphere 0 1 - {i}) {x. i \<bullet> x = 0} f g"
       
   822     by (force simp: homeomorphic_def)
       
   823   have "h ` op + (- a) ` S \<subseteq> T"
       
   824     using heq span_clauses(1) span_linear_image by blast
       
   825   then have "g ` h ` op + (- a) ` S \<subseteq> g ` {x. i \<bullet> x = 0}"
       
   826     using Tsub by (simp add: image_mono)
       
   827   also have "... \<subseteq> sphere 0 1 - {i}"
       
   828     by (simp add: fg [unfolded homeomorphism_def])
       
   829   finally have gh_sub_sph: "(g \<circ> h) ` op + (- a) ` S \<subseteq> sphere 0 1 - {i}"
       
   830     by (metis image_comp)
       
   831   then have gh_sub_cb: "(g \<circ> h) ` op + (- a) ` S \<subseteq> cball 0 1"
       
   832     by (metis Diff_subset order_trans sphere_cball)
       
   833   have [simp]: "\<And>u. u \<in> S \<Longrightarrow> norm (g (h (u - a))) = 1"
       
   834     using gh_sub_sph [THEN subsetD] by (auto simp: o_def)
       
   835   have ghcont: "continuous_on (op + (- a) ` S) (\<lambda>x. g (h x))"
       
   836     apply (rule continuous_on_compose2 [OF homeomorphism_cont2 [OF fg] hcont], force)
       
   837     done
       
   838   have kfcont: "continuous_on ((g \<circ> h \<circ> op + (- a)) ` S) (\<lambda>x. k (f x))"
       
   839     apply (rule continuous_on_compose2 [OF kcont])
       
   840     using homeomorphism_cont1 [OF fg] gh_sub_sph apply (force intro: continuous_on_subset, blast)
       
   841     done
       
   842   have "S homeomorphic op + (- a) ` S"
       
   843     by (simp add: homeomorphic_translation)
       
   844   also have Shom: "\<dots> homeomorphic (g \<circ> h) ` op + (- a) ` S"
       
   845     apply (simp add: homeomorphic_def homeomorphism_def)
       
   846     apply (rule_tac x="g \<circ> h" in exI)
       
   847     apply (rule_tac x="k \<circ> f" in exI)
       
   848     apply (auto simp: ghcont kfcont span_clauses(1) homeomorphism_apply2 [OF fg] image_comp)
       
   849     apply (force simp: o_def homeomorphism_apply2 [OF fg] span_clauses(1))
       
   850     done
       
   851   finally have Shom: "S homeomorphic (g \<circ> h) ` op + (- a) ` S" .
       
   852   show ?thesis
       
   853     apply (rule_tac U = "ball 0 1 \<union> image (g o h) (op + (- a) ` S)"
       
   854                 and T = "image (g o h) (op + (- a) ` S)"
       
   855                     in that)
       
   856     apply (rule convex_intermediate_ball [of 0 1], force)
       
   857     using gh_sub_cb apply force
       
   858     apply force
       
   859     apply (simp add: closedin_closed)
       
   860     apply (rule_tac x="sphere 0 1" in exI)
       
   861     apply (auto simp: Shom)
       
   862     done
       
   863 qed
       
   864 
       
   865 subsection\<open>Locally compact sets in an open set\<close>
       
   866 
       
   867 text\<open> Locally compact sets are closed in an open set and are homeomorphic
       
   868   to an absolutely closed set if we have one more dimension to play with.\<close>
       
   869 
       
   870 lemma locally_compact_open_Int_closure:
       
   871   fixes S :: "'a :: metric_space set"
       
   872   assumes "locally compact S"
       
   873   obtains T where "open T" "S = T \<inter> closure S"
       
   874 proof -
       
   875   have "\<forall>x\<in>S. \<exists>T v u. u = S \<inter> T \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> S \<and> open T \<and> compact v"
       
   876     by (metis assms locally_compact openin_open)
       
   877   then obtain t v where
       
   878         tv: "\<And>x. x \<in> S
       
   879              \<Longrightarrow> v x \<subseteq> S \<and> open (t x) \<and> compact (v x) \<and> (\<exists>u. x \<in> u \<and> u \<subseteq> v x \<and> u = S \<inter> t x)"
       
   880     by metis
       
   881   then have o: "open (UNION S t)"
       
   882     by blast
       
   883   have "S = \<Union> (v ` S)"
       
   884     using tv by blast
       
   885   also have "... = UNION S t \<inter> closure S"
       
   886   proof
       
   887     show "UNION S v \<subseteq> UNION S t \<inter> closure S"
       
   888       apply safe
       
   889        apply (metis Int_iff subsetD UN_iff tv)
       
   890       apply (simp add: closure_def rev_subsetD tv)
       
   891       done
       
   892     have "t x \<inter> closure S \<subseteq> v x" if "x \<in> S" for x
       
   893     proof -
       
   894       have "t x \<inter> closure S \<subseteq> closure (t x \<inter> S)"
       
   895         by (simp add: open_Int_closure_subset that tv)
       
   896       also have "... \<subseteq> v x"
       
   897         by (metis Int_commute closure_minimal compact_imp_closed that tv)
       
   898       finally show ?thesis .
       
   899     qed
       
   900     then show "UNION S t \<inter> closure S \<subseteq> UNION S v"
       
   901       by blast
       
   902   qed
       
   903   finally have e: "S = UNION S t \<inter> closure S" .
       
   904   show ?thesis
       
   905     by (rule that [OF o e])
       
   906 qed
       
   907 
       
   908 
       
   909 lemma locally_compact_closedin_open:
       
   910     fixes S :: "'a :: metric_space set"
       
   911     assumes "locally compact S"
       
   912     obtains T where "open T" "closedin (subtopology euclidean T) S"
       
   913   by (metis locally_compact_open_Int_closure [OF assms] closed_closure closedin_closed_Int)
       
   914 
       
   915 
       
   916 lemma locally_compact_homeomorphism_projection_closed:
       
   917   assumes "locally compact S"
       
   918   obtains T and f :: "'a \<Rightarrow> 'a :: euclidean_space \<times> 'b :: euclidean_space"
       
   919     where "closed T" "homeomorphism S T f fst"
       
   920 proof (cases "closed S")
       
   921   case True
       
   922     then show ?thesis
       
   923       apply (rule_tac T = "S \<times> {0}" and f = "\<lambda>x. (x, 0)" in that)
       
   924       apply (auto simp: closed_Times homeomorphism_def continuous_intros)
       
   925       done
       
   926 next
       
   927   case False
       
   928     obtain U where "open U" and US: "U \<inter> closure S = S"
       
   929       by (metis locally_compact_open_Int_closure [OF assms])
       
   930     with False have Ucomp: "-U \<noteq> {}"
       
   931       using closure_eq by auto
       
   932     have [simp]: "closure (- U) = -U"
       
   933       by (simp add: \<open>open U\<close> closed_Compl)
       
   934     define f :: "'a \<Rightarrow> 'a \<times> 'b" where "f \<equiv> \<lambda>x. (x, One /\<^sub>R setdist {x} (- U))"
       
   935     have "continuous_on U (\<lambda>x. (x, One /\<^sub>R setdist {x} (- U)))"
       
   936       by (auto simp: Ucomp continuous_intros continuous_on_setdist)
       
   937     then have homU: "homeomorphism U (f`U) f fst"
       
   938       by (auto simp: f_def homeomorphism_def image_iff continuous_intros)
       
   939     have cloS: "closedin (subtopology euclidean U) S"
       
   940       by (metis US closed_closure closedin_closed_Int)
       
   941     have cont: "isCont ((\<lambda>x. setdist {x} (- U)) o fst) z" for z :: "'a \<times> 'b"
       
   942       by (rule isCont_o continuous_intros continuous_at_setdist)+
       
   943     have setdist1D: "setdist {a} (- U) *\<^sub>R b = One \<Longrightarrow> setdist {a} (- U) \<noteq> 0" for a::'a and b::'b
       
   944       by force
       
   945     have *: "r *\<^sub>R b = One \<Longrightarrow> b = (1 / r) *\<^sub>R One" for r and b::'b
       
   946       by (metis One_non_0 nonzero_divide_eq_eq real_vector.scale_eq_0_iff real_vector.scale_scale scaleR_one)
       
   947     have "f ` U = {z. (setdist {fst z} (- U) *\<^sub>R snd z) \<in> {One}}"
       
   948       apply (auto simp: f_def field_simps Ucomp)
       
   949       apply (rule_tac x=a in image_eqI)
       
   950       apply (auto simp: * dest: setdist1D)
       
   951       done
       
   952     then have clfU: "closed (f ` U)"
       
   953       apply (rule ssubst)
       
   954       apply (rule continuous_closed_preimage_univ)
       
   955       apply (auto intro: continuous_intros cont [unfolded o_def])
       
   956       done
       
   957     have "closed (f ` S)"
       
   958        apply (rule closedin_closed_trans [OF _ clfU])
       
   959        apply (rule homeomorphism_imp_closed_map [OF homU cloS])
       
   960        done
       
   961     then show ?thesis
       
   962       apply (rule that)
       
   963       apply (rule homeomorphism_of_subsets [OF homU])
       
   964       using US apply auto
       
   965       done
       
   966 qed
       
   967 
       
   968 lemma locally_compact_closed_Int_open:
       
   969   fixes S :: "'a :: euclidean_space set"
       
   970   shows
       
   971     "locally compact S \<longleftrightarrow> (\<exists>U u. closed U \<and> open u \<and> S = U \<inter> u)"
       
   972 by (metis closed_closure closed_imp_locally_compact inf_commute locally_compact_Int locally_compact_open_Int_closure open_imp_locally_compact)
       
   973 
       
   974 
       
   975 proposition locally_compact_homeomorphic_closed:
       
   976   fixes S :: "'a::euclidean_space set"
       
   977   assumes "locally compact S" and dimlt: "DIM('a) < DIM('b)"
       
   978   obtains T :: "'b::euclidean_space set" where "closed T" "S homeomorphic T"
       
   979 proof -
       
   980   obtain U:: "('a*real)set" and h
       
   981     where "closed U" and homU: "homeomorphism S U h fst"
       
   982     using locally_compact_homeomorphism_projection_closed assms by metis
       
   983   let ?BP = "Basis :: ('a*real) set"
       
   984   have "DIM('a * real) \<le> DIM('b)"
       
   985     by (simp add: Suc_leI dimlt)
       
   986   then obtain basf :: "'a*real \<Rightarrow> 'b" where surbf: "basf ` ?BP \<subseteq> Basis" and injbf: "inj_on basf Basis"
       
   987     by (metis finite_Basis card_le_inj)
       
   988   define basg:: "'b \<Rightarrow> 'a * real" where
       
   989     "basg \<equiv> \<lambda>i. inv_into Basis basf i"
       
   990   have bgf[simp]: "basg (basf i) = i" if "i \<in> Basis" for i
       
   991     using inv_into_f_f injbf that by (force simp: basg_def)
       
   992   define f :: "'a*real \<Rightarrow> 'b" where "f \<equiv> \<lambda>u. \<Sum>j\<in>Basis. (u \<bullet> basg j) *\<^sub>R j"
       
   993   have "linear f"
       
   994     unfolding f_def
       
   995     apply (intro linear_compose_setsum linearI ballI)
       
   996     apply (auto simp: algebra_simps)
       
   997     done
       
   998   define g :: "'b \<Rightarrow> 'a*real" where "g \<equiv> \<lambda>z. (\<Sum>i\<in>Basis. (z \<bullet> basf i) *\<^sub>R i)"
       
   999   have "linear g"
       
  1000     unfolding g_def
       
  1001     apply (intro linear_compose_setsum linearI ballI)
       
  1002     apply (auto simp: algebra_simps)
       
  1003     done
       
  1004   have *: "(\<Sum>a \<in> Basis. a \<bullet> basf b * (x \<bullet> basg a)) = x \<bullet> b" if "b \<in> Basis" for x b
       
  1005     using surbf that by auto
       
  1006   have gf[simp]: "g (f x) = x" for x
       
  1007     apply (rule euclidean_eqI)
       
  1008     apply (simp add: f_def g_def inner_setsum_left scaleR_setsum_left algebra_simps)
       
  1009     apply (simp add: Groups_Big.setsum_right_distrib [symmetric] *)
       
  1010     done
       
  1011   then have "inj f" by (metis injI)
       
  1012   have gfU: "g ` f ` U = U"
       
  1013     by (rule set_eqI) (auto simp: image_def)
       
  1014   have "S homeomorphic U"
       
  1015     using homU homeomorphic_def by blast
       
  1016   also have "... homeomorphic f ` U"
       
  1017     apply (rule homeomorphicI [OF refl gfU])
       
  1018        apply (meson \<open>inj f\<close> \<open>linear f\<close> homeomorphism_cont2 linear_homeomorphism_image)
       
  1019       using \<open>linear g\<close> linear_continuous_on linear_conv_bounded_linear apply blast
       
  1020      apply auto
       
  1021      done
       
  1022   finally show ?thesis
       
  1023     apply (rule_tac T = "f ` U" in that)
       
  1024     apply (rule closed_injective_linear_image [OF \<open>closed U\<close> \<open>linear f\<close> \<open>inj f\<close>], assumption)
       
  1025     done
       
  1026 qed
       
  1027 
       
  1028 
       
  1029 lemma homeomorphic_convex_compact_lemma:
       
  1030   fixes s :: "'a::euclidean_space set"
       
  1031   assumes "convex s"
       
  1032     and "compact s"
       
  1033     and "cball 0 1 \<subseteq> s"
       
  1034   shows "s homeomorphic (cball (0::'a) 1)"
       
  1035 proof (rule starlike_compact_projective_special[OF assms(2-3)])
       
  1036   fix x u
       
  1037   assume "x \<in> s" and "0 \<le> u" and "u < (1::real)"
       
  1038   have "open (ball (u *\<^sub>R x) (1 - u))"
       
  1039     by (rule open_ball)
       
  1040   moreover have "u *\<^sub>R x \<in> ball (u *\<^sub>R x) (1 - u)"
       
  1041     unfolding centre_in_ball using \<open>u < 1\<close> by simp
       
  1042   moreover have "ball (u *\<^sub>R x) (1 - u) \<subseteq> s"
       
  1043   proof
       
  1044     fix y
       
  1045     assume "y \<in> ball (u *\<^sub>R x) (1 - u)"
       
  1046     then have "dist (u *\<^sub>R x) y < 1 - u"
       
  1047       unfolding mem_ball .
       
  1048     with \<open>u < 1\<close> have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \<in> cball 0 1"
       
  1049       by (simp add: dist_norm inverse_eq_divide norm_minus_commute)
       
  1050     with assms(3) have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \<in> s" ..
       
  1051     with assms(1) have "(1 - u) *\<^sub>R ((y - u *\<^sub>R x) /\<^sub>R (1 - u)) + u *\<^sub>R x \<in> s"
       
  1052       using \<open>x \<in> s\<close> \<open>0 \<le> u\<close> \<open>u < 1\<close> [THEN less_imp_le] by (rule convexD_alt)
       
  1053     then show "y \<in> s" using \<open>u < 1\<close>
       
  1054       by simp
       
  1055   qed
       
  1056   ultimately have "u *\<^sub>R x \<in> interior s" ..
       
  1057   then show "u *\<^sub>R x \<in> s - frontier s"
       
  1058     using frontier_def and interior_subset by auto
       
  1059 qed
       
  1060 
       
  1061 proposition homeomorphic_convex_compact_cball:
       
  1062   fixes e :: real
       
  1063     and s :: "'a::euclidean_space set"
       
  1064   assumes "convex s"
       
  1065     and "compact s"
       
  1066     and "interior s \<noteq> {}"
       
  1067     and "e > 0"
       
  1068   shows "s homeomorphic (cball (b::'a) e)"
       
  1069 proof -
       
  1070   obtain a where "a \<in> interior s"
       
  1071     using assms(3) by auto
       
  1072   then obtain d where "d > 0" and d: "cball a d \<subseteq> s"
       
  1073     unfolding mem_interior_cball by auto
       
  1074   let ?d = "inverse d" and ?n = "0::'a"
       
  1075   have "cball ?n 1 \<subseteq> (\<lambda>x. inverse d *\<^sub>R (x - a)) ` s"
       
  1076     apply rule
       
  1077     apply (rule_tac x="d *\<^sub>R x + a" in image_eqI)
       
  1078     defer
       
  1079     apply (rule d[unfolded subset_eq, rule_format])
       
  1080     using \<open>d > 0\<close>
       
  1081     unfolding mem_cball dist_norm
       
  1082     apply (auto simp add: mult_right_le_one_le)
       
  1083     done
       
  1084   then have "(\<lambda>x. inverse d *\<^sub>R (x - a)) ` s homeomorphic cball ?n 1"
       
  1085     using homeomorphic_convex_compact_lemma[of "(\<lambda>x. ?d *\<^sub>R -a + ?d *\<^sub>R x) ` s",
       
  1086       OF convex_affinity compact_affinity]
       
  1087     using assms(1,2)
       
  1088     by (auto simp add: scaleR_right_diff_distrib)
       
  1089   then show ?thesis
       
  1090     apply (rule_tac homeomorphic_trans[OF _ homeomorphic_balls(2)[of 1 _ ?n]])
       
  1091     apply (rule homeomorphic_trans[OF homeomorphic_affinity[of "?d" s "?d *\<^sub>R -a"]])
       
  1092     using \<open>d>0\<close> \<open>e>0\<close>
       
  1093     apply (auto simp add: scaleR_right_diff_distrib)
       
  1094     done
       
  1095 qed
       
  1096 
       
  1097 corollary homeomorphic_convex_compact:
       
  1098   fixes s :: "'a::euclidean_space set"
       
  1099     and t :: "'a set"
       
  1100   assumes "convex s" "compact s" "interior s \<noteq> {}"
       
  1101     and "convex t" "compact t" "interior t \<noteq> {}"
       
  1102   shows "s homeomorphic t"
       
  1103   using assms
       
  1104   by (meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym)
       
  1105 
       
  1106 end