src/HOL/Analysis/Retracts.thy
changeset 70642 de9c4ed2d5df
child 70643 93a7a85de312
equal deleted inserted replaced
70641:0e2a065d6f31 70642:de9c4ed2d5df
       
     1 section \<open>Absolute Retracts, Absolute Neighbourhood Retracts and Euclidean Neighbourhood Retracts\<close>
       
     2 
       
     3 theory Retracts
       
     4   imports Brouwer_Fixpoint Ordered_Euclidean_Space
       
     5 begin
       
     6 
       
     7 
       
     8 text \<open>Absolute retracts (AR), absolute neighbourhood retracts (ANR) and also Euclidean neighbourhood
       
     9 retracts (ENR). We define AR and ANR by specializing the standard definitions for a set to embedding
       
    10 in spaces of higher dimension.
       
    11 
       
    12 John Harrison writes: "This turns out to be sufficient (since any set in \<open>\<real>\<^sup>n\<close> can be
       
    13 embedded as a closed subset of a convex subset of \<open>\<real>\<^sup>n\<^sup>+\<^sup>1\<close>) to derive the usual
       
    14 definitions, but we need to split them into two implications because of the lack of type
       
    15 quantifiers. Then ENR turns out to be equivalent to ANR plus local compactness."\<close>
       
    16 
       
    17 definition\<^marker>\<open>tag important\<close> AR :: "'a::topological_space set \<Rightarrow> bool" where
       
    18 "AR S \<equiv> \<forall>U. \<forall>S'::('a * real) set.
       
    19   S homeomorphic S' \<and> closedin (top_of_set U) S' \<longrightarrow> S' retract_of U"
       
    20 
       
    21 definition\<^marker>\<open>tag important\<close> ANR :: "'a::topological_space set \<Rightarrow> bool" where
       
    22 "ANR S \<equiv> \<forall>U. \<forall>S'::('a * real) set.
       
    23   S homeomorphic S' \<and> closedin (top_of_set U) S'
       
    24   \<longrightarrow> (\<exists>T. openin (top_of_set U) T \<and> S' retract_of T)"
       
    25 
       
    26 definition\<^marker>\<open>tag important\<close> ENR :: "'a::topological_space set \<Rightarrow> bool" where
       
    27 "ENR S \<equiv> \<exists>U. open U \<and> S retract_of U"
       
    28 
       
    29 text \<open>First, show that we do indeed get the "usual" properties of ARs and ANRs.\<close>
       
    30 
       
    31 lemma AR_imp_absolute_extensor:
       
    32   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
    33   assumes "AR S" and contf: "continuous_on T f" and "f ` T \<subseteq> S"
       
    34       and cloUT: "closedin (top_of_set U) T"
       
    35   obtains g where "continuous_on U g" "g ` U \<subseteq> S" "\<And>x. x \<in> T \<Longrightarrow> g x = f x"
       
    36 proof -
       
    37   have "aff_dim S < int (DIM('b \<times> real))"
       
    38     using aff_dim_le_DIM [of S] by simp
       
    39   then obtain C and S' :: "('b * real) set"
       
    40           where C: "convex C" "C \<noteq> {}"
       
    41             and cloCS: "closedin (top_of_set C) S'"
       
    42             and hom: "S homeomorphic S'"
       
    43     by (metis that homeomorphic_closedin_convex)
       
    44   then have "S' retract_of C"
       
    45     using \<open>AR S\<close> by (simp add: AR_def)
       
    46   then obtain r where "S' \<subseteq> C" and contr: "continuous_on C r"
       
    47                   and "r ` C \<subseteq> S'" and rid: "\<And>x. x\<in>S' \<Longrightarrow> r x = x"
       
    48     by (auto simp: retraction_def retract_of_def)
       
    49   obtain g h where "homeomorphism S S' g h"
       
    50     using hom by (force simp: homeomorphic_def)
       
    51   then have "continuous_on (f ` T) g"
       
    52     by (meson \<open>f ` T \<subseteq> S\<close> continuous_on_subset homeomorphism_def)
       
    53   then have contgf: "continuous_on T (g \<circ> f)"
       
    54     by (metis continuous_on_compose contf)
       
    55   have gfTC: "(g \<circ> f) ` T \<subseteq> C"
       
    56   proof -
       
    57     have "g ` S = S'"
       
    58       by (metis (no_types) \<open>homeomorphism S S' g h\<close> homeomorphism_def)
       
    59     with \<open>S' \<subseteq> C\<close> \<open>f ` T \<subseteq> S\<close> show ?thesis by force
       
    60   qed
       
    61   obtain f' where f': "continuous_on U f'"  "f' ` U \<subseteq> C"
       
    62                       "\<And>x. x \<in> T \<Longrightarrow> f' x = (g \<circ> f) x"
       
    63     by (metis Dugundji [OF C cloUT contgf gfTC])
       
    64   show ?thesis
       
    65   proof (rule_tac g = "h \<circ> r \<circ> f'" in that)
       
    66     show "continuous_on U (h \<circ> r \<circ> f')"
       
    67       apply (intro continuous_on_compose f')
       
    68        using continuous_on_subset contr f' apply blast
       
    69       by (meson \<open>homeomorphism S S' g h\<close> \<open>r ` C \<subseteq> S'\<close> continuous_on_subset \<open>f' ` U \<subseteq> C\<close> homeomorphism_def image_mono)
       
    70     show "(h \<circ> r \<circ> f') ` U \<subseteq> S"
       
    71       using \<open>homeomorphism S S' g h\<close> \<open>r ` C \<subseteq> S'\<close> \<open>f' ` U \<subseteq> C\<close>
       
    72       by (fastforce simp: homeomorphism_def)
       
    73     show "\<And>x. x \<in> T \<Longrightarrow> (h \<circ> r \<circ> f') x = f x"
       
    74       using \<open>homeomorphism S S' g h\<close> \<open>f ` T \<subseteq> S\<close> f'
       
    75       by (auto simp: rid homeomorphism_def)
       
    76   qed
       
    77 qed
       
    78 
       
    79 lemma AR_imp_absolute_retract:
       
    80   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
    81   assumes "AR S" "S homeomorphic S'"
       
    82       and clo: "closedin (top_of_set U) S'"
       
    83     shows "S' retract_of U"
       
    84 proof -
       
    85   obtain g h where hom: "homeomorphism S S' g h"
       
    86     using assms by (force simp: homeomorphic_def)
       
    87   have h: "continuous_on S' h" " h ` S' \<subseteq> S"
       
    88     using hom homeomorphism_def apply blast
       
    89     apply (metis hom equalityE homeomorphism_def)
       
    90     done
       
    91   obtain h' where h': "continuous_on U h'" "h' ` U \<subseteq> S"
       
    92               and h'h: "\<And>x. x \<in> S' \<Longrightarrow> h' x = h x"
       
    93     by (blast intro: AR_imp_absolute_extensor [OF \<open>AR S\<close> h clo])
       
    94   have [simp]: "S' \<subseteq> U" using clo closedin_limpt by blast
       
    95   show ?thesis
       
    96   proof (simp add: retraction_def retract_of_def, intro exI conjI)
       
    97     show "continuous_on U (g \<circ> h')"
       
    98       apply (intro continuous_on_compose h')
       
    99       apply (meson hom continuous_on_subset h' homeomorphism_cont1)
       
   100       done
       
   101     show "(g \<circ> h') ` U \<subseteq> S'"
       
   102       using h'  by clarsimp (metis hom subsetD homeomorphism_def imageI)
       
   103     show "\<forall>x\<in>S'. (g \<circ> h') x = x"
       
   104       by clarsimp (metis h'h hom homeomorphism_def)
       
   105   qed
       
   106 qed
       
   107 
       
   108 lemma AR_imp_absolute_retract_UNIV:
       
   109   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
   110   assumes "AR S" and hom: "S homeomorphic S'"
       
   111       and clo: "closed S'"
       
   112     shows "S' retract_of UNIV"
       
   113 apply (rule AR_imp_absolute_retract [OF \<open>AR S\<close> hom])
       
   114 using clo closed_closedin by auto
       
   115 
       
   116 lemma absolute_extensor_imp_AR:
       
   117   fixes S :: "'a::euclidean_space set"
       
   118   assumes "\<And>f :: 'a * real \<Rightarrow> 'a.
       
   119            \<And>U T. \<lbrakk>continuous_on T f;  f ` T \<subseteq> S;
       
   120                   closedin (top_of_set U) T\<rbrakk>
       
   121                  \<Longrightarrow> \<exists>g. continuous_on U g \<and> g ` U \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)"
       
   122   shows "AR S"
       
   123 proof (clarsimp simp: AR_def)
       
   124   fix U and T :: "('a * real) set"
       
   125   assume "S homeomorphic T" and clo: "closedin (top_of_set U) T"
       
   126   then obtain g h where hom: "homeomorphism S T g h"
       
   127     by (force simp: homeomorphic_def)
       
   128   have h: "continuous_on T h" " h ` T \<subseteq> S"
       
   129     using hom homeomorphism_def apply blast
       
   130     apply (metis hom equalityE homeomorphism_def)
       
   131     done
       
   132   obtain h' where h': "continuous_on U h'" "h' ` U \<subseteq> S"
       
   133               and h'h: "\<forall>x\<in>T. h' x = h x"
       
   134     using assms [OF h clo] by blast
       
   135   have [simp]: "T \<subseteq> U"
       
   136     using clo closedin_imp_subset by auto
       
   137   show "T retract_of U"
       
   138   proof (simp add: retraction_def retract_of_def, intro exI conjI)
       
   139     show "continuous_on U (g \<circ> h')"
       
   140       apply (intro continuous_on_compose h')
       
   141       apply (meson hom continuous_on_subset h' homeomorphism_cont1)
       
   142       done
       
   143     show "(g \<circ> h') ` U \<subseteq> T"
       
   144       using h'  by clarsimp (metis hom subsetD homeomorphism_def imageI)
       
   145     show "\<forall>x\<in>T. (g \<circ> h') x = x"
       
   146       by clarsimp (metis h'h hom homeomorphism_def)
       
   147   qed
       
   148 qed
       
   149 
       
   150 lemma AR_eq_absolute_extensor:
       
   151   fixes S :: "'a::euclidean_space set"
       
   152   shows "AR S \<longleftrightarrow>
       
   153        (\<forall>f :: 'a * real \<Rightarrow> 'a.
       
   154         \<forall>U T. continuous_on T f \<longrightarrow> f ` T \<subseteq> S \<longrightarrow>
       
   155                closedin (top_of_set U) T \<longrightarrow>
       
   156                 (\<exists>g. continuous_on U g \<and> g ` U \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)))"
       
   157 apply (rule iffI)
       
   158  apply (metis AR_imp_absolute_extensor)
       
   159 apply (simp add: absolute_extensor_imp_AR)
       
   160 done
       
   161 
       
   162 lemma AR_imp_retract:
       
   163   fixes S :: "'a::euclidean_space set"
       
   164   assumes "AR S \<and> closedin (top_of_set U) S"
       
   165     shows "S retract_of U"
       
   166 using AR_imp_absolute_retract assms homeomorphic_refl by blast
       
   167 
       
   168 lemma AR_homeomorphic_AR:
       
   169   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
   170   assumes "AR T" "S homeomorphic T"
       
   171     shows "AR S"
       
   172 unfolding AR_def
       
   173 by (metis assms AR_imp_absolute_retract homeomorphic_trans [of _ S] homeomorphic_sym)
       
   174 
       
   175 lemma homeomorphic_AR_iff_AR:
       
   176   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
   177   shows "S homeomorphic T \<Longrightarrow> AR S \<longleftrightarrow> AR T"
       
   178 by (metis AR_homeomorphic_AR homeomorphic_sym)
       
   179 
       
   180 
       
   181 lemma ANR_imp_absolute_neighbourhood_extensor:
       
   182   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
   183   assumes "ANR S" and contf: "continuous_on T f" and "f ` T \<subseteq> S"
       
   184       and cloUT: "closedin (top_of_set U) T"
       
   185   obtains V g where "T \<subseteq> V" "openin (top_of_set U) V"
       
   186                     "continuous_on V g"
       
   187                     "g ` V \<subseteq> S" "\<And>x. x \<in> T \<Longrightarrow> g x = f x"
       
   188 proof -
       
   189   have "aff_dim S < int (DIM('b \<times> real))"
       
   190     using aff_dim_le_DIM [of S] by simp
       
   191   then obtain C and S' :: "('b * real) set"
       
   192           where C: "convex C" "C \<noteq> {}"
       
   193             and cloCS: "closedin (top_of_set C) S'"
       
   194             and hom: "S homeomorphic S'"
       
   195     by (metis that homeomorphic_closedin_convex)
       
   196   then obtain D where opD: "openin (top_of_set C) D" and "S' retract_of D"
       
   197     using \<open>ANR S\<close> by (auto simp: ANR_def)
       
   198   then obtain r where "S' \<subseteq> D" and contr: "continuous_on D r"
       
   199                   and "r ` D \<subseteq> S'" and rid: "\<And>x. x \<in> S' \<Longrightarrow> r x = x"
       
   200     by (auto simp: retraction_def retract_of_def)
       
   201   obtain g h where homgh: "homeomorphism S S' g h"
       
   202     using hom by (force simp: homeomorphic_def)
       
   203   have "continuous_on (f ` T) g"
       
   204     by (meson \<open>f ` T \<subseteq> S\<close> continuous_on_subset homeomorphism_def homgh)
       
   205   then have contgf: "continuous_on T (g \<circ> f)"
       
   206     by (intro continuous_on_compose contf)
       
   207   have gfTC: "(g \<circ> f) ` T \<subseteq> C"
       
   208   proof -
       
   209     have "g ` S = S'"
       
   210       by (metis (no_types) homeomorphism_def homgh)
       
   211     then show ?thesis
       
   212       by (metis (no_types) assms(3) cloCS closedin_def image_comp image_mono order.trans topspace_euclidean_subtopology)
       
   213   qed
       
   214   obtain f' where contf': "continuous_on U f'"
       
   215               and "f' ` U \<subseteq> C"
       
   216               and eq: "\<And>x. x \<in> T \<Longrightarrow> f' x = (g \<circ> f) x"
       
   217     by (metis Dugundji [OF C cloUT contgf gfTC])
       
   218   show ?thesis
       
   219   proof (rule_tac V = "U \<inter> f' -` D" and g = "h \<circ> r \<circ> f'" in that)
       
   220     show "T \<subseteq> U \<inter> f' -` D"
       
   221       using cloUT closedin_imp_subset \<open>S' \<subseteq> D\<close> \<open>f ` T \<subseteq> S\<close> eq homeomorphism_image1 homgh
       
   222       by fastforce
       
   223     show ope: "openin (top_of_set U) (U \<inter> f' -` D)"
       
   224       using  \<open>f' ` U \<subseteq> C\<close> by (auto simp: opD contf' continuous_openin_preimage)
       
   225     have conth: "continuous_on (r ` f' ` (U \<inter> f' -` D)) h"
       
   226       apply (rule continuous_on_subset [of S'])
       
   227       using homeomorphism_def homgh apply blast
       
   228       using \<open>r ` D \<subseteq> S'\<close> by blast
       
   229     show "continuous_on (U \<inter> f' -` D) (h \<circ> r \<circ> f')"
       
   230       apply (intro continuous_on_compose conth
       
   231                    continuous_on_subset [OF contr] continuous_on_subset [OF contf'], auto)
       
   232       done
       
   233     show "(h \<circ> r \<circ> f') ` (U \<inter> f' -` D) \<subseteq> S"
       
   234       using \<open>homeomorphism S S' g h\<close>  \<open>f' ` U \<subseteq> C\<close>  \<open>r ` D \<subseteq> S'\<close>
       
   235       by (auto simp: homeomorphism_def)
       
   236     show "\<And>x. x \<in> T \<Longrightarrow> (h \<circ> r \<circ> f') x = f x"
       
   237       using \<open>homeomorphism S S' g h\<close> \<open>f ` T \<subseteq> S\<close> eq
       
   238       by (auto simp: rid homeomorphism_def)
       
   239   qed
       
   240 qed
       
   241 
       
   242 
       
   243 corollary ANR_imp_absolute_neighbourhood_retract:
       
   244   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
   245   assumes "ANR S" "S homeomorphic S'"
       
   246       and clo: "closedin (top_of_set U) S'"
       
   247   obtains V where "openin (top_of_set U) V" "S' retract_of V"
       
   248 proof -
       
   249   obtain g h where hom: "homeomorphism S S' g h"
       
   250     using assms by (force simp: homeomorphic_def)
       
   251   have h: "continuous_on S' h" " h ` S' \<subseteq> S"
       
   252     using hom homeomorphism_def apply blast
       
   253     apply (metis hom equalityE homeomorphism_def)
       
   254     done
       
   255     from ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> h clo]
       
   256   obtain V h' where "S' \<subseteq> V" and opUV: "openin (top_of_set U) V"
       
   257                 and h': "continuous_on V h'" "h' ` V \<subseteq> S"
       
   258                 and h'h:"\<And>x. x \<in> S' \<Longrightarrow> h' x = h x"
       
   259     by (blast intro: ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> h clo])
       
   260   have "S' retract_of V"
       
   261   proof (simp add: retraction_def retract_of_def, intro exI conjI \<open>S' \<subseteq> V\<close>)
       
   262     show "continuous_on V (g \<circ> h')"
       
   263       apply (intro continuous_on_compose h')
       
   264       apply (meson hom continuous_on_subset h' homeomorphism_cont1)
       
   265       done
       
   266     show "(g \<circ> h') ` V \<subseteq> S'"
       
   267       using h'  by clarsimp (metis hom subsetD homeomorphism_def imageI)
       
   268     show "\<forall>x\<in>S'. (g \<circ> h') x = x"
       
   269       by clarsimp (metis h'h hom homeomorphism_def)
       
   270   qed
       
   271   then show ?thesis
       
   272     by (rule that [OF opUV])
       
   273 qed
       
   274 
       
   275 corollary ANR_imp_absolute_neighbourhood_retract_UNIV:
       
   276   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
   277   assumes "ANR S" and hom: "S homeomorphic S'" and clo: "closed S'"
       
   278   obtains V where "open V" "S' retract_of V"
       
   279   using ANR_imp_absolute_neighbourhood_retract [OF \<open>ANR S\<close> hom]
       
   280 by (metis clo closed_closedin open_openin subtopology_UNIV)
       
   281 
       
   282 corollary neighbourhood_extension_into_ANR:
       
   283   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
   284   assumes contf: "continuous_on S f" and fim: "f ` S \<subseteq> T" and "ANR T" "closed S"
       
   285   obtains V g where "S \<subseteq> V" "open V" "continuous_on V g"
       
   286                     "g ` V \<subseteq> T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
       
   287   using ANR_imp_absolute_neighbourhood_extensor [OF  \<open>ANR T\<close> contf fim]
       
   288   by (metis \<open>closed S\<close> closed_closedin open_openin subtopology_UNIV)
       
   289 
       
   290 lemma absolute_neighbourhood_extensor_imp_ANR:
       
   291   fixes S :: "'a::euclidean_space set"
       
   292   assumes "\<And>f :: 'a * real \<Rightarrow> 'a.
       
   293            \<And>U T. \<lbrakk>continuous_on T f;  f ` T \<subseteq> S;
       
   294                   closedin (top_of_set U) T\<rbrakk>
       
   295                  \<Longrightarrow> \<exists>V g. T \<subseteq> V \<and> openin (top_of_set U) V \<and>
       
   296                        continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)"
       
   297   shows "ANR S"
       
   298 proof (clarsimp simp: ANR_def)
       
   299   fix U and T :: "('a * real) set"
       
   300   assume "S homeomorphic T" and clo: "closedin (top_of_set U) T"
       
   301   then obtain g h where hom: "homeomorphism S T g h"
       
   302     by (force simp: homeomorphic_def)
       
   303   have h: "continuous_on T h" " h ` T \<subseteq> S"
       
   304     using hom homeomorphism_def apply blast
       
   305     apply (metis hom equalityE homeomorphism_def)
       
   306     done
       
   307   obtain V h' where "T \<subseteq> V" and opV: "openin (top_of_set U) V"
       
   308                 and h': "continuous_on V h'" "h' ` V \<subseteq> S"
       
   309               and h'h: "\<forall>x\<in>T. h' x = h x"
       
   310     using assms [OF h clo] by blast
       
   311   have [simp]: "T \<subseteq> U"
       
   312     using clo closedin_imp_subset by auto
       
   313   have "T retract_of V"
       
   314   proof (simp add: retraction_def retract_of_def, intro exI conjI \<open>T \<subseteq> V\<close>)
       
   315     show "continuous_on V (g \<circ> h')"
       
   316       apply (intro continuous_on_compose h')
       
   317       apply (meson hom continuous_on_subset h' homeomorphism_cont1)
       
   318       done
       
   319     show "(g \<circ> h') ` V \<subseteq> T"
       
   320       using h'  by clarsimp (metis hom subsetD homeomorphism_def imageI)
       
   321     show "\<forall>x\<in>T. (g \<circ> h') x = x"
       
   322       by clarsimp (metis h'h hom homeomorphism_def)
       
   323   qed
       
   324   then show "\<exists>V. openin (top_of_set U) V \<and> T retract_of V"
       
   325     using opV by blast
       
   326 qed
       
   327 
       
   328 lemma ANR_eq_absolute_neighbourhood_extensor:
       
   329   fixes S :: "'a::euclidean_space set"
       
   330   shows "ANR S \<longleftrightarrow>
       
   331          (\<forall>f :: 'a * real \<Rightarrow> 'a.
       
   332           \<forall>U T. continuous_on T f \<longrightarrow> f ` T \<subseteq> S \<longrightarrow>
       
   333                 closedin (top_of_set U) T \<longrightarrow>
       
   334                (\<exists>V g. T \<subseteq> V \<and> openin (top_of_set U) V \<and>
       
   335                        continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)))"
       
   336 apply (rule iffI)
       
   337  apply (metis ANR_imp_absolute_neighbourhood_extensor)
       
   338 apply (simp add: absolute_neighbourhood_extensor_imp_ANR)
       
   339 done
       
   340 
       
   341 lemma ANR_imp_neighbourhood_retract:
       
   342   fixes S :: "'a::euclidean_space set"
       
   343   assumes "ANR S" "closedin (top_of_set U) S"
       
   344   obtains V where "openin (top_of_set U) V" "S retract_of V"
       
   345 using ANR_imp_absolute_neighbourhood_retract assms homeomorphic_refl by blast
       
   346 
       
   347 lemma ANR_imp_absolute_closed_neighbourhood_retract:
       
   348   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
   349   assumes "ANR S" "S homeomorphic S'" and US': "closedin (top_of_set U) S'"
       
   350   obtains V W
       
   351     where "openin (top_of_set U) V"
       
   352           "closedin (top_of_set U) W"
       
   353           "S' \<subseteq> V" "V \<subseteq> W" "S' retract_of W"
       
   354 proof -
       
   355   obtain Z where "openin (top_of_set U) Z" and S'Z: "S' retract_of Z"
       
   356     by (blast intro: assms ANR_imp_absolute_neighbourhood_retract)
       
   357   then have UUZ: "closedin (top_of_set U) (U - Z)"
       
   358     by auto
       
   359   have "S' \<inter> (U - Z) = {}"
       
   360     using \<open>S' retract_of Z\<close> closedin_retract closedin_subtopology by fastforce
       
   361   then obtain V W
       
   362       where "openin (top_of_set U) V"
       
   363         and "openin (top_of_set U) W"
       
   364         and "S' \<subseteq> V" "U - Z \<subseteq> W" "V \<inter> W = {}"
       
   365       using separation_normal_local [OF US' UUZ]  by auto
       
   366   moreover have "S' retract_of U - W"
       
   367     apply (rule retract_of_subset [OF S'Z])
       
   368     using US' \<open>S' \<subseteq> V\<close> \<open>V \<inter> W = {}\<close> closedin_subset apply fastforce
       
   369     using Diff_subset_conv \<open>U - Z \<subseteq> W\<close> by blast
       
   370   ultimately show ?thesis
       
   371     apply (rule_tac V=V and W = "U-W" in that)
       
   372     using openin_imp_subset apply force+
       
   373     done
       
   374 qed
       
   375 
       
   376 lemma ANR_imp_closed_neighbourhood_retract:
       
   377   fixes S :: "'a::euclidean_space set"
       
   378   assumes "ANR S" "closedin (top_of_set U) S"
       
   379   obtains V W where "openin (top_of_set U) V"
       
   380                     "closedin (top_of_set U) W"
       
   381                     "S \<subseteq> V" "V \<subseteq> W" "S retract_of W"
       
   382 by (meson ANR_imp_absolute_closed_neighbourhood_retract assms homeomorphic_refl)
       
   383 
       
   384 lemma ANR_homeomorphic_ANR:
       
   385   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
   386   assumes "ANR T" "S homeomorphic T"
       
   387     shows "ANR S"
       
   388 unfolding ANR_def
       
   389 by (metis assms ANR_imp_absolute_neighbourhood_retract homeomorphic_trans [of _ S] homeomorphic_sym)
       
   390 
       
   391 lemma homeomorphic_ANR_iff_ANR:
       
   392   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
   393   shows "S homeomorphic T \<Longrightarrow> ANR S \<longleftrightarrow> ANR T"
       
   394 by (metis ANR_homeomorphic_ANR homeomorphic_sym)
       
   395 
       
   396 subsection \<open>Analogous properties of ENRs\<close>
       
   397 
       
   398 lemma ENR_imp_absolute_neighbourhood_retract:
       
   399   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
   400   assumes "ENR S" and hom: "S homeomorphic S'"
       
   401       and "S' \<subseteq> U"
       
   402   obtains V where "openin (top_of_set U) V" "S' retract_of V"
       
   403 proof -
       
   404   obtain X where "open X" "S retract_of X"
       
   405     using \<open>ENR S\<close> by (auto simp: ENR_def)
       
   406   then obtain r where "retraction X S r"
       
   407     by (auto simp: retract_of_def)
       
   408   have "locally compact S'"
       
   409     using retract_of_locally_compact open_imp_locally_compact
       
   410           homeomorphic_local_compactness \<open>S retract_of X\<close> \<open>open X\<close> hom by blast
       
   411   then obtain W where UW: "openin (top_of_set U) W"
       
   412                   and WS': "closedin (top_of_set W) S'"
       
   413     apply (rule locally_compact_closedin_open)
       
   414     apply (rename_tac W)
       
   415     apply (rule_tac W = "U \<inter> W" in that, blast)
       
   416     by (simp add: \<open>S' \<subseteq> U\<close> closedin_limpt)
       
   417   obtain f g where hom: "homeomorphism S S' f g"
       
   418     using assms by (force simp: homeomorphic_def)
       
   419   have contg: "continuous_on S' g"
       
   420     using hom homeomorphism_def by blast
       
   421   moreover have "g ` S' \<subseteq> S" by (metis hom equalityE homeomorphism_def)
       
   422   ultimately obtain h where conth: "continuous_on W h" and hg: "\<And>x. x \<in> S' \<Longrightarrow> h x = g x"
       
   423     using Tietze_unbounded [of S' g W] WS' by blast
       
   424   have "W \<subseteq> U" using UW openin_open by auto
       
   425   have "S' \<subseteq> W" using WS' closedin_closed by auto
       
   426   have him: "\<And>x. x \<in> S' \<Longrightarrow> h x \<in> X"
       
   427     by (metis (no_types) \<open>S retract_of X\<close> hg hom homeomorphism_def image_insert insert_absorb insert_iff retract_of_imp_subset subset_eq)
       
   428   have "S' retract_of (W \<inter> h -` X)"
       
   429   proof (simp add: retraction_def retract_of_def, intro exI conjI)
       
   430     show "S' \<subseteq> W" "S' \<subseteq> h -` X"
       
   431       using him WS' closedin_imp_subset by blast+
       
   432     show "continuous_on (W \<inter> h -` X) (f \<circ> r \<circ> h)"
       
   433     proof (intro continuous_on_compose)
       
   434       show "continuous_on (W \<inter> h -` X) h"
       
   435         by (meson conth continuous_on_subset inf_le1)
       
   436       show "continuous_on (h ` (W \<inter> h -` X)) r"
       
   437       proof -
       
   438         have "h ` (W \<inter> h -` X) \<subseteq> X"
       
   439           by blast
       
   440         then show "continuous_on (h ` (W \<inter> h -` X)) r"
       
   441           by (meson \<open>retraction X S r\<close> continuous_on_subset retraction)
       
   442       qed
       
   443       show "continuous_on (r ` h ` (W \<inter> h -` X)) f"
       
   444         apply (rule continuous_on_subset [of S])
       
   445          using hom homeomorphism_def apply blast
       
   446         apply clarify
       
   447         apply (meson \<open>retraction X S r\<close> subsetD imageI retraction_def)
       
   448         done
       
   449     qed
       
   450     show "(f \<circ> r \<circ> h) ` (W \<inter> h -` X) \<subseteq> S'"
       
   451       using \<open>retraction X S r\<close> hom
       
   452       by (auto simp: retraction_def homeomorphism_def)
       
   453     show "\<forall>x\<in>S'. (f \<circ> r \<circ> h) x = x"
       
   454       using \<open>retraction X S r\<close> hom by (auto simp: retraction_def homeomorphism_def hg)
       
   455   qed
       
   456   then show ?thesis
       
   457     apply (rule_tac V = "W \<inter> h -` X" in that)
       
   458      apply (rule openin_trans [OF _ UW])
       
   459      using \<open>continuous_on W h\<close> \<open>open X\<close> continuous_openin_preimage_eq apply blast+
       
   460      done
       
   461 qed
       
   462 
       
   463 corollary ENR_imp_absolute_neighbourhood_retract_UNIV:
       
   464   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
   465   assumes "ENR S" "S homeomorphic S'"
       
   466   obtains T' where "open T'" "S' retract_of T'"
       
   467 by (metis ENR_imp_absolute_neighbourhood_retract UNIV_I assms(1) assms(2) open_openin subsetI subtopology_UNIV)
       
   468 
       
   469 lemma ENR_homeomorphic_ENR:
       
   470   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
   471   assumes "ENR T" "S homeomorphic T"
       
   472     shows "ENR S"
       
   473 unfolding ENR_def
       
   474 by (meson ENR_imp_absolute_neighbourhood_retract_UNIV assms homeomorphic_sym)
       
   475 
       
   476 lemma homeomorphic_ENR_iff_ENR:
       
   477   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
   478   assumes "S homeomorphic T"
       
   479     shows "ENR S \<longleftrightarrow> ENR T"
       
   480 by (meson ENR_homeomorphic_ENR assms homeomorphic_sym)
       
   481 
       
   482 lemma ENR_translation:
       
   483   fixes S :: "'a::euclidean_space set"
       
   484   shows "ENR(image (\<lambda>x. a + x) S) \<longleftrightarrow> ENR S"
       
   485 by (meson homeomorphic_sym homeomorphic_translation homeomorphic_ENR_iff_ENR)
       
   486 
       
   487 lemma ENR_linear_image_eq:
       
   488   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
   489   assumes "linear f" "inj f"
       
   490   shows "ENR (image f S) \<longleftrightarrow> ENR S"
       
   491 apply (rule homeomorphic_ENR_iff_ENR)
       
   492 using assms homeomorphic_sym linear_homeomorphic_image by auto
       
   493 
       
   494 text \<open>Some relations among the concepts. We also relate AR to being a retract of UNIV, which is
       
   495 often a more convenient proxy in the closed case.\<close>
       
   496 
       
   497 lemma AR_imp_ANR: "AR S \<Longrightarrow> ANR S"
       
   498   using ANR_def AR_def by fastforce
       
   499 
       
   500 lemma ENR_imp_ANR:
       
   501   fixes S :: "'a::euclidean_space set"
       
   502   shows "ENR S \<Longrightarrow> ANR S"
       
   503 apply (simp add: ANR_def)
       
   504 by (metis ENR_imp_absolute_neighbourhood_retract closedin_imp_subset)
       
   505 
       
   506 lemma ENR_ANR:
       
   507   fixes S :: "'a::euclidean_space set"
       
   508   shows "ENR S \<longleftrightarrow> ANR S \<and> locally compact S"
       
   509 proof
       
   510   assume "ENR S"
       
   511   then have "locally compact S"
       
   512     using ENR_def open_imp_locally_compact retract_of_locally_compact by auto
       
   513   then show "ANR S \<and> locally compact S"
       
   514     using ENR_imp_ANR \<open>ENR S\<close> by blast
       
   515 next
       
   516   assume "ANR S \<and> locally compact S"
       
   517   then have "ANR S" "locally compact S" by auto
       
   518   then obtain T :: "('a * real) set" where "closed T" "S homeomorphic T"
       
   519     using locally_compact_homeomorphic_closed
       
   520     by (metis DIM_prod DIM_real Suc_eq_plus1 lessI)
       
   521   then show "ENR S"
       
   522     using \<open>ANR S\<close>
       
   523     apply (simp add: ANR_def)
       
   524     apply (drule_tac x=UNIV in spec)
       
   525     apply (drule_tac x=T in spec, clarsimp)
       
   526     apply (meson ENR_def ENR_homeomorphic_ENR open_openin)
       
   527     done
       
   528 qed
       
   529 
       
   530 
       
   531 lemma AR_ANR:
       
   532   fixes S :: "'a::euclidean_space set"
       
   533   shows "AR S \<longleftrightarrow> ANR S \<and> contractible S \<and> S \<noteq> {}"
       
   534         (is "?lhs = ?rhs")
       
   535 proof
       
   536   assume ?lhs
       
   537   obtain C and S' :: "('a * real) set"
       
   538     where "convex C" "C \<noteq> {}" "closedin (top_of_set C) S'" "S homeomorphic S'"
       
   539       apply (rule homeomorphic_closedin_convex [of S, where 'n = "'a * real"])
       
   540       using aff_dim_le_DIM [of S] by auto
       
   541   with \<open>AR S\<close> have "contractible S"
       
   542     apply (simp add: AR_def)
       
   543     apply (drule_tac x=C in spec)
       
   544     apply (drule_tac x="S'" in spec, simp)
       
   545     using convex_imp_contractible homeomorphic_contractible_eq retract_of_contractible by fastforce
       
   546   with \<open>AR S\<close> show ?rhs
       
   547     apply (auto simp: AR_imp_ANR)
       
   548     apply (force simp: AR_def)
       
   549     done
       
   550 next
       
   551   assume ?rhs
       
   552   then obtain a and h:: "real \<times> 'a \<Rightarrow> 'a"
       
   553       where conth: "continuous_on ({0..1} \<times> S) h"
       
   554         and hS: "h ` ({0..1} \<times> S) \<subseteq> S"
       
   555         and [simp]: "\<And>x. h(0, x) = x"
       
   556         and [simp]: "\<And>x. h(1, x) = a"
       
   557         and "ANR S" "S \<noteq> {}"
       
   558     by (auto simp: contractible_def homotopic_with_def)
       
   559   then have "a \<in> S"
       
   560     by (metis all_not_in_conv atLeastAtMost_iff image_subset_iff mem_Sigma_iff order_refl zero_le_one)
       
   561   have "\<exists>g. continuous_on W g \<and> g ` W \<subseteq> S \<and> (\<forall>x\<in>T. g x = f x)"
       
   562          if      f: "continuous_on T f" "f ` T \<subseteq> S"
       
   563             and WT: "closedin (top_of_set W) T"
       
   564          for W T and f :: "'a \<times> real \<Rightarrow> 'a"
       
   565   proof -
       
   566     obtain U g
       
   567       where "T \<subseteq> U" and WU: "openin (top_of_set W) U"
       
   568         and contg: "continuous_on U g"
       
   569         and "g ` U \<subseteq> S" and gf: "\<And>x. x \<in> T \<Longrightarrow> g x = f x"
       
   570       using iffD1 [OF ANR_eq_absolute_neighbourhood_extensor \<open>ANR S\<close>, rule_format, OF f WT]
       
   571       by auto
       
   572     have WWU: "closedin (top_of_set W) (W - U)"
       
   573       using WU closedin_diff by fastforce
       
   574     moreover have "(W - U) \<inter> T = {}"
       
   575       using \<open>T \<subseteq> U\<close> by auto
       
   576     ultimately obtain V V'
       
   577       where WV': "openin (top_of_set W) V'"
       
   578         and WV: "openin (top_of_set W) V"
       
   579         and "W - U \<subseteq> V'" "T \<subseteq> V" "V' \<inter> V = {}"
       
   580       using separation_normal_local [of W "W-U" T] WT by blast
       
   581     then have WVT: "T \<inter> (W - V) = {}"
       
   582       by auto
       
   583     have WWV: "closedin (top_of_set W) (W - V)"
       
   584       using WV closedin_diff by fastforce
       
   585     obtain j :: " 'a \<times> real \<Rightarrow> real"
       
   586       where contj: "continuous_on W j"
       
   587         and j:  "\<And>x. x \<in> W \<Longrightarrow> j x \<in> {0..1}"
       
   588         and j0: "\<And>x. x \<in> W - V \<Longrightarrow> j x = 1"
       
   589         and j1: "\<And>x. x \<in> T \<Longrightarrow> j x = 0"
       
   590       by (rule Urysohn_local [OF WT WWV WVT, of 0 "1::real"]) (auto simp: in_segment)
       
   591     have Weq: "W = (W - V) \<union> (W - V')"
       
   592       using \<open>V' \<inter> V = {}\<close> by force
       
   593     show ?thesis
       
   594     proof (intro conjI exI)
       
   595       have *: "continuous_on (W - V') (\<lambda>x. h (j x, g x))"
       
   596         apply (rule continuous_on_compose2 [OF conth continuous_on_Pair])
       
   597           apply (rule continuous_on_subset [OF contj Diff_subset])
       
   598          apply (rule continuous_on_subset [OF contg])
       
   599          apply (metis Diff_subset_conv Un_commute \<open>W - U \<subseteq> V'\<close>)
       
   600         using j \<open>g ` U \<subseteq> S\<close> \<open>W - U \<subseteq> V'\<close> apply fastforce
       
   601         done
       
   602       show "continuous_on W (\<lambda>x. if x \<in> W - V then a else h (j x, g x))"
       
   603         apply (subst Weq)
       
   604         apply (rule continuous_on_cases_local)
       
   605             apply (simp_all add: Weq [symmetric] WWV continuous_on_const *)
       
   606           using WV' closedin_diff apply fastforce
       
   607          apply (auto simp: j0 j1)
       
   608         done
       
   609     next
       
   610       have "h (j (x, y), g (x, y)) \<in> S" if "(x, y) \<in> W" "(x, y) \<in> V" for x y
       
   611       proof -
       
   612         have "j(x, y) \<in> {0..1}"
       
   613           using j that by blast
       
   614         moreover have "g(x, y) \<in> S"
       
   615           using \<open>V' \<inter> V = {}\<close> \<open>W - U \<subseteq> V'\<close> \<open>g ` U \<subseteq> S\<close> that by fastforce
       
   616         ultimately show ?thesis
       
   617           using hS by blast
       
   618       qed
       
   619       with \<open>a \<in> S\<close> \<open>g ` U \<subseteq> S\<close>
       
   620       show "(\<lambda>x. if x \<in> W - V then a else h (j x, g x)) ` W \<subseteq> S"
       
   621         by auto
       
   622     next
       
   623       show "\<forall>x\<in>T. (if x \<in> W - V then a else h (j x, g x)) = f x"
       
   624         using \<open>T \<subseteq> V\<close> by (auto simp: j0 j1 gf)
       
   625     qed
       
   626   qed
       
   627   then show ?lhs
       
   628     by (simp add: AR_eq_absolute_extensor)
       
   629 qed
       
   630 
       
   631 
       
   632 lemma ANR_retract_of_ANR:
       
   633   fixes S :: "'a::euclidean_space set"
       
   634   assumes "ANR T" "S retract_of T"
       
   635   shows "ANR S"
       
   636 using assms
       
   637 apply (simp add: ANR_eq_absolute_neighbourhood_extensor retract_of_def retraction_def)
       
   638 apply (clarsimp elim!: all_forward)
       
   639 apply (erule impCE, metis subset_trans)
       
   640 apply (clarsimp elim!: ex_forward)
       
   641 apply (rule_tac x="r \<circ> g" in exI)
       
   642 by (metis comp_apply continuous_on_compose continuous_on_subset subsetD imageI image_comp image_mono subset_trans)
       
   643 
       
   644 lemma AR_retract_of_AR:
       
   645   fixes S :: "'a::euclidean_space set"
       
   646   shows "\<lbrakk>AR T; S retract_of T\<rbrakk> \<Longrightarrow> AR S"
       
   647 using ANR_retract_of_ANR AR_ANR retract_of_contractible by fastforce
       
   648 
       
   649 lemma ENR_retract_of_ENR:
       
   650    "\<lbrakk>ENR T; S retract_of T\<rbrakk> \<Longrightarrow> ENR S"
       
   651 by (meson ENR_def retract_of_trans)
       
   652 
       
   653 lemma retract_of_UNIV:
       
   654   fixes S :: "'a::euclidean_space set"
       
   655   shows "S retract_of UNIV \<longleftrightarrow> AR S \<and> closed S"
       
   656 by (metis AR_ANR AR_imp_retract ENR_def ENR_imp_ANR closed_UNIV closed_closedin contractible_UNIV empty_not_UNIV open_UNIV retract_of_closed retract_of_contractible retract_of_empty(1) subtopology_UNIV)
       
   657 
       
   658 lemma compact_AR:
       
   659   fixes S :: "'a::euclidean_space set"
       
   660   shows "compact S \<and> AR S \<longleftrightarrow> compact S \<and> S retract_of UNIV"
       
   661 using compact_imp_closed retract_of_UNIV by blast
       
   662 
       
   663 text \<open>More properties of ARs, ANRs and ENRs\<close>
       
   664 
       
   665 lemma not_AR_empty [simp]: "\<not> AR({})"
       
   666   by (auto simp: AR_def)
       
   667 
       
   668 lemma ENR_empty [simp]: "ENR {}"
       
   669   by (simp add: ENR_def)
       
   670 
       
   671 lemma ANR_empty [simp]: "ANR ({} :: 'a::euclidean_space set)"
       
   672   by (simp add: ENR_imp_ANR)
       
   673 
       
   674 lemma convex_imp_AR:
       
   675   fixes S :: "'a::euclidean_space set"
       
   676   shows "\<lbrakk>convex S; S \<noteq> {}\<rbrakk> \<Longrightarrow> AR S"
       
   677 apply (rule absolute_extensor_imp_AR)
       
   678 apply (rule Dugundji, assumption+)
       
   679 by blast
       
   680 
       
   681 lemma convex_imp_ANR:
       
   682   fixes S :: "'a::euclidean_space set"
       
   683   shows "convex S \<Longrightarrow> ANR S"
       
   684 using ANR_empty AR_imp_ANR convex_imp_AR by blast
       
   685 
       
   686 lemma ENR_convex_closed:
       
   687   fixes S :: "'a::euclidean_space set"
       
   688   shows "\<lbrakk>closed S; convex S\<rbrakk> \<Longrightarrow> ENR S"
       
   689 using ENR_def ENR_empty convex_imp_AR retract_of_UNIV by blast
       
   690 
       
   691 lemma AR_UNIV [simp]: "AR (UNIV :: 'a::euclidean_space set)"
       
   692   using retract_of_UNIV by auto
       
   693 
       
   694 lemma ANR_UNIV [simp]: "ANR (UNIV :: 'a::euclidean_space set)"
       
   695   by (simp add: AR_imp_ANR)
       
   696 
       
   697 lemma ENR_UNIV [simp]:"ENR UNIV"
       
   698   using ENR_def by blast
       
   699 
       
   700 lemma AR_singleton:
       
   701     fixes a :: "'a::euclidean_space"
       
   702     shows "AR {a}"
       
   703   using retract_of_UNIV by blast
       
   704 
       
   705 lemma ANR_singleton:
       
   706     fixes a :: "'a::euclidean_space"
       
   707     shows "ANR {a}"
       
   708   by (simp add: AR_imp_ANR AR_singleton)
       
   709 
       
   710 lemma ENR_singleton: "ENR {a}"
       
   711   using ENR_def by blast
       
   712 
       
   713 text \<open>ARs closed under union\<close>
       
   714 
       
   715 lemma AR_closed_Un_local_aux:
       
   716   fixes U :: "'a::euclidean_space set"
       
   717   assumes "closedin (top_of_set U) S"
       
   718           "closedin (top_of_set U) T"
       
   719           "AR S" "AR T" "AR(S \<inter> T)"
       
   720   shows "(S \<union> T) retract_of U"
       
   721 proof -
       
   722   have "S \<inter> T \<noteq> {}"
       
   723     using assms AR_def by fastforce
       
   724   have "S \<subseteq> U" "T \<subseteq> U"
       
   725     using assms by (auto simp: closedin_imp_subset)
       
   726   define S' where "S' \<equiv> {x \<in> U. setdist {x} S \<le> setdist {x} T}"
       
   727   define T' where "T' \<equiv> {x \<in> U. setdist {x} T \<le> setdist {x} S}"
       
   728   define W  where "W \<equiv> {x \<in> U. setdist {x} S = setdist {x} T}"
       
   729   have US': "closedin (top_of_set U) S'"
       
   730     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} S - setdist {x} T" "{..0}"]
       
   731     by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist)
       
   732   have UT': "closedin (top_of_set U) T'"
       
   733     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} T - setdist {x} S" "{..0}"]
       
   734     by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist)
       
   735   have "S \<subseteq> S'"
       
   736     using S'_def \<open>S \<subseteq> U\<close> setdist_sing_in_set by fastforce
       
   737   have "T \<subseteq> T'"
       
   738     using T'_def \<open>T \<subseteq> U\<close> setdist_sing_in_set by fastforce
       
   739   have "S \<inter> T \<subseteq> W" "W \<subseteq> U"
       
   740     using \<open>S \<subseteq> U\<close> by (auto simp: W_def setdist_sing_in_set)
       
   741   have "(S \<inter> T) retract_of W"
       
   742     apply (rule AR_imp_absolute_retract [OF \<open>AR(S \<inter> T)\<close>])
       
   743      apply (simp add: homeomorphic_refl)
       
   744     apply (rule closedin_subset_trans [of U])
       
   745     apply (simp_all add: assms closedin_Int \<open>S \<inter> T \<subseteq> W\<close> \<open>W \<subseteq> U\<close>)
       
   746     done
       
   747   then obtain r0
       
   748     where "S \<inter> T \<subseteq> W" and contr0: "continuous_on W r0"
       
   749       and "r0 ` W \<subseteq> S \<inter> T"
       
   750       and r0 [simp]: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> r0 x = x"
       
   751       by (auto simp: retract_of_def retraction_def)
       
   752   have ST: "x \<in> W \<Longrightarrow> x \<in> S \<longleftrightarrow> x \<in> T" for x
       
   753     using setdist_eq_0_closedin \<open>S \<inter> T \<noteq> {}\<close> assms
       
   754     by (force simp: W_def setdist_sing_in_set)
       
   755   have "S' \<inter> T' = W"
       
   756     by (auto simp: S'_def T'_def W_def)
       
   757   then have cloUW: "closedin (top_of_set U) W"
       
   758     using closedin_Int US' UT' by blast
       
   759   define r where "r \<equiv> \<lambda>x. if x \<in> W then r0 x else x"
       
   760   have "r ` (W \<union> S) \<subseteq> S" "r ` (W \<union> T) \<subseteq> T"
       
   761     using \<open>r0 ` W \<subseteq> S \<inter> T\<close> r_def by auto
       
   762   have contr: "continuous_on (W \<union> (S \<union> T)) r"
       
   763   unfolding r_def
       
   764   proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id])
       
   765     show "closedin (top_of_set (W \<union> (S \<union> T))) W"
       
   766       using \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> \<open>W \<subseteq> U\<close> \<open>closedin (top_of_set U) W\<close> closedin_subset_trans by fastforce
       
   767     show "closedin (top_of_set (W \<union> (S \<union> T))) (S \<union> T)"
       
   768       by (meson \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> \<open>W \<subseteq> U\<close> assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2)
       
   769     show "\<And>x. x \<in> W \<and> x \<notin> W \<or> x \<in> S \<union> T \<and> x \<in> W \<Longrightarrow> r0 x = x"
       
   770       by (auto simp: ST)
       
   771   qed
       
   772   have cloUWS: "closedin (top_of_set U) (W \<union> S)"
       
   773     by (simp add: cloUW assms closedin_Un)
       
   774   obtain g where contg: "continuous_on U g"
       
   775              and "g ` U \<subseteq> S" and geqr: "\<And>x. x \<in> W \<union> S \<Longrightarrow> g x = r x"
       
   776     apply (rule AR_imp_absolute_extensor [OF \<open>AR S\<close> _ _ cloUWS])
       
   777       apply (rule continuous_on_subset [OF contr])
       
   778       using \<open>r ` (W \<union> S) \<subseteq> S\<close> apply auto
       
   779     done
       
   780   have cloUWT: "closedin (top_of_set U) (W \<union> T)"
       
   781     by (simp add: cloUW assms closedin_Un)
       
   782   obtain h where conth: "continuous_on U h"
       
   783              and "h ` U \<subseteq> T" and heqr: "\<And>x. x \<in> W \<union> T \<Longrightarrow> h x = r x"
       
   784     apply (rule AR_imp_absolute_extensor [OF \<open>AR T\<close> _ _ cloUWT])
       
   785       apply (rule continuous_on_subset [OF contr])
       
   786       using \<open>r ` (W \<union> T) \<subseteq> T\<close> apply auto
       
   787     done
       
   788   have "U = S' \<union> T'"
       
   789     by (force simp: S'_def T'_def)
       
   790   then have cont: "continuous_on U (\<lambda>x. if x \<in> S' then g x else h x)"
       
   791     apply (rule ssubst)
       
   792     apply (rule continuous_on_cases_local)
       
   793     using US' UT' \<open>S' \<inter> T' = W\<close> \<open>U = S' \<union> T'\<close>
       
   794           contg conth continuous_on_subset geqr heqr apply auto
       
   795     done
       
   796   have UST: "(\<lambda>x. if x \<in> S' then g x else h x) ` U \<subseteq> S \<union> T"
       
   797     using \<open>g ` U \<subseteq> S\<close> \<open>h ` U \<subseteq> T\<close> by auto
       
   798   show ?thesis
       
   799     apply (simp add: retract_of_def retraction_def \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close>)
       
   800     apply (rule_tac x="\<lambda>x. if x \<in> S' then g x else h x" in exI)
       
   801     apply (intro conjI cont UST)
       
   802     by (metis IntI ST Un_iff \<open>S \<subseteq> S'\<close> \<open>S' \<inter> T' = W\<close> \<open>T \<subseteq> T'\<close> subsetD geqr heqr r0 r_def)
       
   803 qed
       
   804 
       
   805 
       
   806 lemma AR_closed_Un_local:
       
   807   fixes S :: "'a::euclidean_space set"
       
   808   assumes STS: "closedin (top_of_set (S \<union> T)) S"
       
   809       and STT: "closedin (top_of_set (S \<union> T)) T"
       
   810       and "AR S" "AR T" "AR(S \<inter> T)"
       
   811     shows "AR(S \<union> T)"
       
   812 proof -
       
   813   have "C retract_of U"
       
   814        if hom: "S \<union> T homeomorphic C" and UC: "closedin (top_of_set U) C"
       
   815        for U and C :: "('a * real) set"
       
   816   proof -
       
   817     obtain f g where hom: "homeomorphism (S \<union> T) C f g"
       
   818       using hom by (force simp: homeomorphic_def)
       
   819     have US: "closedin (top_of_set U) (C \<inter> g -` S)"
       
   820       apply (rule closedin_trans [OF _ UC])
       
   821       apply (rule continuous_closedin_preimage_gen [OF _ _ STS])
       
   822       using hom homeomorphism_def apply blast
       
   823       apply (metis hom homeomorphism_def set_eq_subset)
       
   824       done
       
   825     have UT: "closedin (top_of_set U) (C \<inter> g -` T)"
       
   826       apply (rule closedin_trans [OF _ UC])
       
   827       apply (rule continuous_closedin_preimage_gen [OF _ _ STT])
       
   828       using hom homeomorphism_def apply blast
       
   829       apply (metis hom homeomorphism_def set_eq_subset)
       
   830       done
       
   831     have ARS: "AR (C \<inter> g -` S)"
       
   832       apply (rule AR_homeomorphic_AR [OF \<open>AR S\<close>])
       
   833       apply (simp add: homeomorphic_def)
       
   834       apply (rule_tac x=g in exI)
       
   835       apply (rule_tac x=f in exI)
       
   836       using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset)
       
   837       apply (rule_tac x="f x" in image_eqI, auto)
       
   838       done
       
   839     have ART: "AR (C \<inter> g -` T)"
       
   840       apply (rule AR_homeomorphic_AR [OF \<open>AR T\<close>])
       
   841       apply (simp add: homeomorphic_def)
       
   842       apply (rule_tac x=g in exI)
       
   843       apply (rule_tac x=f in exI)
       
   844       using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset)
       
   845       apply (rule_tac x="f x" in image_eqI, auto)
       
   846       done
       
   847     have ARI: "AR ((C \<inter> g -` S) \<inter> (C \<inter> g -` T))"
       
   848       apply (rule AR_homeomorphic_AR [OF \<open>AR (S \<inter> T)\<close>])
       
   849       apply (simp add: homeomorphic_def)
       
   850       apply (rule_tac x=g in exI)
       
   851       apply (rule_tac x=f in exI)
       
   852       using hom
       
   853       apply (auto simp: homeomorphism_def elim!: continuous_on_subset)
       
   854       apply (rule_tac x="f x" in image_eqI, auto)
       
   855       done
       
   856     have "C = (C \<inter> g -` S) \<union> (C \<inter> g -` T)"
       
   857       using hom  by (auto simp: homeomorphism_def)
       
   858     then show ?thesis
       
   859       by (metis AR_closed_Un_local_aux [OF US UT ARS ART ARI])
       
   860   qed
       
   861   then show ?thesis
       
   862     by (force simp: AR_def)
       
   863 qed
       
   864 
       
   865 corollary AR_closed_Un:
       
   866   fixes S :: "'a::euclidean_space set"
       
   867   shows "\<lbrakk>closed S; closed T; AR S; AR T; AR (S \<inter> T)\<rbrakk> \<Longrightarrow> AR (S \<union> T)"
       
   868 by (metis AR_closed_Un_local_aux closed_closedin retract_of_UNIV subtopology_UNIV)
       
   869 
       
   870 text \<open>ANRs closed under union\<close>
       
   871 
       
   872 lemma ANR_closed_Un_local_aux:
       
   873   fixes U :: "'a::euclidean_space set"
       
   874   assumes US: "closedin (top_of_set U) S"
       
   875       and UT: "closedin (top_of_set U) T"
       
   876       and "ANR S" "ANR T" "ANR(S \<inter> T)"
       
   877   obtains V where "openin (top_of_set U) V" "(S \<union> T) retract_of V"
       
   878 proof (cases "S = {} \<or> T = {}")
       
   879   case True with assms that show ?thesis
       
   880     by (metis ANR_imp_neighbourhood_retract Un_commute inf_bot_right sup_inf_absorb)
       
   881 next
       
   882   case False
       
   883   then have [simp]: "S \<noteq> {}" "T \<noteq> {}" by auto
       
   884   have "S \<subseteq> U" "T \<subseteq> U"
       
   885     using assms by (auto simp: closedin_imp_subset)
       
   886   define S' where "S' \<equiv> {x \<in> U. setdist {x} S \<le> setdist {x} T}"
       
   887   define T' where "T' \<equiv> {x \<in> U. setdist {x} T \<le> setdist {x} S}"
       
   888   define W  where "W \<equiv> {x \<in> U. setdist {x} S = setdist {x} T}"
       
   889   have cloUS': "closedin (top_of_set U) S'"
       
   890     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} S - setdist {x} T" "{..0}"]
       
   891     by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist)
       
   892   have cloUT': "closedin (top_of_set U) T'"
       
   893     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} T - setdist {x} S" "{..0}"]
       
   894     by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist)
       
   895   have "S \<subseteq> S'"
       
   896     using S'_def \<open>S \<subseteq> U\<close> setdist_sing_in_set by fastforce
       
   897   have "T \<subseteq> T'"
       
   898     using T'_def \<open>T \<subseteq> U\<close> setdist_sing_in_set by fastforce
       
   899   have "S' \<union> T' = U"
       
   900     by (auto simp: S'_def T'_def)
       
   901   have "W \<subseteq> S'"
       
   902     by (simp add: Collect_mono S'_def W_def)
       
   903   have "W \<subseteq> T'"
       
   904     by (simp add: Collect_mono T'_def W_def)
       
   905   have ST_W: "S \<inter> T \<subseteq> W" and "W \<subseteq> U"
       
   906     using \<open>S \<subseteq> U\<close> by (force simp: W_def setdist_sing_in_set)+
       
   907   have "S' \<inter> T' = W"
       
   908     by (auto simp: S'_def T'_def W_def)
       
   909   then have cloUW: "closedin (top_of_set U) W"
       
   910     using closedin_Int cloUS' cloUT' by blast
       
   911   obtain W' W0 where "openin (top_of_set W) W'"
       
   912                  and cloWW0: "closedin (top_of_set W) W0"
       
   913                  and "S \<inter> T \<subseteq> W'" "W' \<subseteq> W0"
       
   914                  and ret: "(S \<inter> T) retract_of W0"
       
   915     apply (rule ANR_imp_closed_neighbourhood_retract [OF \<open>ANR(S \<inter> T)\<close>])
       
   916     apply (rule closedin_subset_trans [of U, OF _ ST_W \<open>W \<subseteq> U\<close>])
       
   917     apply (blast intro: assms)+
       
   918     done
       
   919   then obtain U0 where opeUU0: "openin (top_of_set U) U0"
       
   920                    and U0: "S \<inter> T \<subseteq> U0" "U0 \<inter> W \<subseteq> W0"
       
   921     unfolding openin_open  using \<open>W \<subseteq> U\<close> by blast
       
   922   have "W0 \<subseteq> U"
       
   923     using \<open>W \<subseteq> U\<close> cloWW0 closedin_subset by fastforce
       
   924   obtain r0
       
   925     where "S \<inter> T \<subseteq> W0" and contr0: "continuous_on W0 r0" and "r0 ` W0 \<subseteq> S \<inter> T"
       
   926       and r0 [simp]: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> r0 x = x"
       
   927     using ret  by (force simp: retract_of_def retraction_def)
       
   928   have ST: "x \<in> W \<Longrightarrow> x \<in> S \<longleftrightarrow> x \<in> T" for x
       
   929     using assms by (auto simp: W_def setdist_sing_in_set dest!: setdist_eq_0_closedin)
       
   930   define r where "r \<equiv> \<lambda>x. if x \<in> W0 then r0 x else x"
       
   931   have "r ` (W0 \<union> S) \<subseteq> S" "r ` (W0 \<union> T) \<subseteq> T"
       
   932     using \<open>r0 ` W0 \<subseteq> S \<inter> T\<close> r_def by auto
       
   933   have contr: "continuous_on (W0 \<union> (S \<union> T)) r"
       
   934   unfolding r_def
       
   935   proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id])
       
   936     show "closedin (top_of_set (W0 \<union> (S \<union> T))) W0"
       
   937       apply (rule closedin_subset_trans [of U])
       
   938       using cloWW0 cloUW closedin_trans \<open>W0 \<subseteq> U\<close> \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> apply blast+
       
   939       done
       
   940     show "closedin (top_of_set (W0 \<union> (S \<union> T))) (S \<union> T)"
       
   941       by (meson \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> \<open>W0 \<subseteq> U\<close> assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2)
       
   942     show "\<And>x. x \<in> W0 \<and> x \<notin> W0 \<or> x \<in> S \<union> T \<and> x \<in> W0 \<Longrightarrow> r0 x = x"
       
   943       using ST cloWW0 closedin_subset by fastforce
       
   944   qed
       
   945   have cloS'WS: "closedin (top_of_set S') (W0 \<union> S)"
       
   946     by (meson closedin_subset_trans US cloUS' \<open>S \<subseteq> S'\<close> \<open>W \<subseteq> S'\<close> cloUW cloWW0 
       
   947               closedin_Un closedin_imp_subset closedin_trans)
       
   948   obtain W1 g where "W0 \<union> S \<subseteq> W1" and contg: "continuous_on W1 g"
       
   949                 and opeSW1: "openin (top_of_set S') W1"
       
   950                 and "g ` W1 \<subseteq> S" and geqr: "\<And>x. x \<in> W0 \<union> S \<Longrightarrow> g x = r x"
       
   951     apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> _ \<open>r ` (W0 \<union> S) \<subseteq> S\<close> cloS'WS])
       
   952      apply (rule continuous_on_subset [OF contr], blast+)
       
   953     done
       
   954   have cloT'WT: "closedin (top_of_set T') (W0 \<union> T)"
       
   955     by (meson closedin_subset_trans UT cloUT' \<open>T \<subseteq> T'\<close> \<open>W \<subseteq> T'\<close> cloUW cloWW0 
       
   956               closedin_Un closedin_imp_subset closedin_trans)
       
   957   obtain W2 h where "W0 \<union> T \<subseteq> W2" and conth: "continuous_on W2 h"
       
   958                 and opeSW2: "openin (top_of_set T') W2"
       
   959                 and "h ` W2 \<subseteq> T" and heqr: "\<And>x. x \<in> W0 \<union> T \<Longrightarrow> h x = r x"
       
   960     apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> _ \<open>r ` (W0 \<union> T) \<subseteq> T\<close> cloT'WT])
       
   961      apply (rule continuous_on_subset [OF contr], blast+)
       
   962     done
       
   963   have "S' \<inter> T' = W"
       
   964     by (force simp: S'_def T'_def W_def)
       
   965   obtain O1 O2 where "open O1" "W1 = S' \<inter> O1" "open O2" "W2 = T' \<inter> O2"
       
   966     using opeSW1 opeSW2 by (force simp: openin_open)
       
   967   show ?thesis
       
   968   proof
       
   969     have eq: "W1 - (W - U0) \<union> (W2 - (W - U0)) =
       
   970          ((U - T') \<inter> O1 \<union> (U - S') \<inter> O2 \<union> U \<inter> O1 \<inter> O2) - (W - U0)"
       
   971      using \<open>U0 \<inter> W \<subseteq> W0\<close> \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W0 \<union> T \<subseteq> W2\<close>
       
   972       by (auto simp: \<open>S' \<union> T' = U\<close> [symmetric] \<open>S' \<inter> T' = W\<close> [symmetric] \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close>)
       
   973     show "openin (top_of_set U) (W1 - (W - U0) \<union> (W2 - (W - U0)))"
       
   974       apply (subst eq)
       
   975       apply (intro openin_Un openin_Int_open openin_diff closedin_diff cloUW opeUU0 cloUS' cloUT' \<open>open O1\<close> \<open>open O2\<close>, simp_all)
       
   976       done
       
   977     have cloW1: "closedin (top_of_set (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W1 - (W - U0))"
       
   978       using cloUS' apply (simp add: closedin_closed)
       
   979       apply (erule ex_forward)
       
   980       using U0 \<open>W0 \<union> S \<subseteq> W1\<close>
       
   981       apply (auto simp: \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<union> T' = U\<close> [symmetric]\<open>S' \<inter> T' = W\<close> [symmetric])
       
   982       done
       
   983     have cloW2: "closedin (top_of_set (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W2 - (W - U0))"
       
   984       using cloUT' apply (simp add: closedin_closed)
       
   985       apply (erule ex_forward)
       
   986       using U0 \<open>W0 \<union> T \<subseteq> W2\<close>
       
   987       apply (auto simp: \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<union> T' = U\<close> [symmetric]\<open>S' \<inter> T' = W\<close> [symmetric])
       
   988       done
       
   989     have *: "\<forall>x\<in>S \<union> T. (if x \<in> S' then g x else h x) = x"
       
   990       using ST \<open>S' \<inter> T' = W\<close> cloT'WT closedin_subset geqr heqr 
       
   991       apply (auto simp: r_def, fastforce)
       
   992       using \<open>S \<subseteq> S'\<close> \<open>T \<subseteq> T'\<close> \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W1 = S' \<inter> O1\<close>  by auto
       
   993     have "\<exists>r. continuous_on (W1 - (W - U0) \<union> (W2 - (W - U0))) r \<and>
       
   994               r ` (W1 - (W - U0) \<union> (W2 - (W - U0))) \<subseteq> S \<union> T \<and> 
       
   995               (\<forall>x\<in>S \<union> T. r x = x)"
       
   996       apply (rule_tac x = "\<lambda>x. if  x \<in> S' then g x else h x" in exI)
       
   997       apply (intro conjI *)
       
   998       apply (rule continuous_on_cases_local 
       
   999                   [OF cloW1 cloW2 continuous_on_subset [OF contg] continuous_on_subset [OF conth]])
       
  1000       using \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<inter> T' = W\<close>
       
  1001             \<open>g ` W1 \<subseteq> S\<close> \<open>h ` W2 \<subseteq> T\<close> apply auto
       
  1002       using \<open>U0 \<inter> W \<subseteq> W0\<close> \<open>W0 \<union> S \<subseteq> W1\<close> apply (fastforce simp add: geqr heqr)+
       
  1003       done
       
  1004     then show "S \<union> T retract_of W1 - (W - U0) \<union> (W2 - (W - U0))"
       
  1005       using  \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W0 \<union> T \<subseteq> W2\<close> ST opeUU0 U0
       
  1006       by (auto simp: retract_of_def retraction_def)
       
  1007   qed
       
  1008 qed
       
  1009 
       
  1010 
       
  1011 lemma ANR_closed_Un_local:
       
  1012   fixes S :: "'a::euclidean_space set"
       
  1013   assumes STS: "closedin (top_of_set (S \<union> T)) S"
       
  1014       and STT: "closedin (top_of_set (S \<union> T)) T"
       
  1015       and "ANR S" "ANR T" "ANR(S \<inter> T)" 
       
  1016     shows "ANR(S \<union> T)"
       
  1017 proof -
       
  1018   have "\<exists>T. openin (top_of_set U) T \<and> C retract_of T"
       
  1019        if hom: "S \<union> T homeomorphic C" and UC: "closedin (top_of_set U) C"
       
  1020        for U and C :: "('a * real) set"
       
  1021   proof -
       
  1022     obtain f g where hom: "homeomorphism (S \<union> T) C f g"
       
  1023       using hom by (force simp: homeomorphic_def)
       
  1024     have US: "closedin (top_of_set U) (C \<inter> g -` S)"
       
  1025       apply (rule closedin_trans [OF _ UC])
       
  1026       apply (rule continuous_closedin_preimage_gen [OF _ _ STS])
       
  1027       using hom [unfolded homeomorphism_def] apply blast
       
  1028       apply (metis hom homeomorphism_def set_eq_subset)
       
  1029       done
       
  1030     have UT: "closedin (top_of_set U) (C \<inter> g -` T)"
       
  1031       apply (rule closedin_trans [OF _ UC])
       
  1032       apply (rule continuous_closedin_preimage_gen [OF _ _ STT])
       
  1033       using hom [unfolded homeomorphism_def] apply blast
       
  1034       apply (metis hom homeomorphism_def set_eq_subset)
       
  1035       done
       
  1036     have ANRS: "ANR (C \<inter> g -` S)"
       
  1037       apply (rule ANR_homeomorphic_ANR [OF \<open>ANR S\<close>])
       
  1038       apply (simp add: homeomorphic_def)
       
  1039       apply (rule_tac x=g in exI)
       
  1040       apply (rule_tac x=f in exI)
       
  1041       using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset)
       
  1042       apply (rule_tac x="f x" in image_eqI, auto)
       
  1043       done
       
  1044     have ANRT: "ANR (C \<inter> g -` T)"
       
  1045       apply (rule ANR_homeomorphic_ANR [OF \<open>ANR T\<close>])
       
  1046       apply (simp add: homeomorphic_def)
       
  1047       apply (rule_tac x=g in exI)
       
  1048       apply (rule_tac x=f in exI)
       
  1049       using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset)
       
  1050       apply (rule_tac x="f x" in image_eqI, auto)
       
  1051       done
       
  1052     have ANRI: "ANR ((C \<inter> g -` S) \<inter> (C \<inter> g -` T))"
       
  1053       apply (rule ANR_homeomorphic_ANR [OF \<open>ANR (S \<inter> T)\<close>])
       
  1054       apply (simp add: homeomorphic_def)
       
  1055       apply (rule_tac x=g in exI)
       
  1056       apply (rule_tac x=f in exI)
       
  1057       using hom
       
  1058       apply (auto simp: homeomorphism_def elim!: continuous_on_subset)
       
  1059       apply (rule_tac x="f x" in image_eqI, auto)
       
  1060       done
       
  1061     have "C = (C \<inter> g -` S) \<union> (C \<inter> g -` T)"
       
  1062       using hom by (auto simp: homeomorphism_def)
       
  1063     then show ?thesis
       
  1064       by (metis ANR_closed_Un_local_aux [OF US UT ANRS ANRT ANRI])
       
  1065   qed
       
  1066   then show ?thesis
       
  1067     by (auto simp: ANR_def)
       
  1068 qed    
       
  1069 
       
  1070 corollary ANR_closed_Un:
       
  1071   fixes S :: "'a::euclidean_space set"
       
  1072   shows "\<lbrakk>closed S; closed T; ANR S; ANR T; ANR (S \<inter> T)\<rbrakk> \<Longrightarrow> ANR (S \<union> T)"
       
  1073 by (simp add: ANR_closed_Un_local closedin_def diff_eq open_Compl openin_open_Int)
       
  1074 
       
  1075 lemma ANR_openin:
       
  1076   fixes S :: "'a::euclidean_space set"
       
  1077   assumes "ANR T" and opeTS: "openin (top_of_set T) S"
       
  1078   shows "ANR S"
       
  1079 proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor)
       
  1080   fix f :: "'a \<times> real \<Rightarrow> 'a" and U C
       
  1081   assume contf: "continuous_on C f" and fim: "f ` C \<subseteq> S"
       
  1082      and cloUC: "closedin (top_of_set U) C"
       
  1083   have "f ` C \<subseteq> T"
       
  1084     using fim opeTS openin_imp_subset by blast
       
  1085   obtain W g where "C \<subseteq> W"
       
  1086                and UW: "openin (top_of_set U) W"
       
  1087                and contg: "continuous_on W g"
       
  1088                and gim: "g ` W \<subseteq> T"
       
  1089                and geq: "\<And>x. x \<in> C \<Longrightarrow> g x = f x"
       
  1090     apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> contf \<open>f ` C \<subseteq> T\<close> cloUC])
       
  1091     using fim by auto
       
  1092   show "\<exists>V g. C \<subseteq> V \<and> openin (top_of_set U) V \<and> continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x\<in>C. g x = f x)"
       
  1093   proof (intro exI conjI)
       
  1094     show "C \<subseteq> W \<inter> g -` S"
       
  1095       using \<open>C \<subseteq> W\<close> fim geq by blast
       
  1096     show "openin (top_of_set U) (W \<inter> g -` S)"
       
  1097       by (metis (mono_tags, lifting) UW contg continuous_openin_preimage gim opeTS openin_trans)
       
  1098     show "continuous_on (W \<inter> g -` S) g"
       
  1099       by (blast intro: continuous_on_subset [OF contg])
       
  1100     show "g ` (W \<inter> g -` S) \<subseteq> S"
       
  1101       using gim by blast
       
  1102     show "\<forall>x\<in>C. g x = f x"
       
  1103       using geq by blast
       
  1104   qed
       
  1105 qed
       
  1106 
       
  1107 lemma ENR_openin:
       
  1108     fixes S :: "'a::euclidean_space set"
       
  1109     assumes "ENR T" and opeTS: "openin (top_of_set T) S"
       
  1110     shows "ENR S"
       
  1111   using assms apply (simp add: ENR_ANR)
       
  1112   using ANR_openin locally_open_subset by blast
       
  1113 
       
  1114 lemma ANR_neighborhood_retract:
       
  1115     fixes S :: "'a::euclidean_space set"
       
  1116     assumes "ANR U" "S retract_of T" "openin (top_of_set U) T"
       
  1117     shows "ANR S"
       
  1118   using ANR_openin ANR_retract_of_ANR assms by blast
       
  1119 
       
  1120 lemma ENR_neighborhood_retract:
       
  1121     fixes S :: "'a::euclidean_space set"
       
  1122     assumes "ENR U" "S retract_of T" "openin (top_of_set U) T"
       
  1123     shows "ENR S"
       
  1124   using ENR_openin ENR_retract_of_ENR assms by blast
       
  1125 
       
  1126 lemma ANR_rel_interior:
       
  1127   fixes S :: "'a::euclidean_space set"
       
  1128   shows "ANR S \<Longrightarrow> ANR(rel_interior S)"
       
  1129    by (blast intro: ANR_openin openin_set_rel_interior)
       
  1130 
       
  1131 lemma ANR_delete:
       
  1132   fixes S :: "'a::euclidean_space set"
       
  1133   shows "ANR S \<Longrightarrow> ANR(S - {a})"
       
  1134    by (blast intro: ANR_openin openin_delete openin_subtopology_self)
       
  1135 
       
  1136 lemma ENR_rel_interior:
       
  1137   fixes S :: "'a::euclidean_space set"
       
  1138   shows "ENR S \<Longrightarrow> ENR(rel_interior S)"
       
  1139    by (blast intro: ENR_openin openin_set_rel_interior)
       
  1140 
       
  1141 lemma ENR_delete:
       
  1142   fixes S :: "'a::euclidean_space set"
       
  1143   shows "ENR S \<Longrightarrow> ENR(S - {a})"
       
  1144    by (blast intro: ENR_openin openin_delete openin_subtopology_self)
       
  1145 
       
  1146 lemma open_imp_ENR: "open S \<Longrightarrow> ENR S"
       
  1147     using ENR_def by blast
       
  1148 
       
  1149 lemma open_imp_ANR:
       
  1150     fixes S :: "'a::euclidean_space set"
       
  1151     shows "open S \<Longrightarrow> ANR S"
       
  1152   by (simp add: ENR_imp_ANR open_imp_ENR)
       
  1153 
       
  1154 lemma ANR_ball [iff]:
       
  1155     fixes a :: "'a::euclidean_space"
       
  1156     shows "ANR(ball a r)"
       
  1157   by (simp add: convex_imp_ANR)
       
  1158 
       
  1159 lemma ENR_ball [iff]: "ENR(ball a r)"
       
  1160   by (simp add: open_imp_ENR)
       
  1161 
       
  1162 lemma AR_ball [simp]:
       
  1163     fixes a :: "'a::euclidean_space"
       
  1164     shows "AR(ball a r) \<longleftrightarrow> 0 < r"
       
  1165   by (auto simp: AR_ANR convex_imp_contractible)
       
  1166 
       
  1167 lemma ANR_cball [iff]:
       
  1168     fixes a :: "'a::euclidean_space"
       
  1169     shows "ANR(cball a r)"
       
  1170   by (simp add: convex_imp_ANR)
       
  1171 
       
  1172 lemma ENR_cball:
       
  1173     fixes a :: "'a::euclidean_space"
       
  1174     shows "ENR(cball a r)"
       
  1175   using ENR_convex_closed by blast
       
  1176 
       
  1177 lemma AR_cball [simp]:
       
  1178     fixes a :: "'a::euclidean_space"
       
  1179     shows "AR(cball a r) \<longleftrightarrow> 0 \<le> r"
       
  1180   by (auto simp: AR_ANR convex_imp_contractible)
       
  1181 
       
  1182 lemma ANR_box [iff]:
       
  1183     fixes a :: "'a::euclidean_space"
       
  1184     shows "ANR(cbox a b)" "ANR(box a b)"
       
  1185   by (auto simp: convex_imp_ANR open_imp_ANR)
       
  1186 
       
  1187 lemma ENR_box [iff]:
       
  1188     fixes a :: "'a::euclidean_space"
       
  1189     shows "ENR(cbox a b)" "ENR(box a b)"
       
  1190 apply (simp add: ENR_convex_closed closed_cbox)
       
  1191 by (simp add: open_box open_imp_ENR)
       
  1192 
       
  1193 lemma AR_box [simp]:
       
  1194     "AR(cbox a b) \<longleftrightarrow> cbox a b \<noteq> {}" "AR(box a b) \<longleftrightarrow> box a b \<noteq> {}"
       
  1195   by (auto simp: AR_ANR convex_imp_contractible)
       
  1196 
       
  1197 lemma ANR_interior:
       
  1198      fixes S :: "'a::euclidean_space set"
       
  1199      shows "ANR(interior S)"
       
  1200   by (simp add: open_imp_ANR)
       
  1201 
       
  1202 lemma ENR_interior:
       
  1203      fixes S :: "'a::euclidean_space set"
       
  1204      shows "ENR(interior S)"
       
  1205   by (simp add: open_imp_ENR)
       
  1206 
       
  1207 lemma AR_imp_contractible:
       
  1208     fixes S :: "'a::euclidean_space set"
       
  1209     shows "AR S \<Longrightarrow> contractible S"
       
  1210   by (simp add: AR_ANR)
       
  1211 
       
  1212 lemma ENR_imp_locally_compact:
       
  1213     fixes S :: "'a::euclidean_space set"
       
  1214     shows "ENR S \<Longrightarrow> locally compact S"
       
  1215   by (simp add: ENR_ANR)
       
  1216 
       
  1217 lemma ANR_imp_locally_path_connected:
       
  1218   fixes S :: "'a::euclidean_space set"
       
  1219   assumes "ANR S"
       
  1220     shows "locally path_connected S"
       
  1221 proof -
       
  1222   obtain U and T :: "('a \<times> real) set"
       
  1223      where "convex U" "U \<noteq> {}"
       
  1224        and UT: "closedin (top_of_set U) T"
       
  1225        and "S homeomorphic T"
       
  1226     apply (rule homeomorphic_closedin_convex [of S])
       
  1227     using aff_dim_le_DIM [of S] apply auto
       
  1228     done
       
  1229   then have "locally path_connected T"
       
  1230     by (meson ANR_imp_absolute_neighbourhood_retract
       
  1231         assms convex_imp_locally_path_connected locally_open_subset retract_of_locally_path_connected)
       
  1232   then have S: "locally path_connected S"
       
  1233       if "openin (top_of_set U) V" "T retract_of V" "U \<noteq> {}" for V
       
  1234     using \<open>S homeomorphic T\<close> homeomorphic_locally homeomorphic_path_connectedness by blast
       
  1235   show ?thesis
       
  1236     using assms
       
  1237     apply (clarsimp simp: ANR_def)
       
  1238     apply (drule_tac x=U in spec)
       
  1239     apply (drule_tac x=T in spec)
       
  1240     using \<open>S homeomorphic T\<close> \<open>U \<noteq> {}\<close> UT  apply (blast intro: S)
       
  1241     done
       
  1242 qed
       
  1243 
       
  1244 lemma ANR_imp_locally_connected:
       
  1245   fixes S :: "'a::euclidean_space set"
       
  1246   assumes "ANR S"
       
  1247     shows "locally connected S"
       
  1248 using locally_path_connected_imp_locally_connected ANR_imp_locally_path_connected assms by auto
       
  1249 
       
  1250 lemma AR_imp_locally_path_connected:
       
  1251   fixes S :: "'a::euclidean_space set"
       
  1252   assumes "AR S"
       
  1253     shows "locally path_connected S"
       
  1254 by (simp add: ANR_imp_locally_path_connected AR_imp_ANR assms)
       
  1255 
       
  1256 lemma AR_imp_locally_connected:
       
  1257   fixes S :: "'a::euclidean_space set"
       
  1258   assumes "AR S"
       
  1259     shows "locally connected S"
       
  1260 using ANR_imp_locally_connected AR_ANR assms by blast
       
  1261 
       
  1262 lemma ENR_imp_locally_path_connected:
       
  1263   fixes S :: "'a::euclidean_space set"
       
  1264   assumes "ENR S"
       
  1265     shows "locally path_connected S"
       
  1266 by (simp add: ANR_imp_locally_path_connected ENR_imp_ANR assms)
       
  1267 
       
  1268 lemma ENR_imp_locally_connected:
       
  1269   fixes S :: "'a::euclidean_space set"
       
  1270   assumes "ENR S"
       
  1271     shows "locally connected S"
       
  1272 using ANR_imp_locally_connected ENR_ANR assms by blast
       
  1273 
       
  1274 lemma ANR_Times:
       
  1275   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  1276   assumes "ANR S" "ANR T" shows "ANR(S \<times> T)"
       
  1277 proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor)
       
  1278   fix f :: " ('a \<times> 'b) \<times> real \<Rightarrow> 'a \<times> 'b" and U C
       
  1279   assume "continuous_on C f" and fim: "f ` C \<subseteq> S \<times> T"
       
  1280      and cloUC: "closedin (top_of_set U) C"
       
  1281   have contf1: "continuous_on C (fst \<circ> f)"
       
  1282     by (simp add: \<open>continuous_on C f\<close> continuous_on_fst)
       
  1283   obtain W1 g where "C \<subseteq> W1"
       
  1284                and UW1: "openin (top_of_set U) W1"
       
  1285                and contg: "continuous_on W1 g"
       
  1286                and gim: "g ` W1 \<subseteq> S"
       
  1287                and geq: "\<And>x. x \<in> C \<Longrightarrow> g x = (fst \<circ> f) x"
       
  1288     apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> contf1 _ cloUC])
       
  1289     using fim apply auto
       
  1290     done
       
  1291   have contf2: "continuous_on C (snd \<circ> f)"
       
  1292     by (simp add: \<open>continuous_on C f\<close> continuous_on_snd)
       
  1293   obtain W2 h where "C \<subseteq> W2"
       
  1294                and UW2: "openin (top_of_set U) W2"
       
  1295                and conth: "continuous_on W2 h"
       
  1296                and him: "h ` W2 \<subseteq> T"
       
  1297                and heq: "\<And>x. x \<in> C \<Longrightarrow> h x = (snd \<circ> f) x"
       
  1298     apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> contf2 _ cloUC])
       
  1299     using fim apply auto
       
  1300     done
       
  1301   show "\<exists>V g. C \<subseteq> V \<and>
       
  1302                openin (top_of_set U) V \<and>
       
  1303                continuous_on V g \<and> g ` V \<subseteq> S \<times> T \<and> (\<forall>x\<in>C. g x = f x)"
       
  1304   proof (intro exI conjI)
       
  1305     show "C \<subseteq> W1 \<inter> W2"
       
  1306       by (simp add: \<open>C \<subseteq> W1\<close> \<open>C \<subseteq> W2\<close>)
       
  1307     show "openin (top_of_set U) (W1 \<inter> W2)"
       
  1308       by (simp add: UW1 UW2 openin_Int)
       
  1309     show  "continuous_on (W1 \<inter> W2) (\<lambda>x. (g x, h x))"
       
  1310       by (metis (no_types) contg conth continuous_on_Pair continuous_on_subset inf_commute inf_le1)
       
  1311     show  "(\<lambda>x. (g x, h x)) ` (W1 \<inter> W2) \<subseteq> S \<times> T"
       
  1312       using gim him by blast
       
  1313     show  "(\<forall>x\<in>C. (g x, h x) = f x)"
       
  1314       using geq heq by auto
       
  1315   qed
       
  1316 qed
       
  1317 
       
  1318 lemma AR_Times:
       
  1319   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  1320   assumes "AR S" "AR T" shows "AR(S \<times> T)"
       
  1321   using assms by (simp add: AR_ANR ANR_Times contractible_Times)
       
  1322 
       
  1323 subsection\<^marker>\<open>tag unimportant\<close>\<open>Retracts and intervals in ordered euclidean space\<close>
       
  1324 
       
  1325 lemma ANR_interval [iff]:
       
  1326   fixes a :: "'a::ordered_euclidean_space"
       
  1327   shows "ANR{a..b}"
       
  1328   by (simp add: interval_cbox)
       
  1329 
       
  1330 lemma ENR_interval [iff]:
       
  1331   fixes a :: "'a::ordered_euclidean_space"
       
  1332   shows "ENR{a..b}"
       
  1333   by (auto simp: interval_cbox)
       
  1334 
       
  1335 subsection \<open>More advanced properties of ANRs and ENRs\<close>
       
  1336 
       
  1337 lemma ENR_rel_frontier_convex:
       
  1338   fixes S :: "'a::euclidean_space set"
       
  1339   assumes "bounded S" "convex S"
       
  1340     shows "ENR(rel_frontier S)"
       
  1341 proof (cases "S = {}")
       
  1342   case True then show ?thesis
       
  1343     by simp
       
  1344 next
       
  1345   case False
       
  1346   with assms have "rel_interior S \<noteq> {}"
       
  1347     by (simp add: rel_interior_eq_empty)
       
  1348   then obtain a where a: "a \<in> rel_interior S"
       
  1349     by auto
       
  1350   have ahS: "affine hull S - {a} \<subseteq> {x. closest_point (affine hull S) x \<noteq> a}"
       
  1351     by (auto simp: closest_point_self)
       
  1352   have "rel_frontier S retract_of affine hull S - {a}"
       
  1353     by (simp add: assms a rel_frontier_retract_of_punctured_affine_hull)
       
  1354   also have "\<dots> retract_of {x. closest_point (affine hull S) x \<noteq> a}"
       
  1355     apply (simp add: retract_of_def retraction_def ahS)
       
  1356     apply (rule_tac x="closest_point (affine hull S)" in exI)
       
  1357     apply (auto simp: False closest_point_self affine_imp_convex closest_point_in_set continuous_on_closest_point)
       
  1358     done
       
  1359   finally have "rel_frontier S retract_of {x. closest_point (affine hull S) x \<noteq> a}" .
       
  1360   moreover have "openin (top_of_set UNIV) (UNIV \<inter> closest_point (affine hull S) -` (- {a}))"
       
  1361     apply (rule continuous_openin_preimage_gen)
       
  1362     apply (auto simp: False affine_imp_convex continuous_on_closest_point)
       
  1363     done
       
  1364   ultimately show ?thesis
       
  1365     unfolding ENR_def
       
  1366     apply (rule_tac x = "closest_point (affine hull S) -` (- {a})" in exI)
       
  1367     apply (simp add: vimage_def)
       
  1368     done
       
  1369 qed
       
  1370 
       
  1371 lemma ANR_rel_frontier_convex:
       
  1372                  fixes S :: "'a::euclidean_space set"
       
  1373   assumes "bounded S" "convex S"
       
  1374     shows "ANR(rel_frontier S)"
       
  1375 by (simp add: ENR_imp_ANR ENR_rel_frontier_convex assms)
       
  1376     
       
  1377 lemma ENR_closedin_Un_local:
       
  1378   fixes S :: "'a::euclidean_space set"
       
  1379   shows "\<lbrakk>ENR S; ENR T; ENR(S \<inter> T);
       
  1380           closedin (top_of_set (S \<union> T)) S; closedin (top_of_set (S \<union> T)) T\<rbrakk>
       
  1381         \<Longrightarrow> ENR(S \<union> T)"
       
  1382 by (simp add: ENR_ANR ANR_closed_Un_local locally_compact_closedin_Un)
       
  1383 
       
  1384 lemma ENR_closed_Un:
       
  1385   fixes S :: "'a::euclidean_space set"
       
  1386   shows "\<lbrakk>closed S; closed T; ENR S; ENR T; ENR(S \<inter> T)\<rbrakk> \<Longrightarrow> ENR(S \<union> T)"
       
  1387 by (auto simp: closed_subset ENR_closedin_Un_local)
       
  1388 
       
  1389 lemma absolute_retract_Un:
       
  1390   fixes S :: "'a::euclidean_space set"
       
  1391   shows "\<lbrakk>S retract_of UNIV; T retract_of UNIV; (S \<inter> T) retract_of UNIV\<rbrakk>
       
  1392          \<Longrightarrow> (S \<union> T) retract_of UNIV"
       
  1393   by (meson AR_closed_Un_local_aux closed_subset retract_of_UNIV retract_of_imp_subset)
       
  1394 
       
  1395 lemma retract_from_Un_Int:
       
  1396   fixes S :: "'a::euclidean_space set"
       
  1397   assumes clS: "closedin (top_of_set (S \<union> T)) S"
       
  1398       and clT: "closedin (top_of_set (S \<union> T)) T"
       
  1399       and Un: "(S \<union> T) retract_of U" and Int: "(S \<inter> T) retract_of T"
       
  1400     shows "S retract_of U"
       
  1401 proof -
       
  1402   obtain r where r: "continuous_on T r" "r ` T \<subseteq> S \<inter> T" "\<forall>x\<in>S \<inter> T. r x = x"
       
  1403     using Int by (auto simp: retraction_def retract_of_def)
       
  1404   have "S retract_of S \<union> T"
       
  1405     unfolding retraction_def retract_of_def
       
  1406   proof (intro exI conjI)
       
  1407     show "continuous_on (S \<union> T) (\<lambda>x. if x \<in> S then x else r x)"
       
  1408       apply (rule continuous_on_cases_local [OF clS clT])
       
  1409       using r by (auto simp: continuous_on_id)
       
  1410   qed (use r in auto)
       
  1411   also have "\<dots> retract_of U"
       
  1412     by (rule Un)
       
  1413   finally show ?thesis .
       
  1414 qed
       
  1415 
       
  1416 lemma AR_from_Un_Int_local:
       
  1417   fixes S :: "'a::euclidean_space set"
       
  1418   assumes clS: "closedin (top_of_set (S \<union> T)) S"
       
  1419       and clT: "closedin (top_of_set (S \<union> T)) T"
       
  1420       and Un: "AR(S \<union> T)" and Int: "AR(S \<inter> T)"
       
  1421     shows "AR S"
       
  1422   apply (rule AR_retract_of_AR [OF Un])
       
  1423   by (meson AR_imp_retract clS clT closedin_closed_subset local.Int retract_from_Un_Int retract_of_refl sup_ge2)
       
  1424 
       
  1425 lemma AR_from_Un_Int_local':
       
  1426   fixes S :: "'a::euclidean_space set"
       
  1427   assumes "closedin (top_of_set (S \<union> T)) S"
       
  1428       and "closedin (top_of_set (S \<union> T)) T"
       
  1429       and "AR(S \<union> T)" "AR(S \<inter> T)"
       
  1430     shows "AR T"
       
  1431   using AR_from_Un_Int_local [of T S] assms by (simp add: Un_commute Int_commute)
       
  1432 
       
  1433 lemma AR_from_Un_Int:
       
  1434   fixes S :: "'a::euclidean_space set"
       
  1435   assumes clo: "closed S" "closed T" and Un: "AR(S \<union> T)" and Int: "AR(S \<inter> T)"
       
  1436   shows "AR S"
       
  1437   by (metis AR_from_Un_Int_local [OF _ _ Un Int] Un_commute clo closed_closedin closedin_closed_subset inf_sup_absorb subtopology_UNIV top_greatest)
       
  1438 
       
  1439 lemma ANR_from_Un_Int_local:
       
  1440   fixes S :: "'a::euclidean_space set"
       
  1441   assumes clS: "closedin (top_of_set (S \<union> T)) S"
       
  1442       and clT: "closedin (top_of_set (S \<union> T)) T"
       
  1443       and Un: "ANR(S \<union> T)" and Int: "ANR(S \<inter> T)"
       
  1444     shows "ANR S"
       
  1445 proof -
       
  1446   obtain V where clo: "closedin (top_of_set (S \<union> T)) (S \<inter> T)"
       
  1447              and ope: "openin (top_of_set (S \<union> T)) V"
       
  1448              and ret: "S \<inter> T retract_of V"
       
  1449     using ANR_imp_neighbourhood_retract [OF Int] by (metis clS clT closedin_Int)
       
  1450   then obtain r where r: "continuous_on V r" and rim: "r ` V \<subseteq> S \<inter> T" and req: "\<forall>x\<in>S \<inter> T. r x = x"
       
  1451     by (auto simp: retraction_def retract_of_def)
       
  1452   have Vsub: "V \<subseteq> S \<union> T"
       
  1453     by (meson ope openin_contains_cball)
       
  1454   have Vsup: "S \<inter> T \<subseteq> V"
       
  1455     by (simp add: retract_of_imp_subset ret)
       
  1456   then have eq: "S \<union> V = ((S \<union> T) - T) \<union> V"
       
  1457     by auto
       
  1458   have eq': "S \<union> V = S \<union> (V \<inter> T)"
       
  1459     using Vsub by blast
       
  1460   have "continuous_on (S \<union> V \<inter> T) (\<lambda>x. if x \<in> S then x else r x)"
       
  1461   proof (rule continuous_on_cases_local)
       
  1462     show "closedin (top_of_set (S \<union> V \<inter> T)) S"
       
  1463       using clS closedin_subset_trans inf.boundedE by blast
       
  1464     show "closedin (top_of_set (S \<union> V \<inter> T)) (V \<inter> T)"
       
  1465       using clT Vsup by (auto simp: closedin_closed)
       
  1466     show "continuous_on (V \<inter> T) r"
       
  1467       by (meson Int_lower1 continuous_on_subset r)
       
  1468   qed (use req continuous_on_id in auto)
       
  1469   with rim have "S retract_of S \<union> V"
       
  1470     unfolding retraction_def retract_of_def
       
  1471     apply (rule_tac x="\<lambda>x. if x \<in> S then x else r x" in exI)
       
  1472     apply (auto simp: eq')
       
  1473     done
       
  1474   then show ?thesis
       
  1475     using ANR_neighborhood_retract [OF Un]
       
  1476     using \<open>S \<union> V = S \<union> T - T \<union> V\<close> clT ope by fastforce
       
  1477 qed
       
  1478 
       
  1479 lemma ANR_from_Un_Int:
       
  1480   fixes S :: "'a::euclidean_space set"
       
  1481   assumes clo: "closed S" "closed T" and Un: "ANR(S \<union> T)" and Int: "ANR(S \<inter> T)"
       
  1482   shows "ANR S"
       
  1483   by (metis ANR_from_Un_Int_local [OF _ _ Un Int] Un_commute clo closed_closedin closedin_closed_subset inf_sup_absorb subtopology_UNIV top_greatest)
       
  1484 
       
  1485 lemma ANR_finite_Union_convex_closed:
       
  1486   fixes \<T> :: "'a::euclidean_space set set"
       
  1487   assumes \<T>: "finite \<T>" and clo: "\<And>C. C \<in> \<T> \<Longrightarrow> closed C" and con: "\<And>C. C \<in> \<T> \<Longrightarrow> convex C"
       
  1488   shows "ANR(\<Union>\<T>)"
       
  1489 proof -
       
  1490   have "ANR(\<Union>\<T>)" if "card \<T> < n" for n
       
  1491   using assms that
       
  1492   proof (induction n arbitrary: \<T>)
       
  1493     case 0 then show ?case by simp
       
  1494   next
       
  1495     case (Suc n)
       
  1496     have "ANR(\<Union>\<U>)" if "finite \<U>" "\<U> \<subseteq> \<T>" for \<U>
       
  1497       using that
       
  1498     proof (induction \<U>)
       
  1499       case empty
       
  1500       then show ?case  by simp
       
  1501     next
       
  1502       case (insert C \<U>)
       
  1503       have "ANR (C \<union> \<Union>\<U>)"
       
  1504       proof (rule ANR_closed_Un)
       
  1505         show "ANR (C \<inter> \<Union>\<U>)"
       
  1506           unfolding Int_Union
       
  1507         proof (rule Suc)
       
  1508           show "finite ((\<inter>) C ` \<U>)"
       
  1509             by (simp add: insert.hyps(1))
       
  1510           show "\<And>Ca. Ca \<in> (\<inter>) C ` \<U> \<Longrightarrow> closed Ca"
       
  1511             by (metis (no_types, hide_lams) Suc.prems(2) closed_Int subsetD imageE insert.prems insertI1 insertI2)
       
  1512           show "\<And>Ca. Ca \<in> (\<inter>) C ` \<U> \<Longrightarrow> convex Ca"
       
  1513             by (metis (mono_tags, lifting) Suc.prems(3) convex_Int imageE insert.prems insert_subset subsetCE)
       
  1514           show "card ((\<inter>) C ` \<U>) < n"
       
  1515           proof -
       
  1516             have "card \<T> \<le> n"
       
  1517               by (meson Suc.prems(4) not_less not_less_eq)
       
  1518             then show ?thesis
       
  1519               by (metis Suc.prems(1) card_image_le card_seteq insert.hyps insert.prems insert_subset le_trans not_less)
       
  1520           qed
       
  1521         qed
       
  1522         show "closed (\<Union>\<U>)"
       
  1523           using Suc.prems(2) insert.hyps(1) insert.prems by blast
       
  1524       qed (use Suc.prems convex_imp_ANR insert.prems insert.IH in auto)
       
  1525       then show ?case
       
  1526         by simp
       
  1527     qed
       
  1528     then show ?case
       
  1529       using Suc.prems(1) by blast
       
  1530   qed
       
  1531   then show ?thesis
       
  1532     by blast
       
  1533 qed
       
  1534 
       
  1535 
       
  1536 lemma finite_imp_ANR:
       
  1537   fixes S :: "'a::euclidean_space set"
       
  1538   assumes "finite S"
       
  1539   shows "ANR S"
       
  1540 proof -
       
  1541   have "ANR(\<Union>x \<in> S. {x})"
       
  1542     by (blast intro: ANR_finite_Union_convex_closed assms)
       
  1543   then show ?thesis
       
  1544     by simp
       
  1545 qed
       
  1546 
       
  1547 lemma ANR_insert:
       
  1548   fixes S :: "'a::euclidean_space set"
       
  1549   assumes "ANR S" "closed S"
       
  1550   shows "ANR(insert a S)"
       
  1551   by (metis ANR_closed_Un ANR_empty ANR_singleton Diff_disjoint Diff_insert_absorb assms closed_singleton insert_absorb insert_is_Un)
       
  1552 
       
  1553 lemma ANR_path_component_ANR:
       
  1554   fixes S :: "'a::euclidean_space set"
       
  1555   shows "ANR S \<Longrightarrow> ANR(path_component_set S x)"
       
  1556   using ANR_imp_locally_path_connected ANR_openin openin_path_component_locally_path_connected by blast
       
  1557 
       
  1558 lemma ANR_connected_component_ANR:
       
  1559   fixes S :: "'a::euclidean_space set"
       
  1560   shows "ANR S \<Longrightarrow> ANR(connected_component_set S x)"
       
  1561   by (metis ANR_openin openin_connected_component_locally_connected ANR_imp_locally_connected)
       
  1562 
       
  1563 lemma ANR_component_ANR:
       
  1564   fixes S :: "'a::euclidean_space set"
       
  1565   assumes "ANR S" "c \<in> components S"
       
  1566   shows "ANR c"
       
  1567   by (metis ANR_connected_component_ANR assms componentsE)
       
  1568 
       
  1569 subsection\<open>Original ANR material, now for ENRs\<close>
       
  1570 
       
  1571 lemma ENR_bounded:
       
  1572   fixes S :: "'a::euclidean_space set"
       
  1573   assumes "bounded S"
       
  1574   shows "ENR S \<longleftrightarrow> (\<exists>U. open U \<and> bounded U \<and> S retract_of U)"
       
  1575          (is "?lhs = ?rhs")
       
  1576 proof
       
  1577   obtain r where "0 < r" and r: "S \<subseteq> ball 0 r"
       
  1578     using bounded_subset_ballD assms by blast
       
  1579   assume ?lhs
       
  1580   then show ?rhs
       
  1581     apply (clarsimp simp: ENR_def)
       
  1582     apply (rule_tac x="ball 0 r \<inter> U" in exI, auto)
       
  1583     using r retract_of_imp_subset retract_of_subset by fastforce
       
  1584 next
       
  1585   assume ?rhs
       
  1586   then show ?lhs
       
  1587     using ENR_def by blast
       
  1588 qed
       
  1589 
       
  1590 lemma absolute_retract_imp_AR_gen:
       
  1591   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
  1592   assumes "S retract_of T" "convex T" "T \<noteq> {}" "S homeomorphic S'" "closedin (top_of_set U) S'"
       
  1593   shows "S' retract_of U"
       
  1594 proof -
       
  1595   have "AR T"
       
  1596     by (simp add: assms convex_imp_AR)
       
  1597   then have "AR S"
       
  1598     using AR_retract_of_AR assms by auto
       
  1599   then show ?thesis
       
  1600     using assms AR_imp_absolute_retract by metis
       
  1601 qed
       
  1602 
       
  1603 lemma absolute_retract_imp_AR:
       
  1604   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
  1605   assumes "S retract_of UNIV" "S homeomorphic S'" "closed S'"
       
  1606   shows "S' retract_of UNIV"
       
  1607   using AR_imp_absolute_retract_UNIV assms retract_of_UNIV by blast
       
  1608 
       
  1609 lemma homeomorphic_compact_arness:
       
  1610   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
  1611   assumes "S homeomorphic S'"
       
  1612   shows "compact S \<and> S retract_of UNIV \<longleftrightarrow> compact S' \<and> S' retract_of UNIV"
       
  1613   using assms homeomorphic_compactness
       
  1614   apply auto
       
  1615    apply (meson assms compact_AR homeomorphic_AR_iff_AR homeomorphic_compactness)+
       
  1616   done
       
  1617 
       
  1618 lemma absolute_retract_from_Un_Int:
       
  1619   fixes S :: "'a::euclidean_space set"
       
  1620   assumes "(S \<union> T) retract_of UNIV" "(S \<inter> T) retract_of UNIV" "closed S" "closed T"
       
  1621   shows "S retract_of UNIV"
       
  1622   using AR_from_Un_Int assms retract_of_UNIV by auto
       
  1623 
       
  1624 lemma ENR_from_Un_Int_gen:
       
  1625   fixes S :: "'a::euclidean_space set"
       
  1626   assumes "closedin (top_of_set (S \<union> T)) S" "closedin (top_of_set (S \<union> T)) T" "ENR(S \<union> T)" "ENR(S \<inter> T)"
       
  1627   shows "ENR S"
       
  1628   apply (simp add: ENR_ANR)
       
  1629   using ANR_from_Un_Int_local ENR_ANR assms locally_compact_closedin by blast
       
  1630 
       
  1631 
       
  1632 lemma ENR_from_Un_Int:
       
  1633   fixes S :: "'a::euclidean_space set"
       
  1634   assumes "closed S" "closed T" "ENR(S \<union> T)" "ENR(S \<inter> T)"
       
  1635   shows "ENR S"
       
  1636   by (meson ENR_from_Un_Int_gen assms closed_subset sup_ge1 sup_ge2)
       
  1637 
       
  1638 
       
  1639 lemma ENR_finite_Union_convex_closed:
       
  1640   fixes \<T> :: "'a::euclidean_space set set"
       
  1641   assumes \<T>: "finite \<T>" and clo: "\<And>C. C \<in> \<T> \<Longrightarrow> closed C" and con: "\<And>C. C \<in> \<T> \<Longrightarrow> convex C"
       
  1642   shows "ENR(\<Union> \<T>)"
       
  1643   by (simp add: ENR_ANR ANR_finite_Union_convex_closed \<T> clo closed_Union closed_imp_locally_compact con)
       
  1644 
       
  1645 lemma finite_imp_ENR:
       
  1646   fixes S :: "'a::euclidean_space set"
       
  1647   shows "finite S \<Longrightarrow> ENR S"
       
  1648   by (simp add: ENR_ANR finite_imp_ANR finite_imp_closed closed_imp_locally_compact)
       
  1649 
       
  1650 lemma ENR_insert:
       
  1651   fixes S :: "'a::euclidean_space set"
       
  1652   assumes "closed S" "ENR S"
       
  1653   shows "ENR(insert a S)"
       
  1654 proof -
       
  1655   have "ENR ({a} \<union> S)"
       
  1656     by (metis ANR_insert ENR_ANR Un_commute Un_insert_right assms closed_imp_locally_compact closed_insert sup_bot_right)
       
  1657   then show ?thesis
       
  1658     by auto
       
  1659 qed
       
  1660 
       
  1661 lemma ENR_path_component_ENR:
       
  1662   fixes S :: "'a::euclidean_space set"
       
  1663   assumes "ENR S"
       
  1664   shows "ENR(path_component_set S x)"
       
  1665   by (metis ANR_imp_locally_path_connected ENR_empty ENR_imp_ANR ENR_openin assms
       
  1666             locally_path_connected_2 openin_subtopology_self path_component_eq_empty)
       
  1667 
       
  1668 (*UNUSED
       
  1669 lemma ENR_Times:
       
  1670   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  1671   assumes "ENR S" "ENR T" shows "ENR(S \<times> T)"
       
  1672 using assms apply (simp add: ENR_ANR ANR_Times)
       
  1673 thm locally_compact_Times
       
  1674 oops
       
  1675   SIMP_TAC[ENR_ANR; ANR_PCROSS; LOCALLY_COMPACT_PCROSS]);;
       
  1676 *)
       
  1677 
       
  1678 subsection\<open>Finally, spheres are ANRs and ENRs\<close>
       
  1679 
       
  1680 lemma absolute_retract_homeomorphic_convex_compact:
       
  1681   fixes S :: "'a::euclidean_space set" and U :: "'b::euclidean_space set"
       
  1682   assumes "S homeomorphic U" "S \<noteq> {}" "S \<subseteq> T" "convex U" "compact U"
       
  1683   shows "S retract_of T"
       
  1684   by (metis UNIV_I assms compact_AR convex_imp_AR homeomorphic_AR_iff_AR homeomorphic_compactness homeomorphic_empty(1) retract_of_subset subsetI)
       
  1685 
       
  1686 lemma frontier_retract_of_punctured_universe:
       
  1687   fixes S :: "'a::euclidean_space set"
       
  1688   assumes "convex S" "bounded S" "a \<in> interior S"
       
  1689   shows "(frontier S) retract_of (- {a})"
       
  1690   using rel_frontier_retract_of_punctured_affine_hull
       
  1691   by (metis Compl_eq_Diff_UNIV affine_hull_nonempty_interior assms empty_iff rel_frontier_frontier rel_interior_nonempty_interior)
       
  1692 
       
  1693 lemma sphere_retract_of_punctured_universe_gen:
       
  1694   fixes a :: "'a::euclidean_space"
       
  1695   assumes "b \<in> ball a r"
       
  1696   shows  "sphere a r retract_of (- {b})"
       
  1697 proof -
       
  1698   have "frontier (cball a r) retract_of (- {b})"
       
  1699     apply (rule frontier_retract_of_punctured_universe)
       
  1700     using assms by auto
       
  1701   then show ?thesis
       
  1702     by simp
       
  1703 qed
       
  1704 
       
  1705 lemma sphere_retract_of_punctured_universe:
       
  1706   fixes a :: "'a::euclidean_space"
       
  1707   assumes "0 < r"
       
  1708   shows "sphere a r retract_of (- {a})"
       
  1709   by (simp add: assms sphere_retract_of_punctured_universe_gen)
       
  1710 
       
  1711 lemma ENR_sphere:
       
  1712   fixes a :: "'a::euclidean_space"
       
  1713   shows "ENR(sphere a r)"
       
  1714 proof (cases "0 < r")
       
  1715   case True
       
  1716   then have "sphere a r retract_of -{a}"
       
  1717     by (simp add: sphere_retract_of_punctured_universe)
       
  1718   with open_delete show ?thesis
       
  1719     by (auto simp: ENR_def)
       
  1720 next
       
  1721   case False
       
  1722   then show ?thesis
       
  1723     using finite_imp_ENR
       
  1724     by (metis finite_insert infinite_imp_nonempty less_linear sphere_eq_empty sphere_trivial)
       
  1725 qed
       
  1726 
       
  1727 corollary\<^marker>\<open>tag unimportant\<close> ANR_sphere:
       
  1728   fixes a :: "'a::euclidean_space"
       
  1729   shows "ANR(sphere a r)"
       
  1730   by (simp add: ENR_imp_ANR ENR_sphere)
       
  1731 
       
  1732 subsection\<open>Spheres are connected, etc\<close>
       
  1733 
       
  1734 lemma locally_path_connected_sphere_gen:
       
  1735   fixes S :: "'a::euclidean_space set"
       
  1736   assumes "bounded S" and "convex S" 
       
  1737   shows "locally path_connected (rel_frontier S)"
       
  1738 proof (cases "rel_interior S = {}")
       
  1739   case True
       
  1740   with assms show ?thesis
       
  1741     by (simp add: rel_interior_eq_empty)
       
  1742 next
       
  1743   case False
       
  1744   then obtain a where a: "a \<in> rel_interior S"
       
  1745     by blast
       
  1746   show ?thesis
       
  1747   proof (rule retract_of_locally_path_connected)
       
  1748     show "locally path_connected (affine hull S - {a})"
       
  1749       by (meson convex_affine_hull convex_imp_locally_path_connected locally_open_subset openin_delete openin_subtopology_self)
       
  1750     show "rel_frontier S retract_of affine hull S - {a}"
       
  1751       using a assms rel_frontier_retract_of_punctured_affine_hull by blast
       
  1752   qed
       
  1753 qed
       
  1754 
       
  1755 lemma locally_connected_sphere_gen:
       
  1756   fixes S :: "'a::euclidean_space set"
       
  1757   assumes "bounded S" and "convex S" 
       
  1758   shows "locally connected (rel_frontier S)"
       
  1759   by (simp add: ANR_imp_locally_connected ANR_rel_frontier_convex assms)
       
  1760 
       
  1761 lemma locally_path_connected_sphere:
       
  1762   fixes a :: "'a::euclidean_space"
       
  1763   shows "locally path_connected (sphere a r)"
       
  1764   using ENR_imp_locally_path_connected ENR_sphere by blast
       
  1765 
       
  1766 lemma locally_connected_sphere:
       
  1767   fixes a :: "'a::euclidean_space"
       
  1768   shows "locally connected(sphere a r)"
       
  1769   using ANR_imp_locally_connected ANR_sphere by blast
       
  1770 
       
  1771 subsection\<open>Borsuk homotopy extension theorem\<close>
       
  1772 
       
  1773 text\<open>It's only this late so we can use the concept of retraction,
       
  1774   saying that the domain sets or range set are ENRs.\<close>
       
  1775 
       
  1776 theorem Borsuk_homotopy_extension_homotopic:
       
  1777   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
  1778   assumes cloTS: "closedin (top_of_set T) S"
       
  1779       and anr: "(ANR S \<and> ANR T) \<or> ANR U"
       
  1780       and contf: "continuous_on T f"
       
  1781       and "f ` T \<subseteq> U"
       
  1782       and "homotopic_with_canon (\<lambda>x. True) S U f g"
       
  1783    obtains g' where "homotopic_with_canon (\<lambda>x. True) T U f g'"
       
  1784                     "continuous_on T g'" "image g' T \<subseteq> U"
       
  1785                     "\<And>x. x \<in> S \<Longrightarrow> g' x = g x"
       
  1786 proof -
       
  1787   have "S \<subseteq> T" using assms closedin_imp_subset by blast
       
  1788   obtain h where conth: "continuous_on ({0..1} \<times> S) h"
       
  1789              and him: "h ` ({0..1} \<times> S) \<subseteq> U"
       
  1790              and [simp]: "\<And>x. h(0, x) = f x" "\<And>x. h(1::real, x) = g x"
       
  1791        using assms by (auto simp: homotopic_with_def)
       
  1792   define h' where "h' \<equiv>  \<lambda>z. if snd z \<in> S then h z else (f \<circ> snd) z"
       
  1793   define B where "B \<equiv> {0::real} \<times> T \<union> {0..1} \<times> S"
       
  1794   have clo0T: "closedin (top_of_set ({0..1} \<times> T)) ({0::real} \<times> T)"
       
  1795     by (simp add: Abstract_Topology.closedin_Times)
       
  1796   moreover have cloT1S: "closedin (top_of_set ({0..1} \<times> T)) ({0..1} \<times> S)"
       
  1797     by (simp add: Abstract_Topology.closedin_Times assms)
       
  1798   ultimately have clo0TB:"closedin (top_of_set ({0..1} \<times> T)) B"
       
  1799     by (auto simp: B_def)
       
  1800   have cloBS: "closedin (top_of_set B) ({0..1} \<times> S)"
       
  1801     by (metis (no_types) Un_subset_iff B_def closedin_subset_trans [OF cloT1S] clo0TB closedin_imp_subset closedin_self)
       
  1802   moreover have cloBT: "closedin (top_of_set B) ({0} \<times> T)"
       
  1803     using \<open>S \<subseteq> T\<close> closedin_subset_trans [OF clo0T]
       
  1804     by (metis B_def Un_upper1 clo0TB closedin_closed inf_le1)
       
  1805   moreover have "continuous_on ({0} \<times> T) (f \<circ> snd)"
       
  1806     apply (rule continuous_intros)+
       
  1807     apply (simp add: contf)
       
  1808     done
       
  1809   ultimately have conth': "continuous_on B h'"
       
  1810     apply (simp add: h'_def B_def Un_commute [of "{0} \<times> T"])
       
  1811     apply (auto intro!: continuous_on_cases_local conth)
       
  1812     done
       
  1813   have "image h' B \<subseteq> U"
       
  1814     using \<open>f ` T \<subseteq> U\<close> him by (auto simp: h'_def B_def)
       
  1815   obtain V k where "B \<subseteq> V" and opeTV: "openin (top_of_set ({0..1} \<times> T)) V"
       
  1816                and contk: "continuous_on V k" and kim: "k ` V \<subseteq> U"
       
  1817                and keq: "\<And>x. x \<in> B \<Longrightarrow> k x = h' x"
       
  1818   using anr
       
  1819   proof
       
  1820     assume ST: "ANR S \<and> ANR T"
       
  1821     have eq: "({0} \<times> T \<inter> {0..1} \<times> S) = {0::real} \<times> S"
       
  1822       using \<open>S \<subseteq> T\<close> by auto
       
  1823     have "ANR B"
       
  1824       apply (simp add: B_def)
       
  1825       apply (rule ANR_closed_Un_local)
       
  1826           apply (metis cloBT B_def)
       
  1827          apply (metis Un_commute cloBS B_def)
       
  1828         apply (simp_all add: ANR_Times convex_imp_ANR ANR_singleton ST eq)
       
  1829       done
       
  1830     note Vk = that
       
  1831     have *: thesis if "openin (top_of_set ({0..1::real} \<times> T)) V"
       
  1832                       "retraction V B r" for V r
       
  1833       using that
       
  1834       apply (clarsimp simp add: retraction_def)
       
  1835       apply (rule Vk [of V "h' \<circ> r"], assumption+)
       
  1836         apply (metis continuous_on_compose conth' continuous_on_subset) 
       
  1837       using \<open>h' ` B \<subseteq> U\<close> apply force+
       
  1838       done
       
  1839     show thesis
       
  1840         apply (rule ANR_imp_neighbourhood_retract [OF \<open>ANR B\<close> clo0TB])
       
  1841         apply (auto simp: ANR_Times ANR_singleton ST retract_of_def *)
       
  1842         done
       
  1843   next
       
  1844     assume "ANR U"
       
  1845     with ANR_imp_absolute_neighbourhood_extensor \<open>h' ` B \<subseteq> U\<close> clo0TB conth' that
       
  1846     show ?thesis by blast
       
  1847   qed
       
  1848   define S' where "S' \<equiv> {x. \<exists>u::real. u \<in> {0..1} \<and> (u, x::'a) \<in> {0..1} \<times> T - V}"
       
  1849   have "closedin (top_of_set T) S'"
       
  1850     unfolding S'_def
       
  1851     apply (rule closedin_compact_projection, blast)
       
  1852     using closedin_self opeTV by blast
       
  1853   have S'_def: "S' = {x. \<exists>u::real.  (u, x::'a) \<in> {0..1} \<times> T - V}"
       
  1854     by (auto simp: S'_def)
       
  1855   have cloTS': "closedin (top_of_set T) S'"
       
  1856     using S'_def \<open>closedin (top_of_set T) S'\<close> by blast
       
  1857   have "S \<inter> S' = {}"
       
  1858     using S'_def B_def \<open>B \<subseteq> V\<close> by force
       
  1859   obtain a :: "'a \<Rightarrow> real" where conta: "continuous_on T a"
       
  1860       and "\<And>x. x \<in> T \<Longrightarrow> a x \<in> closed_segment 1 0"
       
  1861       and a1: "\<And>x. x \<in> S \<Longrightarrow> a x = 1"
       
  1862       and a0: "\<And>x. x \<in> S' \<Longrightarrow> a x = 0"
       
  1863     apply (rule Urysohn_local [OF cloTS cloTS' \<open>S \<inter> S' = {}\<close>, of 1 0], blast)
       
  1864     done
       
  1865   then have ain: "\<And>x. x \<in> T \<Longrightarrow> a x \<in> {0..1}"
       
  1866     using closed_segment_eq_real_ivl by auto
       
  1867   have inV: "(u * a t, t) \<in> V" if "t \<in> T" "0 \<le> u" "u \<le> 1" for t u
       
  1868   proof (rule ccontr)
       
  1869     assume "(u * a t, t) \<notin> V"
       
  1870     with ain [OF \<open>t \<in> T\<close>] have "a t = 0"
       
  1871       apply simp
       
  1872       apply (rule a0)
       
  1873       by (metis (no_types, lifting) Diff_iff S'_def SigmaI atLeastAtMost_iff mem_Collect_eq mult_le_one mult_nonneg_nonneg that)
       
  1874     show False
       
  1875       using B_def \<open>(u * a t, t) \<notin> V\<close> \<open>B \<subseteq> V\<close> \<open>a t = 0\<close> that by auto
       
  1876   qed
       
  1877   show ?thesis
       
  1878   proof
       
  1879     show hom: "homotopic_with_canon (\<lambda>x. True) T U f (\<lambda>x. k (a x, x))"
       
  1880     proof (simp add: homotopic_with, intro exI conjI)
       
  1881       show "continuous_on ({0..1} \<times> T) (k \<circ> (\<lambda>z. (fst z *\<^sub>R (a \<circ> snd) z, snd z)))"
       
  1882         apply (intro continuous_on_compose continuous_intros)
       
  1883         apply (rule continuous_on_subset [OF conta], force)
       
  1884         apply (rule continuous_on_subset [OF contk])
       
  1885         apply (force intro: inV)
       
  1886         done
       
  1887       show "(k \<circ> (\<lambda>z. (fst z *\<^sub>R (a \<circ> snd) z, snd z))) ` ({0..1} \<times> T) \<subseteq> U"
       
  1888         using inV kim by auto
       
  1889       show "\<forall>x\<in>T. (k \<circ> (\<lambda>z. (fst z *\<^sub>R (a \<circ> snd) z, snd z))) (0, x) = f x"
       
  1890         by (simp add: B_def h'_def keq)
       
  1891       show "\<forall>x\<in>T. (k \<circ> (\<lambda>z. (fst z *\<^sub>R (a \<circ> snd) z, snd z))) (1, x) = k (a x, x)"
       
  1892         by auto
       
  1893     qed
       
  1894   show "continuous_on T (\<lambda>x. k (a x, x))"
       
  1895     using homotopic_with_imp_continuous_maps [OF hom] by auto
       
  1896   show "(\<lambda>x. k (a x, x)) ` T \<subseteq> U"
       
  1897   proof clarify
       
  1898     fix t
       
  1899     assume "t \<in> T"
       
  1900     show "k (a t, t) \<in> U"
       
  1901       by (metis \<open>t \<in> T\<close> image_subset_iff inV kim not_one_le_zero linear mult_cancel_right1)
       
  1902   qed
       
  1903   show "\<And>x. x \<in> S \<Longrightarrow> k (a x, x) = g x"
       
  1904     by (simp add: B_def a1 h'_def keq)
       
  1905   qed
       
  1906 qed
       
  1907 
       
  1908 
       
  1909 corollary\<^marker>\<open>tag unimportant\<close> nullhomotopic_into_ANR_extension:
       
  1910   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
  1911   assumes "closed S"
       
  1912       and contf: "continuous_on S f"
       
  1913       and "ANR T"
       
  1914       and fim: "f ` S \<subseteq> T"
       
  1915       and "S \<noteq> {}"
       
  1916    shows "(\<exists>c. homotopic_with_canon (\<lambda>x. True) S T f (\<lambda>x. c)) \<longleftrightarrow>
       
  1917           (\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> T \<and> (\<forall>x \<in> S. g x = f x))"
       
  1918        (is "?lhs = ?rhs")
       
  1919 proof
       
  1920   assume ?lhs
       
  1921   then obtain c where c: "homotopic_with_canon (\<lambda>x. True) S T (\<lambda>x. c) f"
       
  1922     by (blast intro: homotopic_with_symD)
       
  1923   have "closedin (top_of_set UNIV) S"
       
  1924     using \<open>closed S\<close> closed_closedin by fastforce
       
  1925   then obtain g where "continuous_on UNIV g" "range g \<subseteq> T"
       
  1926                       "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
       
  1927     apply (rule Borsuk_homotopy_extension_homotopic [OF _ _ continuous_on_const _ c, where T=UNIV])
       
  1928     using \<open>ANR T\<close> \<open>S \<noteq> {}\<close> c homotopic_with_imp_subset1 apply fastforce+
       
  1929     done
       
  1930   then show ?rhs by blast
       
  1931 next
       
  1932   assume ?rhs
       
  1933   then obtain g where "continuous_on UNIV g" "range g \<subseteq> T" "\<And>x. x\<in>S \<Longrightarrow> g x = f x"
       
  1934     by blast
       
  1935   then obtain c where "homotopic_with_canon (\<lambda>h. True) UNIV T g (\<lambda>x. c)"
       
  1936     using nullhomotopic_from_contractible [of UNIV g T] contractible_UNIV by blast
       
  1937   then have "homotopic_with_canon (\<lambda>x. True) S T g (\<lambda>x. c)"
       
  1938     by (simp add: homotopic_from_subtopology)
       
  1939   then show ?lhs
       
  1940     by (force elim: homotopic_with_eq [of _ _ _ g "\<lambda>x. c"] simp: \<open>\<And>x. x \<in> S \<Longrightarrow> g x = f x\<close>)
       
  1941 qed
       
  1942 
       
  1943 corollary\<^marker>\<open>tag unimportant\<close> nullhomotopic_into_rel_frontier_extension:
       
  1944   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
  1945   assumes "closed S"
       
  1946       and contf: "continuous_on S f"
       
  1947       and "convex T" "bounded T"
       
  1948       and fim: "f ` S \<subseteq> rel_frontier T"
       
  1949       and "S \<noteq> {}"
       
  1950    shows "(\<exists>c. homotopic_with_canon (\<lambda>x. True) S (rel_frontier T) f (\<lambda>x. c)) \<longleftrightarrow>
       
  1951           (\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> rel_frontier T \<and> (\<forall>x \<in> S. g x = f x))"
       
  1952 by (simp add: nullhomotopic_into_ANR_extension assms ANR_rel_frontier_convex)
       
  1953 
       
  1954 corollary\<^marker>\<open>tag unimportant\<close> nullhomotopic_into_sphere_extension:
       
  1955   fixes f :: "'a::euclidean_space \<Rightarrow> 'b :: euclidean_space"
       
  1956   assumes "closed S" and contf: "continuous_on S f"
       
  1957       and "S \<noteq> {}" and fim: "f ` S \<subseteq> sphere a r"
       
  1958     shows "((\<exists>c. homotopic_with_canon (\<lambda>x. True) S (sphere a r) f (\<lambda>x. c)) \<longleftrightarrow>
       
  1959            (\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> sphere a r \<and> (\<forall>x \<in> S. g x = f x)))"
       
  1960            (is "?lhs = ?rhs")
       
  1961 proof (cases "r = 0")
       
  1962   case True with fim show ?thesis
       
  1963     apply auto
       
  1964     using fim continuous_on_const apply fastforce
       
  1965     by (metis contf contractible_sing nullhomotopic_into_contractible)
       
  1966 next
       
  1967   case False
       
  1968   then have eq: "sphere a r = rel_frontier (cball a r)" by simp
       
  1969   show ?thesis
       
  1970     using fim unfolding eq
       
  1971     apply (rule nullhomotopic_into_rel_frontier_extension [OF \<open>closed S\<close> contf convex_cball bounded_cball])
       
  1972     apply (rule \<open>S \<noteq> {}\<close>)
       
  1973     done
       
  1974 qed
       
  1975 
       
  1976 proposition\<^marker>\<open>tag unimportant\<close> Borsuk_map_essential_bounded_component:
       
  1977   fixes a :: "'a :: euclidean_space"
       
  1978   assumes "compact S" and "a \<notin> S"
       
  1979    shows "bounded (connected_component_set (- S) a) \<longleftrightarrow>
       
  1980           \<not>(\<exists>c. homotopic_with_canon (\<lambda>x. True) S (sphere 0 1)
       
  1981                                (\<lambda>x. inverse(norm(x - a)) *\<^sub>R (x - a)) (\<lambda>x. c))"
       
  1982    (is "?lhs = ?rhs")
       
  1983 proof (cases "S = {}")
       
  1984   case True then show ?thesis
       
  1985     by simp
       
  1986 next
       
  1987   case False
       
  1988   have "closed S" "bounded S"
       
  1989     using \<open>compact S\<close> compact_eq_bounded_closed by auto
       
  1990   have s01: "(\<lambda>x. (x - a) /\<^sub>R norm (x - a)) ` S \<subseteq> sphere 0 1"
       
  1991     using \<open>a \<notin> S\<close>  by clarsimp (metis dist_eq_0_iff dist_norm mult.commute right_inverse)
       
  1992   have aincc: "a \<in> connected_component_set (- S) a"
       
  1993     by (simp add: \<open>a \<notin> S\<close>)
       
  1994   obtain r where "r>0" and r: "S \<subseteq> ball 0 r"
       
  1995     using bounded_subset_ballD \<open>bounded S\<close> by blast
       
  1996   have "\<not> ?rhs \<longleftrightarrow> \<not> ?lhs"
       
  1997   proof
       
  1998     assume notr: "\<not> ?rhs"
       
  1999     have nog: "\<nexists>g. continuous_on (S \<union> connected_component_set (- S) a) g \<and>
       
  2000                    g ` (S \<union> connected_component_set (- S) a) \<subseteq> sphere 0 1 \<and>
       
  2001                    (\<forall>x\<in>S. g x = (x - a) /\<^sub>R norm (x - a))"
       
  2002          if "bounded (connected_component_set (- S) a)"
       
  2003       apply (rule non_extensible_Borsuk_map [OF \<open>compact S\<close> componentsI _ aincc])
       
  2004       using  \<open>a \<notin> S\<close> that by auto
       
  2005     obtain g where "range g \<subseteq> sphere 0 1" "continuous_on UNIV g"
       
  2006                         "\<And>x. x \<in> S \<Longrightarrow> g x = (x - a) /\<^sub>R norm (x - a)"
       
  2007       using notr
       
  2008       by (auto simp: nullhomotopic_into_sphere_extension
       
  2009                  [OF \<open>closed S\<close> continuous_on_Borsuk_map [OF \<open>a \<notin> S\<close>] False s01])
       
  2010     with \<open>a \<notin> S\<close> show  "\<not> ?lhs"
       
  2011       apply (clarsimp simp: Borsuk_map_into_sphere [of a S, symmetric] dest!: nog)
       
  2012       apply (drule_tac x=g in spec)
       
  2013       using continuous_on_subset by fastforce 
       
  2014   next
       
  2015     assume "\<not> ?lhs"
       
  2016     then obtain b where b: "b \<in> connected_component_set (- S) a" and "r \<le> norm b"
       
  2017       using bounded_iff linear by blast
       
  2018     then have bnot: "b \<notin> ball 0 r"
       
  2019       by simp
       
  2020     have "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. (x - a) /\<^sub>R norm (x - a))
       
  2021                                                    (\<lambda>x. (x - b) /\<^sub>R norm (x - b))"
       
  2022       apply (rule Borsuk_maps_homotopic_in_path_component)
       
  2023       using \<open>closed S\<close> b open_Compl open_path_connected_component apply fastforce
       
  2024       done
       
  2025     moreover
       
  2026     obtain c where "homotopic_with_canon (\<lambda>x. True) (ball 0 r) (sphere 0 1)
       
  2027                                    (\<lambda>x. inverse (norm (x - b)) *\<^sub>R (x - b)) (\<lambda>x. c)"
       
  2028     proof (rule nullhomotopic_from_contractible)
       
  2029       show "contractible (ball (0::'a) r)"
       
  2030         by (metis convex_imp_contractible convex_ball)
       
  2031       show "continuous_on (ball 0 r) (\<lambda>x. inverse(norm (x - b)) *\<^sub>R (x - b))"
       
  2032         by (rule continuous_on_Borsuk_map [OF bnot])
       
  2033       show "(\<lambda>x. (x - b) /\<^sub>R norm (x - b)) ` ball 0 r \<subseteq> sphere 0 1"
       
  2034         using bnot Borsuk_map_into_sphere by blast
       
  2035     qed blast
       
  2036     ultimately have "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. (x - a) /\<^sub>R norm (x - a)) (\<lambda>x. c)"
       
  2037       by (meson homotopic_with_subset_left homotopic_with_trans r)
       
  2038     then show "\<not> ?rhs"
       
  2039       by blast
       
  2040   qed
       
  2041   then show ?thesis by blast
       
  2042 qed
       
  2043 
       
  2044 lemma homotopic_Borsuk_maps_in_bounded_component:
       
  2045   fixes a :: "'a :: euclidean_space"
       
  2046   assumes "compact S" and "a \<notin> S"and "b \<notin> S"
       
  2047       and boc: "bounded (connected_component_set (- S) a)"
       
  2048       and hom: "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1)
       
  2049                                (\<lambda>x. (x - a) /\<^sub>R norm (x - a))
       
  2050                                (\<lambda>x. (x - b) /\<^sub>R norm (x - b))"
       
  2051    shows "connected_component (- S) a b"
       
  2052 proof (rule ccontr)
       
  2053   assume notcc: "\<not> connected_component (- S) a b"
       
  2054   let ?T = "S \<union> connected_component_set (- S) a"
       
  2055   have "\<nexists>g. continuous_on (S \<union> connected_component_set (- S) a) g \<and>
       
  2056             g ` (S \<union> connected_component_set (- S) a) \<subseteq> sphere 0 1 \<and>
       
  2057             (\<forall>x\<in>S. g x = (x - a) /\<^sub>R norm (x - a))"
       
  2058     by (simp add: \<open>a \<notin> S\<close> componentsI non_extensible_Borsuk_map [OF \<open>compact S\<close> _ boc])
       
  2059   moreover obtain g where "continuous_on (S \<union> connected_component_set (- S) a) g"
       
  2060                           "g ` (S \<union> connected_component_set (- S) a) \<subseteq> sphere 0 1"
       
  2061                           "\<And>x. x \<in> S \<Longrightarrow> g x = (x - a) /\<^sub>R norm (x - a)"
       
  2062   proof (rule Borsuk_homotopy_extension_homotopic)
       
  2063     show "closedin (top_of_set ?T) S"
       
  2064       by (simp add: \<open>compact S\<close> closed_subset compact_imp_closed)
       
  2065     show "continuous_on ?T (\<lambda>x. (x - b) /\<^sub>R norm (x - b))"
       
  2066       by (simp add: \<open>b \<notin> S\<close> notcc continuous_on_Borsuk_map)
       
  2067     show "(\<lambda>x. (x - b) /\<^sub>R norm (x - b)) ` ?T \<subseteq> sphere 0 1"
       
  2068       by (simp add: \<open>b \<notin> S\<close> notcc Borsuk_map_into_sphere)
       
  2069     show "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1)
       
  2070              (\<lambda>x. (x - b) /\<^sub>R norm (x - b)) (\<lambda>x. (x - a) /\<^sub>R norm (x - a))"
       
  2071       by (simp add: hom homotopic_with_symD)
       
  2072     qed (auto simp: ANR_sphere intro: that)
       
  2073   ultimately show False by blast
       
  2074 qed
       
  2075 
       
  2076 
       
  2077 lemma Borsuk_maps_homotopic_in_connected_component_eq:
       
  2078   fixes a :: "'a :: euclidean_space"
       
  2079   assumes S: "compact S" "a \<notin> S" "b \<notin> S" and 2: "2 \<le> DIM('a)"
       
  2080     shows "(homotopic_with_canon (\<lambda>x. True) S (sphere 0 1)
       
  2081                    (\<lambda>x. (x - a) /\<^sub>R norm (x - a))
       
  2082                    (\<lambda>x. (x - b) /\<^sub>R norm (x - b)) \<longleftrightarrow>
       
  2083            connected_component (- S) a b)"
       
  2084          (is "?lhs = ?rhs")
       
  2085 proof
       
  2086   assume L: ?lhs
       
  2087   show ?rhs
       
  2088   proof (cases "bounded(connected_component_set (- S) a)")
       
  2089     case True
       
  2090     show ?thesis
       
  2091       by (rule homotopic_Borsuk_maps_in_bounded_component [OF S True L])
       
  2092   next
       
  2093     case not_bo_a: False
       
  2094     show ?thesis
       
  2095     proof (cases "bounded(connected_component_set (- S) b)")
       
  2096       case True
       
  2097       show ?thesis
       
  2098         using homotopic_Borsuk_maps_in_bounded_component [OF S]
       
  2099         by (simp add: L True assms connected_component_sym homotopic_Borsuk_maps_in_bounded_component homotopic_with_sym)
       
  2100     next
       
  2101       case False
       
  2102       then show ?thesis
       
  2103         using cobounded_unique_unbounded_component [of "-S" a b] \<open>compact S\<close> not_bo_a
       
  2104         by (auto simp: compact_eq_bounded_closed assms connected_component_eq_eq)
       
  2105     qed
       
  2106   qed
       
  2107 next
       
  2108   assume R: ?rhs
       
  2109   then have "path_component (- S) a b"
       
  2110     using assms(1) compact_eq_bounded_closed open_Compl open_path_connected_component_set by fastforce
       
  2111   then show ?lhs
       
  2112     by (simp add: Borsuk_maps_homotopic_in_path_component)
       
  2113 qed
       
  2114 
       
  2115 subsection\<open>More extension theorems\<close>
       
  2116 
       
  2117 lemma extension_from_clopen:
       
  2118   assumes ope: "openin (top_of_set S) T"
       
  2119       and clo: "closedin (top_of_set S) T"
       
  2120       and contf: "continuous_on T f" and fim: "f ` T \<subseteq> U" and null: "U = {} \<Longrightarrow> S = {}"
       
  2121  obtains g where "continuous_on S g" "g ` S \<subseteq> U" "\<And>x. x \<in> T \<Longrightarrow> g x = f x"
       
  2122 proof (cases "U = {}")
       
  2123   case True
       
  2124   then show ?thesis
       
  2125     by (simp add: null that)
       
  2126 next
       
  2127   case False
       
  2128   then obtain a where "a \<in> U"
       
  2129     by auto
       
  2130   let ?g = "\<lambda>x. if x \<in> T then f x else a"
       
  2131   have Seq: "S = T \<union> (S - T)"
       
  2132     using clo closedin_imp_subset by fastforce
       
  2133   show ?thesis
       
  2134   proof
       
  2135     have "continuous_on (T \<union> (S - T)) ?g"
       
  2136       apply (rule continuous_on_cases_local)
       
  2137       using Seq clo ope by (auto simp: contf continuous_on_const intro: continuous_on_cases_local)
       
  2138     with Seq show "continuous_on S ?g"
       
  2139       by metis
       
  2140     show "?g ` S \<subseteq> U"
       
  2141       using \<open>a \<in> U\<close> fim by auto
       
  2142     show "\<And>x. x \<in> T \<Longrightarrow> ?g x = f x"
       
  2143       by auto
       
  2144   qed
       
  2145 qed
       
  2146 
       
  2147 
       
  2148 lemma extension_from_component:
       
  2149   fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
       
  2150   assumes S: "locally connected S \<or> compact S" and "ANR U"
       
  2151      and C: "C \<in> components S" and contf: "continuous_on C f" and fim: "f ` C \<subseteq> U"
       
  2152  obtains g where "continuous_on S g" "g ` S \<subseteq> U" "\<And>x. x \<in> C \<Longrightarrow> g x = f x"
       
  2153 proof -
       
  2154   obtain T g where ope: "openin (top_of_set S) T"
       
  2155                and clo: "closedin (top_of_set S) T"
       
  2156                and "C \<subseteq> T" and contg: "continuous_on T g" and gim: "g ` T \<subseteq> U"
       
  2157                and gf: "\<And>x. x \<in> C \<Longrightarrow> g x = f x"
       
  2158     using S
       
  2159   proof
       
  2160     assume "locally connected S"
       
  2161     show ?thesis
       
  2162       by (metis C \<open>locally connected S\<close> openin_components_locally_connected closedin_component contf fim order_refl that)
       
  2163   next
       
  2164     assume "compact S"
       
  2165     then obtain W g where "C \<subseteq> W" and opeW: "openin (top_of_set S) W"
       
  2166                  and contg: "continuous_on W g"
       
  2167                  and gim: "g ` W \<subseteq> U" and gf: "\<And>x. x \<in> C \<Longrightarrow> g x = f x"
       
  2168       using ANR_imp_absolute_neighbourhood_extensor [of U C f S] C \<open>ANR U\<close> closedin_component contf fim by blast
       
  2169     then obtain V where "open V" and V: "W = S \<inter> V"
       
  2170       by (auto simp: openin_open)
       
  2171     moreover have "locally compact S"
       
  2172       by (simp add: \<open>compact S\<close> closed_imp_locally_compact compact_imp_closed)
       
  2173     ultimately obtain K where opeK: "openin (top_of_set S) K" and "compact K" "C \<subseteq> K" "K \<subseteq> V"
       
  2174       by (metis C Int_subset_iff \<open>C \<subseteq> W\<close> \<open>compact S\<close> compact_components Sura_Bura_clopen_subset)
       
  2175     show ?thesis
       
  2176     proof
       
  2177       show "closedin (top_of_set S) K"
       
  2178         by (meson \<open>compact K\<close> \<open>compact S\<close> closedin_compact_eq opeK openin_imp_subset)
       
  2179       show "continuous_on K g"
       
  2180         by (metis Int_subset_iff V \<open>K \<subseteq> V\<close> contg continuous_on_subset opeK openin_subtopology subset_eq)
       
  2181       show "g ` K \<subseteq> U"
       
  2182         using V \<open>K \<subseteq> V\<close> gim opeK openin_imp_subset by fastforce
       
  2183     qed (use opeK gf \<open>C \<subseteq> K\<close> in auto)
       
  2184   qed
       
  2185   obtain h where "continuous_on S h" "h ` S \<subseteq> U" "\<And>x. x \<in> T \<Longrightarrow> h x = g x"
       
  2186     using extension_from_clopen
       
  2187     by (metis C bot.extremum_uniqueI clo contg gim fim image_is_empty in_components_nonempty ope)
       
  2188   then show ?thesis
       
  2189     by (metis \<open>C \<subseteq> T\<close> gf subset_eq that)
       
  2190 qed
       
  2191 
       
  2192 
       
  2193 lemma tube_lemma:
       
  2194   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  2195   assumes "compact S" and S: "S \<noteq> {}" "(\<lambda>x. (x,a)) ` S \<subseteq> U" 
       
  2196       and ope: "openin (top_of_set (S \<times> T)) U"
       
  2197   obtains V where "openin (top_of_set T) V" "a \<in> V" "S \<times> V \<subseteq> U"
       
  2198 proof -
       
  2199   let ?W = "{y. \<exists>x. x \<in> S \<and> (x, y) \<in> (S \<times> T - U)}"
       
  2200   have "U \<subseteq> S \<times> T" "closedin (top_of_set (S \<times> T)) (S \<times> T - U)"
       
  2201     using ope by (auto simp: openin_closedin_eq)
       
  2202   then have "closedin (top_of_set T) ?W"
       
  2203     using \<open>compact S\<close> closedin_compact_projection by blast
       
  2204   moreover have "a \<in> T - ?W"
       
  2205     using \<open>U \<subseteq> S \<times> T\<close> S by auto
       
  2206   moreover have "S \<times> (T - ?W) \<subseteq> U"
       
  2207     by auto
       
  2208   ultimately show ?thesis
       
  2209     by (metis (no_types, lifting) Sigma_cong closedin_def that topspace_euclidean_subtopology)
       
  2210 qed
       
  2211 
       
  2212 lemma tube_lemma_gen:
       
  2213   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  2214   assumes "compact S" "S \<noteq> {}" "T \<subseteq> T'" "S \<times> T \<subseteq> U"
       
  2215       and ope: "openin (top_of_set (S \<times> T')) U"
       
  2216   obtains V where "openin (top_of_set T') V" "T \<subseteq> V" "S \<times> V \<subseteq> U"
       
  2217 proof -
       
  2218   have "\<And>x. x \<in> T \<Longrightarrow> \<exists>V. openin (top_of_set T') V \<and> x \<in> V \<and> S \<times> V \<subseteq> U"
       
  2219     using assms by (auto intro:  tube_lemma [OF \<open>compact S\<close>])
       
  2220   then obtain F where F: "\<And>x. x \<in> T \<Longrightarrow> openin (top_of_set T') (F x) \<and> x \<in> F x \<and> S \<times> F x \<subseteq> U"
       
  2221     by metis
       
  2222   show ?thesis
       
  2223   proof
       
  2224     show "openin (top_of_set T') (\<Union>(F ` T))"
       
  2225       using F by blast
       
  2226     show "T \<subseteq> \<Union>(F ` T)"
       
  2227       using F by blast
       
  2228     show "S \<times> \<Union>(F ` T) \<subseteq> U"
       
  2229       using F by auto
       
  2230   qed
       
  2231 qed
       
  2232 
       
  2233 proposition\<^marker>\<open>tag unimportant\<close> homotopic_neighbourhood_extension:
       
  2234   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
  2235   assumes contf: "continuous_on S f" and fim: "f ` S \<subseteq> U"
       
  2236       and contg: "continuous_on S g" and gim: "g ` S \<subseteq> U"
       
  2237       and clo: "closedin (top_of_set S) T"
       
  2238       and "ANR U" and hom: "homotopic_with_canon (\<lambda>x. True) T U f g"
       
  2239     obtains V where "T \<subseteq> V" "openin (top_of_set S) V"
       
  2240                     "homotopic_with_canon (\<lambda>x. True) V U f g"
       
  2241 proof -
       
  2242   have "T \<subseteq> S"
       
  2243     using clo closedin_imp_subset by blast
       
  2244   obtain h where conth: "continuous_on ({0..1::real} \<times> T) h"
       
  2245              and him: "h ` ({0..1} \<times> T) \<subseteq> U"
       
  2246              and h0: "\<And>x. h(0, x) = f x" and h1: "\<And>x. h(1, x) = g x"
       
  2247     using hom by (auto simp: homotopic_with_def)
       
  2248   define h' where "h' \<equiv> \<lambda>z. if fst z \<in> {0} then f(snd z)
       
  2249                              else if fst z \<in> {1} then g(snd z)
       
  2250                              else h z"
       
  2251   let ?S0 = "{0::real} \<times> S" and ?S1 = "{1::real} \<times> S"
       
  2252   have "continuous_on(?S0 \<union> (?S1 \<union> {0..1} \<times> T)) h'"
       
  2253     unfolding h'_def
       
  2254   proof (intro continuous_on_cases_local)
       
  2255     show "closedin (top_of_set (?S0 \<union> (?S1 \<union> {0..1} \<times> T))) ?S0"
       
  2256          "closedin (top_of_set (?S1 \<union> {0..1} \<times> T)) ?S1"
       
  2257       using \<open>T \<subseteq> S\<close> by (force intro: closedin_Times closedin_subset_trans [of "{0..1} \<times> S"])+
       
  2258     show "closedin (top_of_set (?S0 \<union> (?S1 \<union> {0..1} \<times> T))) (?S1 \<union> {0..1} \<times> T)"
       
  2259          "closedin (top_of_set (?S1 \<union> {0..1} \<times> T)) ({0..1} \<times> T)"
       
  2260       using \<open>T \<subseteq> S\<close> by (force intro: clo closedin_Times closedin_subset_trans [of "{0..1} \<times> S"])+
       
  2261     show "continuous_on (?S0) (\<lambda>x. f (snd x))"
       
  2262       by (intro continuous_intros continuous_on_compose2 [OF contf]) auto
       
  2263     show "continuous_on (?S1) (\<lambda>x. g (snd x))"
       
  2264       by (intro continuous_intros continuous_on_compose2 [OF contg]) auto
       
  2265   qed (use h0 h1 conth in auto)
       
  2266   then have "continuous_on ({0,1} \<times> S \<union> ({0..1} \<times> T)) h'"
       
  2267     by (metis Sigma_Un_distrib1 Un_assoc insert_is_Un) 
       
  2268   moreover have "h' ` ({0,1} \<times> S \<union> {0..1} \<times> T) \<subseteq> U"
       
  2269     using fim gim him \<open>T \<subseteq> S\<close> unfolding h'_def by force
       
  2270   moreover have "closedin (top_of_set ({0..1::real} \<times> S)) ({0,1} \<times> S \<union> {0..1::real} \<times> T)"
       
  2271     by (intro closedin_Times closedin_Un clo) (simp_all add: closed_subset)
       
  2272   ultimately
       
  2273   obtain W k where W: "({0,1} \<times> S) \<union> ({0..1} \<times> T) \<subseteq> W"
       
  2274                and opeW: "openin (top_of_set ({0..1} \<times> S)) W"
       
  2275                and contk: "continuous_on W k"
       
  2276                and kim: "k ` W \<subseteq> U"
       
  2277                and kh': "\<And>x. x \<in> ({0,1} \<times> S) \<union> ({0..1} \<times> T) \<Longrightarrow> k x = h' x"
       
  2278     by (metis ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR U\<close>, of "({0,1} \<times> S) \<union> ({0..1} \<times> T)" h' "{0..1} \<times> S"])
       
  2279   obtain T' where opeT': "openin (top_of_set S) T'" 
       
  2280               and "T \<subseteq> T'" and TW: "{0..1} \<times> T' \<subseteq> W"
       
  2281     using tube_lemma_gen [of "{0..1::real}" T S W] W \<open>T \<subseteq> S\<close> opeW by auto
       
  2282   moreover have "homotopic_with_canon (\<lambda>x. True) T' U f g"
       
  2283   proof (simp add: homotopic_with, intro exI conjI)
       
  2284     show "continuous_on ({0..1} \<times> T') k"
       
  2285       using TW continuous_on_subset contk by auto
       
  2286     show "k ` ({0..1} \<times> T') \<subseteq> U"
       
  2287       using TW kim by fastforce
       
  2288     have "T' \<subseteq> S"
       
  2289       by (meson opeT' subsetD openin_imp_subset)
       
  2290     then show "\<forall>x\<in>T'. k (0, x) = f x" "\<forall>x\<in>T'. k (1, x) = g x"
       
  2291       by (auto simp: kh' h'_def)
       
  2292   qed
       
  2293   ultimately show ?thesis
       
  2294     by (blast intro: that)
       
  2295 qed
       
  2296 
       
  2297 text\<open> Homotopy on a union of closed-open sets.\<close>
       
  2298 proposition\<^marker>\<open>tag unimportant\<close> homotopic_on_clopen_Union:
       
  2299   fixes \<F> :: "'a::euclidean_space set set"
       
  2300   assumes "\<And>S. S \<in> \<F> \<Longrightarrow> closedin (top_of_set (\<Union>\<F>)) S"
       
  2301       and "\<And>S. S \<in> \<F> \<Longrightarrow> openin (top_of_set (\<Union>\<F>)) S"
       
  2302       and "\<And>S. S \<in> \<F> \<Longrightarrow> homotopic_with_canon (\<lambda>x. True) S T f g"
       
  2303   shows "homotopic_with_canon (\<lambda>x. True) (\<Union>\<F>) T f g"
       
  2304 proof -
       
  2305   obtain \<V> where "\<V> \<subseteq> \<F>" "countable \<V>" and eqU: "\<Union>\<V> = \<Union>\<F>"
       
  2306     using Lindelof_openin assms by blast
       
  2307   show ?thesis
       
  2308   proof (cases "\<V> = {}")
       
  2309     case True
       
  2310     then show ?thesis
       
  2311       by (metis Union_empty eqU homotopic_with_canon_on_empty)
       
  2312   next
       
  2313     case False
       
  2314     then obtain V :: "nat \<Rightarrow> 'a set" where V: "range V = \<V>"
       
  2315       using range_from_nat_into \<open>countable \<V>\<close> by metis
       
  2316     with \<open>\<V> \<subseteq> \<F>\<close> have clo: "\<And>n. closedin (top_of_set (\<Union>\<F>)) (V n)"
       
  2317                   and ope: "\<And>n. openin (top_of_set (\<Union>\<F>)) (V n)"
       
  2318                   and hom: "\<And>n. homotopic_with_canon (\<lambda>x. True) (V n) T f g"
       
  2319       using assms by auto 
       
  2320     then obtain h where conth: "\<And>n. continuous_on ({0..1::real} \<times> V n) (h n)"
       
  2321                   and him: "\<And>n. h n ` ({0..1} \<times> V n) \<subseteq> T" 
       
  2322                   and h0: "\<And>n. \<And>x. x \<in> V n \<Longrightarrow> h n (0, x) = f x" 
       
  2323                   and h1: "\<And>n. \<And>x. x \<in> V n \<Longrightarrow> h n (1, x) = g x"
       
  2324       by (simp add: homotopic_with) metis
       
  2325     have wop: "b \<in> V x \<Longrightarrow> \<exists>k. b \<in> V k \<and> (\<forall>j<k. b \<notin> V j)" for b x
       
  2326         using nat_less_induct [where P = "\<lambda>i. b \<notin> V i"] by meson
       
  2327     obtain \<zeta> where cont: "continuous_on ({0..1} \<times> \<Union>(V ` UNIV)) \<zeta>"
       
  2328               and eq: "\<And>x i. \<lbrakk>x \<in> {0..1} \<times> \<Union>(V ` UNIV) \<inter>
       
  2329                                    {0..1} \<times> (V i - (\<Union>m<i. V m))\<rbrakk> \<Longrightarrow> \<zeta> x = h i x"
       
  2330     proof (rule pasting_lemma_exists)
       
  2331       let ?X = "top_of_set ({0..1::real} \<times> \<Union>(range V))"
       
  2332       show "topspace ?X \<subseteq> (\<Union>i. {0..1::real} \<times> (V i - (\<Union>m<i. V m)))"
       
  2333         by (force simp: Ball_def dest: wop)
       
  2334       show "openin (top_of_set ({0..1} \<times> \<Union>(V ` UNIV))) 
       
  2335                    ({0..1::real} \<times> (V i - (\<Union>m<i. V m)))" for i
       
  2336       proof (intro openin_Times openin_subtopology_self openin_diff)
       
  2337         show "openin (top_of_set (\<Union>(V ` UNIV))) (V i)"
       
  2338           using ope V eqU by auto
       
  2339         show "closedin (top_of_set (\<Union>(V ` UNIV))) (\<Union>m<i. V m)"
       
  2340           using V clo eqU by (force intro: closedin_Union)
       
  2341       qed
       
  2342       show "continuous_map (subtopology ?X ({0..1} \<times> (V i - \<Union> (V ` {..<i})))) euclidean (h i)"  for i
       
  2343         by (auto simp add: subtopology_subtopology intro!: continuous_on_subset [OF conth])
       
  2344       show "\<And>i j x. x \<in> topspace ?X \<inter> {0..1} \<times> (V i - (\<Union>m<i. V m)) \<inter> {0..1} \<times> (V j - (\<Union>m<j. V m))
       
  2345                     \<Longrightarrow> h i x = h j x"
       
  2346         by clarsimp (metis lessThan_iff linorder_neqE_nat)
       
  2347     qed auto
       
  2348     show ?thesis
       
  2349     proof (simp add: homotopic_with eqU [symmetric], intro exI conjI ballI)
       
  2350       show "continuous_on ({0..1} \<times> \<Union>\<V>) \<zeta>"
       
  2351         using V eqU by (blast intro!:  continuous_on_subset [OF cont])
       
  2352       show "\<zeta>` ({0..1} \<times> \<Union>\<V>) \<subseteq> T"
       
  2353       proof clarsimp
       
  2354         fix t :: real and y :: "'a" and X :: "'a set"
       
  2355         assume "y \<in> X" "X \<in> \<V>" and t: "0 \<le> t" "t \<le> 1"
       
  2356         then obtain k where "y \<in> V k" and j: "\<forall>j<k. y \<notin> V j"
       
  2357           by (metis image_iff V wop)
       
  2358         with him t show "\<zeta>(t, y) \<in> T"
       
  2359           by (subst eq) force+
       
  2360       qed
       
  2361       fix X y
       
  2362       assume "X \<in> \<V>" "y \<in> X"
       
  2363       then obtain k where "y \<in> V k" and j: "\<forall>j<k. y \<notin> V j"
       
  2364         by (metis image_iff V wop)
       
  2365       then show "\<zeta>(0, y) = f y" and "\<zeta>(1, y) = g y"
       
  2366         by (subst eq [where i=k]; force simp: h0 h1)+ 
       
  2367     qed
       
  2368   qed
       
  2369 qed
       
  2370 
       
  2371 lemma homotopic_on_components_eq:
       
  2372   fixes S :: "'a :: euclidean_space set" and T :: "'b :: euclidean_space set"
       
  2373   assumes S: "locally connected S \<or> compact S" and "ANR T"
       
  2374   shows "homotopic_with_canon (\<lambda>x. True) S T f g \<longleftrightarrow>
       
  2375          (continuous_on S f \<and> f ` S \<subseteq> T \<and> continuous_on S g \<and> g ` S \<subseteq> T) \<and>
       
  2376          (\<forall>C \<in> components S. homotopic_with_canon (\<lambda>x. True) C T f g)"
       
  2377     (is "?lhs \<longleftrightarrow> ?C \<and> ?rhs")
       
  2378 proof -
       
  2379   have "continuous_on S f" "f ` S \<subseteq> T" "continuous_on S g" "g ` S \<subseteq> T" if ?lhs
       
  2380     using homotopic_with_imp_continuous homotopic_with_imp_subset1 homotopic_with_imp_subset2 that by blast+
       
  2381   moreover have "?lhs \<longleftrightarrow> ?rhs"
       
  2382     if contf: "continuous_on S f" and fim: "f ` S \<subseteq> T" and contg: "continuous_on S g" and gim: "g ` S \<subseteq> T"
       
  2383   proof
       
  2384     assume ?lhs
       
  2385     with that show ?rhs
       
  2386       by (simp add: homotopic_with_subset_left in_components_subset)
       
  2387   next
       
  2388     assume R: ?rhs
       
  2389     have "\<exists>U. C \<subseteq> U \<and> closedin (top_of_set S) U \<and>
       
  2390               openin (top_of_set S) U \<and>
       
  2391               homotopic_with_canon (\<lambda>x. True) U T f g" if C: "C \<in> components S" for C
       
  2392     proof -
       
  2393       have "C \<subseteq> S"
       
  2394         by (simp add: in_components_subset that)
       
  2395       show ?thesis
       
  2396         using S
       
  2397       proof
       
  2398         assume "locally connected S"
       
  2399         show ?thesis
       
  2400         proof (intro exI conjI)
       
  2401           show "closedin (top_of_set S) C"
       
  2402             by (simp add: closedin_component that)
       
  2403           show "openin (top_of_set S) C"
       
  2404             by (simp add: \<open>locally connected S\<close> openin_components_locally_connected that)
       
  2405           show "homotopic_with_canon (\<lambda>x. True) C T f g"
       
  2406             by (simp add: R that)
       
  2407         qed auto
       
  2408       next
       
  2409         assume "compact S"
       
  2410         have hom: "homotopic_with_canon (\<lambda>x. True) C T f g"
       
  2411           using R that by blast
       
  2412         obtain U where "C \<subseteq> U" and opeU: "openin (top_of_set S) U"
       
  2413                   and hom: "homotopic_with_canon (\<lambda>x. True) U T f g"
       
  2414           using homotopic_neighbourhood_extension [OF contf fim contg gim _ \<open>ANR T\<close> hom]
       
  2415             \<open>C \<in> components S\<close> closedin_component by blast
       
  2416         then obtain V where "open V" and V: "U = S \<inter> V"
       
  2417           by (auto simp: openin_open)
       
  2418         moreover have "locally compact S"
       
  2419           by (simp add: \<open>compact S\<close> closed_imp_locally_compact compact_imp_closed)
       
  2420         ultimately obtain K where opeK: "openin (top_of_set S) K" and "compact K" "C \<subseteq> K" "K \<subseteq> V"
       
  2421           by (metis C Int_subset_iff Sura_Bura_clopen_subset \<open>C \<subseteq> U\<close> \<open>compact S\<close> compact_components)
       
  2422         show ?thesis
       
  2423         proof (intro exI conjI)
       
  2424           show "closedin (top_of_set S) K"
       
  2425             by (meson \<open>compact K\<close> \<open>compact S\<close> closedin_compact_eq opeK openin_imp_subset)
       
  2426           show "homotopic_with_canon (\<lambda>x. True) K T f g"
       
  2427             using V \<open>K \<subseteq> V\<close> hom homotopic_with_subset_left opeK openin_imp_subset by fastforce
       
  2428         qed (use opeK \<open>C \<subseteq> K\<close> in auto)
       
  2429       qed
       
  2430     qed
       
  2431     then obtain \<phi> where \<phi>: "\<And>C. C \<in> components S \<Longrightarrow> C \<subseteq> \<phi> C"
       
  2432                   and clo\<phi>: "\<And>C. C \<in> components S \<Longrightarrow> closedin (top_of_set S) (\<phi> C)"
       
  2433                   and ope\<phi>: "\<And>C. C \<in> components S \<Longrightarrow> openin (top_of_set S) (\<phi> C)"
       
  2434                   and hom\<phi>: "\<And>C. C \<in> components S \<Longrightarrow> homotopic_with_canon (\<lambda>x. True) (\<phi> C) T f g"
       
  2435       by metis
       
  2436     have Seq: "S = \<Union> (\<phi> ` components S)"
       
  2437     proof
       
  2438       show "S \<subseteq> \<Union> (\<phi> ` components S)"
       
  2439         by (metis Sup_mono Union_components \<phi> imageI)
       
  2440       show "\<Union> (\<phi> ` components S) \<subseteq> S"
       
  2441         using ope\<phi> openin_imp_subset by fastforce
       
  2442     qed
       
  2443     show ?lhs
       
  2444       apply (subst Seq)
       
  2445       apply (rule homotopic_on_clopen_Union)
       
  2446       using Seq clo\<phi> ope\<phi> hom\<phi> by auto
       
  2447   qed
       
  2448   ultimately show ?thesis by blast
       
  2449 qed
       
  2450 
       
  2451 
       
  2452 lemma cohomotopically_trivial_on_components:
       
  2453   fixes S :: "'a :: euclidean_space set" and T :: "'b :: euclidean_space set"
       
  2454   assumes S: "locally connected S \<or> compact S" and "ANR T"
       
  2455   shows
       
  2456    "(\<forall>f g. continuous_on S f \<longrightarrow> f ` S \<subseteq> T \<longrightarrow> continuous_on S g \<longrightarrow> g ` S \<subseteq> T \<longrightarrow>
       
  2457            homotopic_with_canon (\<lambda>x. True) S T f g)
       
  2458     \<longleftrightarrow>
       
  2459     (\<forall>C\<in>components S.
       
  2460         \<forall>f g. continuous_on C f \<longrightarrow> f ` C \<subseteq> T \<longrightarrow> continuous_on C g \<longrightarrow> g ` C \<subseteq> T \<longrightarrow>
       
  2461               homotopic_with_canon (\<lambda>x. True) C T f g)"
       
  2462      (is "?lhs = ?rhs")
       
  2463 proof
       
  2464   assume L[rule_format]: ?lhs
       
  2465   show ?rhs
       
  2466   proof clarify
       
  2467     fix C f g
       
  2468     assume contf: "continuous_on C f" and fim: "f ` C \<subseteq> T"
       
  2469        and contg: "continuous_on C g" and gim: "g ` C \<subseteq> T" and C: "C \<in> components S"
       
  2470     obtain f' where contf': "continuous_on S f'" and f'im: "f' ` S \<subseteq> T" and f'f: "\<And>x. x \<in> C \<Longrightarrow> f' x = f x"
       
  2471       using extension_from_component [OF S \<open>ANR T\<close> C contf fim] by metis
       
  2472     obtain g' where contg': "continuous_on S g'" and g'im: "g' ` S \<subseteq> T" and g'g: "\<And>x. x \<in> C \<Longrightarrow> g' x = g x"
       
  2473       using extension_from_component [OF S \<open>ANR T\<close> C contg gim] by metis
       
  2474     have "homotopic_with_canon (\<lambda>x. True) C T f' g'"
       
  2475       using L [OF contf' f'im contg' g'im] homotopic_with_subset_left C in_components_subset by fastforce
       
  2476     then show "homotopic_with_canon (\<lambda>x. True) C T f g"
       
  2477       using f'f g'g homotopic_with_eq by force
       
  2478   qed
       
  2479 next
       
  2480   assume R [rule_format]: ?rhs
       
  2481   show ?lhs
       
  2482   proof clarify
       
  2483     fix f g
       
  2484     assume contf: "continuous_on S f" and fim: "f ` S \<subseteq> T"
       
  2485       and contg: "continuous_on S g" and gim: "g ` S \<subseteq> T"
       
  2486     moreover have "homotopic_with_canon (\<lambda>x. True) C T f g" if "C \<in> components S" for C
       
  2487       using R [OF that]
       
  2488       by (meson contf contg continuous_on_subset fim gim image_mono in_components_subset order.trans that)
       
  2489     ultimately show "homotopic_with_canon (\<lambda>x. True) S T f g"
       
  2490       by (subst homotopic_on_components_eq [OF S \<open>ANR T\<close>]) auto
       
  2491   qed
       
  2492 qed
       
  2493 
       
  2494 subsection\<open>The complement of a set and path-connectedness\<close>
       
  2495 
       
  2496 text\<open>Complement in dimension N > 1 of set homeomorphic to any interval in
       
  2497  any dimension is (path-)connected. This naively generalizes the argument
       
  2498  in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer fixed point theorem",
       
  2499 American Mathematical Monthly 1984.\<close>
       
  2500 
       
  2501 lemma unbounded_components_complement_absolute_retract:
       
  2502   fixes S :: "'a::euclidean_space set"
       
  2503   assumes C: "C \<in> components(- S)" and S: "compact S" "AR S"
       
  2504     shows "\<not> bounded C"
       
  2505 proof -
       
  2506   obtain y where y: "C = connected_component_set (- S) y" and "y \<notin> S"
       
  2507     using C by (auto simp: components_def)
       
  2508   have "open(- S)"
       
  2509     using S by (simp add: closed_open compact_eq_bounded_closed)
       
  2510   have "S retract_of UNIV"
       
  2511     using S compact_AR by blast
       
  2512   then obtain r where contr: "continuous_on UNIV r" and ontor: "range r \<subseteq> S"
       
  2513                   and r: "\<And>x. x \<in> S \<Longrightarrow> r x = x"
       
  2514     by (auto simp: retract_of_def retraction_def)
       
  2515   show ?thesis
       
  2516   proof
       
  2517     assume "bounded C"
       
  2518     have "connected_component_set (- S) y \<subseteq> S"
       
  2519     proof (rule frontier_subset_retraction)
       
  2520       show "bounded (connected_component_set (- S) y)"
       
  2521         using \<open>bounded C\<close> y by blast
       
  2522       show "frontier (connected_component_set (- S) y) \<subseteq> S"
       
  2523         using C \<open>compact S\<close> compact_eq_bounded_closed frontier_of_components_closed_complement y by blast
       
  2524       show "continuous_on (closure (connected_component_set (- S) y)) r"
       
  2525         by (blast intro: continuous_on_subset [OF contr])
       
  2526     qed (use ontor r in auto)
       
  2527     with \<open>y \<notin> S\<close> show False by force
       
  2528   qed
       
  2529 qed
       
  2530 
       
  2531 lemma connected_complement_absolute_retract:
       
  2532   fixes S :: "'a::euclidean_space set"
       
  2533   assumes S: "compact S" "AR S" and 2: "2 \<le> DIM('a)"
       
  2534     shows "connected(- S)"
       
  2535 proof -
       
  2536   have "S retract_of UNIV"
       
  2537     using S compact_AR by blast
       
  2538   show ?thesis
       
  2539     apply (clarsimp simp: connected_iff_connected_component_eq)
       
  2540     apply (rule cobounded_unique_unbounded_component [OF _ 2])
       
  2541       apply (simp add: \<open>compact S\<close> compact_imp_bounded)
       
  2542      apply (meson ComplI S componentsI unbounded_components_complement_absolute_retract)+
       
  2543     done
       
  2544 qed
       
  2545 
       
  2546 lemma path_connected_complement_absolute_retract:
       
  2547   fixes S :: "'a::euclidean_space set"
       
  2548   assumes "compact S" "AR S" "2 \<le> DIM('a)"
       
  2549     shows "path_connected(- S)"
       
  2550   using connected_complement_absolute_retract [OF assms]
       
  2551   using \<open>compact S\<close> compact_eq_bounded_closed connected_open_path_connected by blast
       
  2552 
       
  2553 theorem connected_complement_homeomorphic_convex_compact:
       
  2554   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  2555   assumes hom: "S homeomorphic T" and T: "convex T" "compact T" and 2: "2 \<le> DIM('a)"
       
  2556     shows "connected(- S)"
       
  2557 proof (cases "S = {}")
       
  2558   case True
       
  2559   then show ?thesis
       
  2560     by (simp add: connected_UNIV)
       
  2561 next
       
  2562   case False
       
  2563   show ?thesis
       
  2564   proof (rule connected_complement_absolute_retract)
       
  2565     show "compact S"
       
  2566       using \<open>compact T\<close> hom homeomorphic_compactness by auto
       
  2567     show "AR S"
       
  2568       by (meson AR_ANR False \<open>convex T\<close> convex_imp_ANR convex_imp_contractible hom homeomorphic_ANR_iff_ANR homeomorphic_contractible_eq)
       
  2569   qed (rule 2)
       
  2570 qed
       
  2571 
       
  2572 corollary path_connected_complement_homeomorphic_convex_compact:
       
  2573   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  2574   assumes hom: "S homeomorphic T" "convex T" "compact T" "2 \<le> DIM('a)"
       
  2575     shows "path_connected(- S)"
       
  2576   using connected_complement_homeomorphic_convex_compact [OF assms]
       
  2577   using \<open>compact T\<close> compact_eq_bounded_closed connected_open_path_connected hom homeomorphic_compactness by blast
       
  2578 
       
  2579 lemma path_connected_complement_homeomorphic_interval:
       
  2580   fixes S :: "'a::euclidean_space set"
       
  2581   assumes "S homeomorphic cbox a b" "2 \<le> DIM('a)"
       
  2582   shows "path_connected(-S)"
       
  2583   using assms compact_cbox convex_box(1) path_connected_complement_homeomorphic_convex_compact by blast
       
  2584 
       
  2585 lemma connected_complement_homeomorphic_interval:
       
  2586   fixes S :: "'a::euclidean_space set"
       
  2587   assumes "S homeomorphic cbox a b" "2 \<le> DIM('a)"
       
  2588   shows "connected(-S)"
       
  2589   using assms path_connected_complement_homeomorphic_interval path_connected_imp_connected by blast
       
  2590 
       
  2591 end