src/HOL/Analysis/Brouwer_Fixpoint.thy
changeset 63627 6ddb43c6b711
parent 63540 f8652d0534fa
child 63928 d81fb5b46a5c
equal deleted inserted replaced
63626:44ce6b524ff3 63627:6ddb43c6b711
       
     1 (*  Author:     John Harrison
       
     2     Author:     Robert Himmelmann, TU Muenchen (Translation from HOL light) and LCP
       
     3 *)
       
     4 
       
     5 (* ========================================================================= *)
       
     6 (* Results connected with topological dimension.                             *)
       
     7 (*                                                                           *)
       
     8 (* At the moment this is just Brouwer's fixpoint theorem. The proof is from  *)
       
     9 (* Kuhn: "some combinatorial lemmas in topology", IBM J. v4. (1960) p. 518   *)
       
    10 (* See "http://www.research.ibm.com/journal/rd/045/ibmrd0405K.pdf".          *)
       
    11 (*                                                                           *)
       
    12 (* The script below is quite messy, but at least we avoid formalizing any    *)
       
    13 (* topological machinery; we don't even use barycentric subdivision; this is *)
       
    14 (* the big advantage of Kuhn's proof over the usual Sperner's lemma one.     *)
       
    15 (*                                                                           *)
       
    16 (*              (c) Copyright, John Harrison 1998-2008                       *)
       
    17 (* ========================================================================= *)
       
    18 
       
    19 section \<open>Results connected with topological dimension.\<close>
       
    20 
       
    21 theory Brouwer_Fixpoint
       
    22 imports Path_Connected Homeomorphism
       
    23 begin
       
    24 
       
    25 lemma bij_betw_singleton_eq:
       
    26   assumes f: "bij_betw f A B" and g: "bij_betw g A B" and a: "a \<in> A"
       
    27   assumes eq: "(\<And>x. x \<in> A \<Longrightarrow> x \<noteq> a \<Longrightarrow> f x = g x)"
       
    28   shows "f a = g a"
       
    29 proof -
       
    30   have "f ` (A - {a}) = g ` (A - {a})"
       
    31     by (intro image_cong) (simp_all add: eq)
       
    32   then have "B - {f a} = B - {g a}"
       
    33     using f g a  by (auto simp: bij_betw_def inj_on_image_set_diff set_eq_iff Diff_subset)
       
    34   moreover have "f a \<in> B" "g a \<in> B"
       
    35     using f g a by (auto simp: bij_betw_def)
       
    36   ultimately show ?thesis
       
    37     by auto
       
    38 qed
       
    39 
       
    40 lemma swap_image:
       
    41   "Fun.swap i j f ` A = (if i \<in> A then (if j \<in> A then f ` A else f ` ((A - {i}) \<union> {j}))
       
    42                                   else (if j \<in> A then f ` ((A - {j}) \<union> {i}) else f ` A))"
       
    43   apply (auto simp: Fun.swap_def image_iff)
       
    44   apply metis
       
    45   apply (metis member_remove remove_def)
       
    46   apply (metis member_remove remove_def)
       
    47   done
       
    48 
       
    49 lemmas swap_apply1 = swap_apply(1)
       
    50 lemmas swap_apply2 = swap_apply(2)
       
    51 lemmas lessThan_empty_iff = Iio_eq_empty_iff_nat
       
    52 lemmas Zero_notin_Suc = zero_notin_Suc_image
       
    53 lemmas atMost_Suc_eq_insert_0 = Iic_Suc_eq_insert_0
       
    54 
       
    55 lemma setsum_union_disjoint':
       
    56   assumes "finite A"
       
    57     and "finite B"
       
    58     and "A \<inter> B = {}"
       
    59     and "A \<union> B = C"
       
    60   shows "setsum g C = setsum g A + setsum g B"
       
    61   using setsum.union_disjoint[OF assms(1-3)] and assms(4) by auto
       
    62 
       
    63 lemma pointwise_minimal_pointwise_maximal:
       
    64   fixes s :: "(nat \<Rightarrow> nat) set"
       
    65   assumes "finite s"
       
    66     and "s \<noteq> {}"
       
    67     and "\<forall>x\<in>s. \<forall>y\<in>s. x \<le> y \<or> y \<le> x"
       
    68   shows "\<exists>a\<in>s. \<forall>x\<in>s. a \<le> x"
       
    69     and "\<exists>a\<in>s. \<forall>x\<in>s. x \<le> a"
       
    70   using assms
       
    71 proof (induct s rule: finite_ne_induct)
       
    72   case (insert b s)
       
    73   assume *: "\<forall>x\<in>insert b s. \<forall>y\<in>insert b s. x \<le> y \<or> y \<le> x"
       
    74   then obtain u l where "l \<in> s" "\<forall>b\<in>s. l \<le> b" "u \<in> s" "\<forall>b\<in>s. b \<le> u"
       
    75     using insert by auto
       
    76   with * show "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. a \<le> x" "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. x \<le> a"
       
    77     using *[rule_format, of b u] *[rule_format, of b l] by (metis insert_iff order.trans)+
       
    78 qed auto
       
    79 
       
    80 lemma brouwer_compactness_lemma:
       
    81   fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
       
    82   assumes "compact s"
       
    83     and "continuous_on s f"
       
    84     and "\<not> (\<exists>x\<in>s. f x = 0)"
       
    85   obtains d where "0 < d" and "\<forall>x\<in>s. d \<le> norm (f x)"
       
    86 proof (cases "s = {}")
       
    87   case True
       
    88   show thesis
       
    89     by (rule that [of 1]) (auto simp: True)
       
    90 next
       
    91   case False
       
    92   have "continuous_on s (norm \<circ> f)"
       
    93     by (rule continuous_intros continuous_on_norm assms(2))+
       
    94   with False obtain x where x: "x \<in> s" "\<forall>y\<in>s. (norm \<circ> f) x \<le> (norm \<circ> f) y"
       
    95     using continuous_attains_inf[OF assms(1), of "norm \<circ> f"]
       
    96     unfolding o_def
       
    97     by auto
       
    98   have "(norm \<circ> f) x > 0"
       
    99     using assms(3) and x(1)
       
   100     by auto
       
   101   then show ?thesis
       
   102     by (rule that) (insert x(2), auto simp: o_def)
       
   103 qed
       
   104 
       
   105 lemma kuhn_labelling_lemma:
       
   106   fixes P Q :: "'a::euclidean_space \<Rightarrow> bool"
       
   107   assumes "\<forall>x. P x \<longrightarrow> P (f x)"
       
   108     and "\<forall>x. P x \<longrightarrow> (\<forall>i\<in>Basis. Q i \<longrightarrow> 0 \<le> x\<bullet>i \<and> x\<bullet>i \<le> 1)"
       
   109   shows "\<exists>l. (\<forall>x.\<forall>i\<in>Basis. l x i \<le> (1::nat)) \<and>
       
   110              (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (x\<bullet>i = 0) \<longrightarrow> (l x i = 0)) \<and>
       
   111              (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (x\<bullet>i = 1) \<longrightarrow> (l x i = 1)) \<and>
       
   112              (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (l x i = 0) \<longrightarrow> x\<bullet>i \<le> f x\<bullet>i) \<and>
       
   113              (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (l x i = 1) \<longrightarrow> f x\<bullet>i \<le> x\<bullet>i)"
       
   114 proof -
       
   115   { fix x i
       
   116     let ?R = "\<lambda>y. (P x \<and> Q i \<and> x \<bullet> i = 0 \<longrightarrow> y = (0::nat)) \<and>
       
   117         (P x \<and> Q i \<and> x \<bullet> i = 1 \<longrightarrow> y = 1) \<and>
       
   118         (P x \<and> Q i \<and> y = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i) \<and>
       
   119         (P x \<and> Q i \<and> y = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i)"
       
   120     { assume "P x" "Q i" "i \<in> Basis" with assms have "0 \<le> f x \<bullet> i \<and> f x \<bullet> i \<le> 1" by auto }
       
   121     then have "i \<in> Basis \<Longrightarrow> ?R 0 \<or> ?R 1" by auto }
       
   122   then show ?thesis
       
   123     unfolding all_conj_distrib[symmetric] Ball_def (* FIXME: shouldn't this work by metis? *)
       
   124     by (subst choice_iff[symmetric])+ blast
       
   125 qed
       
   126 
       
   127 
       
   128 subsection \<open>The key "counting" observation, somewhat abstracted.\<close>
       
   129 
       
   130 lemma kuhn_counting_lemma:
       
   131   fixes bnd compo compo' face S F
       
   132   defines "nF s == card {f\<in>F. face f s \<and> compo' f}"
       
   133   assumes [simp, intro]: "finite F" \<comment> "faces" and [simp, intro]: "finite S" \<comment> "simplices"
       
   134     and "\<And>f. f \<in> F \<Longrightarrow> bnd f \<Longrightarrow> card {s\<in>S. face f s} = 1"
       
   135     and "\<And>f. f \<in> F \<Longrightarrow> \<not> bnd f \<Longrightarrow> card {s\<in>S. face f s} = 2"
       
   136     and "\<And>s. s \<in> S \<Longrightarrow> compo s \<Longrightarrow> nF s = 1"
       
   137     and "\<And>s. s \<in> S \<Longrightarrow> \<not> compo s \<Longrightarrow> nF s = 0 \<or> nF s = 2"
       
   138     and "odd (card {f\<in>F. compo' f \<and> bnd f})"
       
   139   shows "odd (card {s\<in>S. compo s})"
       
   140 proof -
       
   141   have "(\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) + (\<Sum>s | s \<in> S \<and> compo s. nF s) = (\<Sum>s\<in>S. nF s)"
       
   142     by (subst setsum.union_disjoint[symmetric]) (auto intro!: setsum.cong)
       
   143   also have "\<dots> = (\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> bnd f}. face f s}) +
       
   144                   (\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> \<not> bnd f}. face f s})"
       
   145     unfolding setsum.distrib[symmetric]
       
   146     by (subst card_Un_disjoint[symmetric])
       
   147        (auto simp: nF_def intro!: setsum.cong arg_cong[where f=card])
       
   148   also have "\<dots> = 1 * card {f\<in>F. compo' f \<and> bnd f} + 2 * card {f\<in>F. compo' f \<and> \<not> bnd f}"
       
   149     using assms(4,5) by (fastforce intro!: arg_cong2[where f="op +"] setsum_multicount)
       
   150   finally have "odd ((\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) + card {s\<in>S. compo s})"
       
   151     using assms(6,8) by simp
       
   152   moreover have "(\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) =
       
   153     (\<Sum>s | s \<in> S \<and> \<not> compo s \<and> nF s = 0. nF s) + (\<Sum>s | s \<in> S \<and> \<not> compo s \<and> nF s = 2. nF s)"
       
   154     using assms(7) by (subst setsum.union_disjoint[symmetric]) (fastforce intro!: setsum.cong)+
       
   155   ultimately show ?thesis
       
   156     by auto
       
   157 qed
       
   158 
       
   159 subsection \<open>The odd/even result for faces of complete vertices, generalized.\<close>
       
   160 
       
   161 lemma kuhn_complete_lemma:
       
   162   assumes [simp]: "finite simplices"
       
   163     and face: "\<And>f s. face f s \<longleftrightarrow> (\<exists>a\<in>s. f = s - {a})"
       
   164     and card_s[simp]:  "\<And>s. s \<in> simplices \<Longrightarrow> card s = n + 2"
       
   165     and rl_bd: "\<And>s. s \<in> simplices \<Longrightarrow> rl ` s \<subseteq> {..Suc n}"
       
   166     and bnd: "\<And>f s. s \<in> simplices \<Longrightarrow> face f s \<Longrightarrow> bnd f \<Longrightarrow> card {s\<in>simplices. face f s} = 1"
       
   167     and nbnd: "\<And>f s. s \<in> simplices \<Longrightarrow> face f s \<Longrightarrow> \<not> bnd f \<Longrightarrow> card {s\<in>simplices. face f s} = 2"
       
   168     and odd_card: "odd (card {f. (\<exists>s\<in>simplices. face f s) \<and> rl ` f = {..n} \<and> bnd f})"
       
   169   shows "odd (card {s\<in>simplices. (rl ` s = {..Suc n})})"
       
   170 proof (rule kuhn_counting_lemma)
       
   171   have finite_s[simp]: "\<And>s. s \<in> simplices \<Longrightarrow> finite s"
       
   172     by (metis add_is_0 zero_neq_numeral card_infinite assms(3))
       
   173 
       
   174   let ?F = "{f. \<exists>s\<in>simplices. face f s}"
       
   175   have F_eq: "?F = (\<Union>s\<in>simplices. \<Union>a\<in>s. {s - {a}})"
       
   176     by (auto simp: face)
       
   177   show "finite ?F"
       
   178     using \<open>finite simplices\<close> unfolding F_eq by auto
       
   179 
       
   180   show "card {s \<in> simplices. face f s} = 1" if "f \<in> ?F" "bnd f" for f
       
   181     using bnd that by auto
       
   182 
       
   183   show "card {s \<in> simplices. face f s} = 2" if "f \<in> ?F" "\<not> bnd f" for f
       
   184     using nbnd that by auto
       
   185 
       
   186   show "odd (card {f \<in> {f. \<exists>s\<in>simplices. face f s}. rl ` f = {..n} \<and> bnd f})"
       
   187     using odd_card by simp
       
   188 
       
   189   fix s assume s[simp]: "s \<in> simplices"
       
   190   let ?S = "{f \<in> {f. \<exists>s\<in>simplices. face f s}. face f s \<and> rl ` f = {..n}}"
       
   191   have "?S = (\<lambda>a. s - {a}) ` {a\<in>s. rl ` (s - {a}) = {..n}}"
       
   192     using s by (fastforce simp: face)
       
   193   then have card_S: "card ?S = card {a\<in>s. rl ` (s - {a}) = {..n}}"
       
   194     by (auto intro!: card_image inj_onI)
       
   195 
       
   196   { assume rl: "rl ` s = {..Suc n}"
       
   197     then have inj_rl: "inj_on rl s"
       
   198       by (intro eq_card_imp_inj_on) auto
       
   199     moreover obtain a where "rl a = Suc n" "a \<in> s"
       
   200       by (metis atMost_iff image_iff le_Suc_eq rl)
       
   201     ultimately have n: "{..n} = rl ` (s - {a})"
       
   202       by (auto simp add: inj_on_image_set_diff Diff_subset rl)
       
   203     have "{a\<in>s. rl ` (s - {a}) = {..n}} = {a}"
       
   204       using inj_rl \<open>a \<in> s\<close> by (auto simp add: n inj_on_image_eq_iff[OF inj_rl] Diff_subset)
       
   205     then show "card ?S = 1"
       
   206       unfolding card_S by simp }
       
   207 
       
   208   { assume rl: "rl ` s \<noteq> {..Suc n}"
       
   209     show "card ?S = 0 \<or> card ?S = 2"
       
   210     proof cases
       
   211       assume *: "{..n} \<subseteq> rl ` s"
       
   212       with rl rl_bd[OF s] have rl_s: "rl ` s = {..n}"
       
   213         by (auto simp add: atMost_Suc subset_insert_iff split: if_split_asm)
       
   214       then have "\<not> inj_on rl s"
       
   215         by (intro pigeonhole) simp
       
   216       then obtain a b where ab: "a \<in> s" "b \<in> s" "rl a = rl b" "a \<noteq> b"
       
   217         by (auto simp: inj_on_def)
       
   218       then have eq: "rl ` (s - {a}) = rl ` s"
       
   219         by auto
       
   220       with ab have inj: "inj_on rl (s - {a})"
       
   221         by (intro eq_card_imp_inj_on) (auto simp add: rl_s card_Diff_singleton_if)
       
   222 
       
   223       { fix x assume "x \<in> s" "x \<notin> {a, b}"
       
   224         then have "rl ` s - {rl x} = rl ` ((s - {a}) - {x})"
       
   225           by (auto simp: eq Diff_subset inj_on_image_set_diff[OF inj])
       
   226         also have "\<dots> = rl ` (s - {x})"
       
   227           using ab \<open>x \<notin> {a, b}\<close> by auto
       
   228         also assume "\<dots> = rl ` s"
       
   229         finally have False
       
   230           using \<open>x\<in>s\<close> by auto }
       
   231       moreover
       
   232       { fix x assume "x \<in> {a, b}" with ab have "x \<in> s \<and> rl ` (s - {x}) = rl ` s"
       
   233           by (simp add: set_eq_iff image_iff Bex_def) metis }
       
   234       ultimately have "{a\<in>s. rl ` (s - {a}) = {..n}} = {a, b}"
       
   235         unfolding rl_s[symmetric] by fastforce
       
   236       with \<open>a \<noteq> b\<close> show "card ?S = 0 \<or> card ?S = 2"
       
   237         unfolding card_S by simp
       
   238     next
       
   239       assume "\<not> {..n} \<subseteq> rl ` s"
       
   240       then have "\<And>x. rl ` (s - {x}) \<noteq> {..n}"
       
   241         by auto
       
   242       then show "card ?S = 0 \<or> card ?S = 2"
       
   243         unfolding card_S by simp
       
   244     qed }
       
   245 qed fact
       
   246 
       
   247 locale kuhn_simplex =
       
   248   fixes p n and base upd and s :: "(nat \<Rightarrow> nat) set"
       
   249   assumes base: "base \<in> {..< n} \<rightarrow> {..< p}"
       
   250   assumes base_out: "\<And>i. n \<le> i \<Longrightarrow> base i = p"
       
   251   assumes upd: "bij_betw upd {..< n} {..< n}"
       
   252   assumes s_pre: "s = (\<lambda>i j. if j \<in> upd`{..< i} then Suc (base j) else base j) ` {.. n}"
       
   253 begin
       
   254 
       
   255 definition "enum i j = (if j \<in> upd`{..< i} then Suc (base j) else base j)"
       
   256 
       
   257 lemma s_eq: "s = enum ` {.. n}"
       
   258   unfolding s_pre enum_def[abs_def] ..
       
   259 
       
   260 lemma upd_space: "i < n \<Longrightarrow> upd i < n"
       
   261   using upd by (auto dest!: bij_betwE)
       
   262 
       
   263 lemma s_space: "s \<subseteq> {..< n} \<rightarrow> {.. p}"
       
   264 proof -
       
   265   { fix i assume "i \<le> n" then have "enum i \<in> {..< n} \<rightarrow> {.. p}"
       
   266     proof (induct i)
       
   267       case 0 then show ?case
       
   268         using base by (auto simp: Pi_iff less_imp_le enum_def)
       
   269     next
       
   270       case (Suc i) with base show ?case
       
   271         by (auto simp: Pi_iff Suc_le_eq less_imp_le enum_def intro: upd_space)
       
   272     qed }
       
   273   then show ?thesis
       
   274     by (auto simp: s_eq)
       
   275 qed
       
   276 
       
   277 lemma inj_upd: "inj_on upd {..< n}"
       
   278   using upd by (simp add: bij_betw_def)
       
   279 
       
   280 lemma inj_enum: "inj_on enum {.. n}"
       
   281 proof -
       
   282   { fix x y :: nat assume "x \<noteq> y" "x \<le> n" "y \<le> n"
       
   283     with upd have "upd ` {..< x} \<noteq> upd ` {..< y}"
       
   284       by (subst inj_on_image_eq_iff[where C="{..< n}"]) (auto simp: bij_betw_def)
       
   285     then have "enum x \<noteq> enum y"
       
   286       by (auto simp add: enum_def fun_eq_iff) }
       
   287   then show ?thesis
       
   288     by (auto simp: inj_on_def)
       
   289 qed
       
   290 
       
   291 lemma enum_0: "enum 0 = base"
       
   292   by (simp add: enum_def[abs_def])
       
   293 
       
   294 lemma base_in_s: "base \<in> s"
       
   295   unfolding s_eq by (subst enum_0[symmetric]) auto
       
   296 
       
   297 lemma enum_in: "i \<le> n \<Longrightarrow> enum i \<in> s"
       
   298   unfolding s_eq by auto
       
   299 
       
   300 lemma one_step:
       
   301   assumes a: "a \<in> s" "j < n"
       
   302   assumes *: "\<And>a'. a' \<in> s \<Longrightarrow> a' \<noteq> a \<Longrightarrow> a' j = p'"
       
   303   shows "a j \<noteq> p'"
       
   304 proof
       
   305   assume "a j = p'"
       
   306   with * a have "\<And>a'. a' \<in> s \<Longrightarrow> a' j = p'"
       
   307     by auto
       
   308   then have "\<And>i. i \<le> n \<Longrightarrow> enum i j = p'"
       
   309     unfolding s_eq by auto
       
   310   from this[of 0] this[of n] have "j \<notin> upd ` {..< n}"
       
   311     by (auto simp: enum_def fun_eq_iff split: if_split_asm)
       
   312   with upd \<open>j < n\<close> show False
       
   313     by (auto simp: bij_betw_def)
       
   314 qed
       
   315 
       
   316 lemma upd_inj: "i < n \<Longrightarrow> j < n \<Longrightarrow> upd i = upd j \<longleftrightarrow> i = j"
       
   317   using upd by (auto simp: bij_betw_def inj_on_eq_iff)
       
   318 
       
   319 lemma upd_surj: "upd ` {..< n} = {..< n}"
       
   320   using upd by (auto simp: bij_betw_def)
       
   321 
       
   322 lemma in_upd_image: "A \<subseteq> {..< n} \<Longrightarrow> i < n \<Longrightarrow> upd i \<in> upd ` A \<longleftrightarrow> i \<in> A"
       
   323   using inj_on_image_mem_iff[of upd "{..< n}"] upd
       
   324   by (auto simp: bij_betw_def)
       
   325 
       
   326 lemma enum_inj: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i = enum j \<longleftrightarrow> i = j"
       
   327   using inj_enum by (auto simp: inj_on_eq_iff)
       
   328 
       
   329 lemma in_enum_image: "A \<subseteq> {.. n} \<Longrightarrow> i \<le> n \<Longrightarrow> enum i \<in> enum ` A \<longleftrightarrow> i \<in> A"
       
   330   using inj_on_image_mem_iff[OF inj_enum] by auto
       
   331 
       
   332 lemma enum_mono: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i \<le> enum j \<longleftrightarrow> i \<le> j"
       
   333   by (auto simp: enum_def le_fun_def in_upd_image Ball_def[symmetric])
       
   334 
       
   335 lemma enum_strict_mono: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i < enum j \<longleftrightarrow> i < j"
       
   336   using enum_mono[of i j] enum_inj[of i j] by (auto simp add: le_less)
       
   337 
       
   338 lemma chain: "a \<in> s \<Longrightarrow> b \<in> s \<Longrightarrow> a \<le> b \<or> b \<le> a"
       
   339   by (auto simp: s_eq enum_mono)
       
   340 
       
   341 lemma less: "a \<in> s \<Longrightarrow> b \<in> s \<Longrightarrow> a i < b i \<Longrightarrow> a < b"
       
   342   using chain[of a b] by (auto simp: less_fun_def le_fun_def not_le[symmetric])
       
   343 
       
   344 lemma enum_0_bot: "a \<in> s \<Longrightarrow> a = enum 0 \<longleftrightarrow> (\<forall>a'\<in>s. a \<le> a')"
       
   345   unfolding s_eq by (auto simp: enum_mono Ball_def)
       
   346 
       
   347 lemma enum_n_top: "a \<in> s \<Longrightarrow> a = enum n \<longleftrightarrow> (\<forall>a'\<in>s. a' \<le> a)"
       
   348   unfolding s_eq by (auto simp: enum_mono Ball_def)
       
   349 
       
   350 lemma enum_Suc: "i < n \<Longrightarrow> enum (Suc i) = (enum i)(upd i := Suc (enum i (upd i)))"
       
   351   by (auto simp: fun_eq_iff enum_def upd_inj)
       
   352 
       
   353 lemma enum_eq_p: "i \<le> n \<Longrightarrow> n \<le> j \<Longrightarrow> enum i j = p"
       
   354   by (induct i) (auto simp: enum_Suc enum_0 base_out upd_space not_less[symmetric])
       
   355 
       
   356 lemma out_eq_p: "a \<in> s \<Longrightarrow> n \<le> j \<Longrightarrow> a j = p"
       
   357   unfolding s_eq by (auto simp add: enum_eq_p)
       
   358 
       
   359 lemma s_le_p: "a \<in> s \<Longrightarrow> a j \<le> p"
       
   360   using out_eq_p[of a j] s_space by (cases "j < n") auto
       
   361 
       
   362 lemma le_Suc_base: "a \<in> s \<Longrightarrow> a j \<le> Suc (base j)"
       
   363   unfolding s_eq by (auto simp: enum_def)
       
   364 
       
   365 lemma base_le: "a \<in> s \<Longrightarrow> base j \<le> a j"
       
   366   unfolding s_eq by (auto simp: enum_def)
       
   367 
       
   368 lemma enum_le_p: "i \<le> n \<Longrightarrow> j < n \<Longrightarrow> enum i j \<le> p"
       
   369   using enum_in[of i] s_space by auto
       
   370 
       
   371 lemma enum_less: "a \<in> s \<Longrightarrow> i < n \<Longrightarrow> enum i < a \<longleftrightarrow> enum (Suc i) \<le> a"
       
   372   unfolding s_eq by (auto simp: enum_strict_mono enum_mono)
       
   373 
       
   374 lemma ksimplex_0:
       
   375   "n = 0 \<Longrightarrow> s = {(\<lambda>x. p)}"
       
   376   using s_eq enum_def base_out by auto
       
   377 
       
   378 lemma replace_0:
       
   379   assumes "j < n" "a \<in> s" and p: "\<forall>x\<in>s - {a}. x j = 0" and "x \<in> s"
       
   380   shows "x \<le> a"
       
   381 proof cases
       
   382   assume "x \<noteq> a"
       
   383   have "a j \<noteq> 0"
       
   384     using assms by (intro one_step[where a=a]) auto
       
   385   with less[OF \<open>x\<in>s\<close> \<open>a\<in>s\<close>, of j] p[rule_format, of x] \<open>x \<in> s\<close> \<open>x \<noteq> a\<close>
       
   386   show ?thesis
       
   387     by auto
       
   388 qed simp
       
   389 
       
   390 lemma replace_1:
       
   391   assumes "j < n" "a \<in> s" and p: "\<forall>x\<in>s - {a}. x j = p" and "x \<in> s"
       
   392   shows "a \<le> x"
       
   393 proof cases
       
   394   assume "x \<noteq> a"
       
   395   have "a j \<noteq> p"
       
   396     using assms by (intro one_step[where a=a]) auto
       
   397   with enum_le_p[of _ j] \<open>j < n\<close> \<open>a\<in>s\<close>
       
   398   have "a j < p"
       
   399     by (auto simp: less_le s_eq)
       
   400   with less[OF \<open>a\<in>s\<close> \<open>x\<in>s\<close>, of j] p[rule_format, of x] \<open>x \<in> s\<close> \<open>x \<noteq> a\<close>
       
   401   show ?thesis
       
   402     by auto
       
   403 qed simp
       
   404 
       
   405 end
       
   406 
       
   407 locale kuhn_simplex_pair = s: kuhn_simplex p n b_s u_s s + t: kuhn_simplex p n b_t u_t t
       
   408   for p n b_s u_s s b_t u_t t
       
   409 begin
       
   410 
       
   411 lemma enum_eq:
       
   412   assumes l: "i \<le> l" "l \<le> j" and "j + d \<le> n"
       
   413   assumes eq: "s.enum ` {i .. j} = t.enum ` {i + d .. j + d}"
       
   414   shows "s.enum l = t.enum (l + d)"
       
   415 using l proof (induct l rule: dec_induct)
       
   416   case base
       
   417   then have s: "s.enum i \<in> t.enum ` {i + d .. j + d}" and t: "t.enum (i + d) \<in> s.enum ` {i .. j}"
       
   418     using eq by auto
       
   419   from t \<open>i \<le> j\<close> \<open>j + d \<le> n\<close> have "s.enum i \<le> t.enum (i + d)"
       
   420     by (auto simp: s.enum_mono)
       
   421   moreover from s \<open>i \<le> j\<close> \<open>j + d \<le> n\<close> have "t.enum (i + d) \<le> s.enum i"
       
   422     by (auto simp: t.enum_mono)
       
   423   ultimately show ?case
       
   424     by auto
       
   425 next
       
   426   case (step l)
       
   427   moreover from step.prems \<open>j + d \<le> n\<close> have
       
   428       "s.enum l < s.enum (Suc l)"
       
   429       "t.enum (l + d) < t.enum (Suc l + d)"
       
   430     by (simp_all add: s.enum_strict_mono t.enum_strict_mono)
       
   431   moreover have
       
   432       "s.enum (Suc l) \<in> t.enum ` {i + d .. j + d}"
       
   433       "t.enum (Suc l + d) \<in> s.enum ` {i .. j}"
       
   434     using step \<open>j + d \<le> n\<close> eq by (auto simp: s.enum_inj t.enum_inj)
       
   435   ultimately have "s.enum (Suc l) = t.enum (Suc (l + d))"
       
   436     using \<open>j + d \<le> n\<close>
       
   437     by (intro antisym s.enum_less[THEN iffD1] t.enum_less[THEN iffD1])
       
   438        (auto intro!: s.enum_in t.enum_in)
       
   439   then show ?case by simp
       
   440 qed
       
   441 
       
   442 lemma ksimplex_eq_bot:
       
   443   assumes a: "a \<in> s" "\<And>a'. a' \<in> s \<Longrightarrow> a \<le> a'"
       
   444   assumes b: "b \<in> t" "\<And>b'. b' \<in> t \<Longrightarrow> b \<le> b'"
       
   445   assumes eq: "s - {a} = t - {b}"
       
   446   shows "s = t"
       
   447 proof cases
       
   448   assume "n = 0" with s.ksimplex_0 t.ksimplex_0 show ?thesis by simp
       
   449 next
       
   450   assume "n \<noteq> 0"
       
   451   have "s.enum 0 = (s.enum (Suc 0)) (u_s 0 := s.enum (Suc 0) (u_s 0) - 1)"
       
   452        "t.enum 0 = (t.enum (Suc 0)) (u_t 0 := t.enum (Suc 0) (u_t 0) - 1)"
       
   453     using \<open>n \<noteq> 0\<close> by (simp_all add: s.enum_Suc t.enum_Suc)
       
   454   moreover have e0: "a = s.enum 0" "b = t.enum 0"
       
   455     using a b by (simp_all add: s.enum_0_bot t.enum_0_bot)
       
   456   moreover
       
   457   { fix j assume "0 < j" "j \<le> n"
       
   458     moreover have "s - {a} = s.enum ` {Suc 0 .. n}" "t - {b} = t.enum ` {Suc 0 .. n}"
       
   459       unfolding s.s_eq t.s_eq e0 by (auto simp: s.enum_inj t.enum_inj)
       
   460     ultimately have "s.enum j = t.enum j"
       
   461       using enum_eq[of "1" j n 0] eq by auto }
       
   462   note enum_eq = this
       
   463   then have "s.enum (Suc 0) = t.enum (Suc 0)"
       
   464     using \<open>n \<noteq> 0\<close> by auto
       
   465   moreover
       
   466   { fix j assume "Suc j < n"
       
   467     with enum_eq[of "Suc j"] enum_eq[of "Suc (Suc j)"]
       
   468     have "u_s (Suc j) = u_t (Suc j)"
       
   469       using s.enum_Suc[of "Suc j"] t.enum_Suc[of "Suc j"]
       
   470       by (auto simp: fun_eq_iff split: if_split_asm) }
       
   471   then have "\<And>j. 0 < j \<Longrightarrow> j < n \<Longrightarrow> u_s j = u_t j"
       
   472     by (auto simp: gr0_conv_Suc)
       
   473   with \<open>n \<noteq> 0\<close> have "u_t 0 = u_s 0"
       
   474     by (intro bij_betw_singleton_eq[OF t.upd s.upd, of 0]) auto
       
   475   ultimately have "a = b"
       
   476     by simp
       
   477   with assms show "s = t"
       
   478     by auto
       
   479 qed
       
   480 
       
   481 lemma ksimplex_eq_top:
       
   482   assumes a: "a \<in> s" "\<And>a'. a' \<in> s \<Longrightarrow> a' \<le> a"
       
   483   assumes b: "b \<in> t" "\<And>b'. b' \<in> t \<Longrightarrow> b' \<le> b"
       
   484   assumes eq: "s - {a} = t - {b}"
       
   485   shows "s = t"
       
   486 proof (cases n)
       
   487   assume "n = 0" with s.ksimplex_0 t.ksimplex_0 show ?thesis by simp
       
   488 next
       
   489   case (Suc n')
       
   490   have "s.enum n = (s.enum n') (u_s n' := Suc (s.enum n' (u_s n')))"
       
   491        "t.enum n = (t.enum n') (u_t n' := Suc (t.enum n' (u_t n')))"
       
   492     using Suc by (simp_all add: s.enum_Suc t.enum_Suc)
       
   493   moreover have en: "a = s.enum n" "b = t.enum n"
       
   494     using a b by (simp_all add: s.enum_n_top t.enum_n_top)
       
   495   moreover
       
   496   { fix j assume "j < n"
       
   497     moreover have "s - {a} = s.enum ` {0 .. n'}" "t - {b} = t.enum ` {0 .. n'}"
       
   498       unfolding s.s_eq t.s_eq en by (auto simp: s.enum_inj t.enum_inj Suc)
       
   499     ultimately have "s.enum j = t.enum j"
       
   500       using enum_eq[of "0" j n' 0] eq Suc by auto }
       
   501   note enum_eq = this
       
   502   then have "s.enum n' = t.enum n'"
       
   503     using Suc by auto
       
   504   moreover
       
   505   { fix j assume "j < n'"
       
   506     with enum_eq[of j] enum_eq[of "Suc j"]
       
   507     have "u_s j = u_t j"
       
   508       using s.enum_Suc[of j] t.enum_Suc[of j]
       
   509       by (auto simp: Suc fun_eq_iff split: if_split_asm) }
       
   510   then have "\<And>j. j < n' \<Longrightarrow> u_s j = u_t j"
       
   511     by (auto simp: gr0_conv_Suc)
       
   512   then have "u_t n' = u_s n'"
       
   513     by (intro bij_betw_singleton_eq[OF t.upd s.upd, of n']) (auto simp: Suc)
       
   514   ultimately have "a = b"
       
   515     by simp
       
   516   with assms show "s = t"
       
   517     by auto
       
   518 qed
       
   519 
       
   520 end
       
   521 
       
   522 inductive ksimplex for p n :: nat where
       
   523   ksimplex: "kuhn_simplex p n base upd s \<Longrightarrow> ksimplex p n s"
       
   524 
       
   525 lemma finite_ksimplexes: "finite {s. ksimplex p n s}"
       
   526 proof (rule finite_subset)
       
   527   { fix a s assume "ksimplex p n s" "a \<in> s"
       
   528     then obtain b u where "kuhn_simplex p n b u s" by (auto elim: ksimplex.cases)
       
   529     then interpret kuhn_simplex p n b u s .
       
   530     from s_space \<open>a \<in> s\<close> out_eq_p[OF \<open>a \<in> s\<close>]
       
   531     have "a \<in> (\<lambda>f x. if n \<le> x then p else f x) ` ({..< n} \<rightarrow>\<^sub>E {.. p})"
       
   532       by (auto simp: image_iff subset_eq Pi_iff split: if_split_asm
       
   533                intro!: bexI[of _ "restrict a {..< n}"]) }
       
   534   then show "{s. ksimplex p n s} \<subseteq> Pow ((\<lambda>f x. if n \<le> x then p else f x) ` ({..< n} \<rightarrow>\<^sub>E {.. p}))"
       
   535     by auto
       
   536 qed (simp add: finite_PiE)
       
   537 
       
   538 lemma ksimplex_card:
       
   539   assumes "ksimplex p n s" shows "card s = Suc n"
       
   540 using assms proof cases
       
   541   case (ksimplex u b)
       
   542   then interpret kuhn_simplex p n u b s .
       
   543   show ?thesis
       
   544     by (simp add: card_image s_eq inj_enum)
       
   545 qed
       
   546 
       
   547 lemma simplex_top_face:
       
   548   assumes "0 < p" "\<forall>x\<in>s'. x n = p"
       
   549   shows "ksimplex p n s' \<longleftrightarrow> (\<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> s' = s - {a})"
       
   550   using assms
       
   551 proof safe
       
   552   fix s a assume "ksimplex p (Suc n) s" and a: "a \<in> s" and na: "\<forall>x\<in>s - {a}. x n = p"
       
   553   then show "ksimplex p n (s - {a})"
       
   554   proof cases
       
   555     case (ksimplex base upd)
       
   556     then interpret kuhn_simplex p "Suc n" base upd "s" .
       
   557 
       
   558     have "a n < p"
       
   559       using one_step[of a n p] na \<open>a\<in>s\<close> s_space by (auto simp: less_le)
       
   560     then have "a = enum 0"
       
   561       using \<open>a \<in> s\<close> na by (subst enum_0_bot) (auto simp: le_less intro!: less[of a _ n])
       
   562     then have s_eq: "s - {a} = enum ` Suc ` {.. n}"
       
   563       using s_eq by (simp add: atMost_Suc_eq_insert_0 insert_ident Zero_notin_Suc in_enum_image subset_eq)
       
   564     then have "enum 1 \<in> s - {a}"
       
   565       by auto
       
   566     then have "upd 0 = n"
       
   567       using \<open>a n < p\<close> \<open>a = enum 0\<close> na[rule_format, of "enum 1"]
       
   568       by (auto simp: fun_eq_iff enum_Suc split: if_split_asm)
       
   569     then have "bij_betw upd (Suc ` {..< n}) {..< n}"
       
   570       using upd
       
   571       by (subst notIn_Un_bij_betw3[where b=0])
       
   572          (auto simp: lessThan_Suc[symmetric] lessThan_Suc_eq_insert_0)
       
   573     then have "bij_betw (upd\<circ>Suc) {..<n} {..<n}"
       
   574       by (rule bij_betw_trans[rotated]) (auto simp: bij_betw_def)
       
   575 
       
   576     have "a n = p - 1"
       
   577       using enum_Suc[of 0] na[rule_format, OF \<open>enum 1 \<in> s - {a}\<close>] \<open>a = enum 0\<close> by (auto simp: \<open>upd 0 = n\<close>)
       
   578 
       
   579     show ?thesis
       
   580     proof (rule ksimplex.intros, standard)
       
   581       show "bij_betw (upd\<circ>Suc) {..< n} {..< n}" by fact
       
   582       show "base(n := p) \<in> {..<n} \<rightarrow> {..<p}" "\<And>i. n\<le>i \<Longrightarrow> (base(n := p)) i = p"
       
   583         using base base_out by (auto simp: Pi_iff)
       
   584 
       
   585       have "\<And>i. Suc ` {..< i} = {..< Suc i} - {0}"
       
   586         by (auto simp: image_iff Ball_def) arith
       
   587       then have upd_Suc: "\<And>i. i \<le> n \<Longrightarrow> (upd\<circ>Suc) ` {..< i} = upd ` {..< Suc i} - {n}"
       
   588         using \<open>upd 0 = n\<close> upd_inj
       
   589         by (auto simp add: image_comp[symmetric] inj_on_image_set_diff[OF inj_upd])
       
   590       have n_in_upd: "\<And>i. n \<in> upd ` {..< Suc i}"
       
   591         using \<open>upd 0 = n\<close> by auto
       
   592 
       
   593       define f' where "f' i j =
       
   594         (if j \<in> (upd\<circ>Suc)`{..< i} then Suc ((base(n := p)) j) else (base(n := p)) j)" for i j
       
   595       { fix x i assume i[arith]: "i \<le> n" then have "enum (Suc i) x = f' i x"
       
   596           unfolding f'_def enum_def using \<open>a n < p\<close> \<open>a = enum 0\<close> \<open>upd 0 = n\<close> \<open>a n = p - 1\<close>
       
   597           by (simp add: upd_Suc enum_0 n_in_upd) }
       
   598       then show "s - {a} = f' ` {.. n}"
       
   599         unfolding s_eq image_comp by (intro image_cong) auto
       
   600     qed
       
   601   qed
       
   602 next
       
   603   assume "ksimplex p n s'" and *: "\<forall>x\<in>s'. x n = p"
       
   604   then show "\<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> s' = s - {a}"
       
   605   proof cases
       
   606     case (ksimplex base upd)
       
   607     then interpret kuhn_simplex p n base upd s' .
       
   608     define b where "b = base (n := p - 1)"
       
   609     define u where "u i = (case i of 0 \<Rightarrow> n | Suc i \<Rightarrow> upd i)" for i
       
   610 
       
   611     have "ksimplex p (Suc n) (s' \<union> {b})"
       
   612     proof (rule ksimplex.intros, standard)
       
   613       show "b \<in> {..<Suc n} \<rightarrow> {..<p}"
       
   614         using base \<open>0 < p\<close> unfolding lessThan_Suc b_def by (auto simp: PiE_iff)
       
   615       show "\<And>i. Suc n \<le> i \<Longrightarrow> b i = p"
       
   616         using base_out by (auto simp: b_def)
       
   617 
       
   618       have "bij_betw u (Suc ` {..< n} \<union> {0}) ({..<n} \<union> {u 0})"
       
   619         using upd
       
   620         by (intro notIn_Un_bij_betw) (auto simp: u_def bij_betw_def image_comp comp_def inj_on_def)
       
   621       then show "bij_betw u {..<Suc n} {..<Suc n}"
       
   622         by (simp add: u_def lessThan_Suc[symmetric] lessThan_Suc_eq_insert_0)
       
   623 
       
   624       define f' where "f' i j = (if j \<in> u`{..< i} then Suc (b j) else b j)" for i j
       
   625 
       
   626       have u_eq: "\<And>i. i \<le> n \<Longrightarrow> u ` {..< Suc i} = upd ` {..< i} \<union> { n }"
       
   627         by (auto simp: u_def image_iff upd_inj Ball_def split: nat.split) arith
       
   628 
       
   629       { fix x have "x \<le> n \<Longrightarrow> n \<notin> upd ` {..<x}"
       
   630           using upd_space by (simp add: image_iff neq_iff) }
       
   631       note n_not_upd = this
       
   632 
       
   633       have *: "f' ` {.. Suc n} = f' ` (Suc ` {.. n} \<union> {0})"
       
   634         unfolding atMost_Suc_eq_insert_0 by simp
       
   635       also have "\<dots> = (f' \<circ> Suc) ` {.. n} \<union> {b}"
       
   636         by (auto simp: f'_def)
       
   637       also have "(f' \<circ> Suc) ` {.. n} = s'"
       
   638         using \<open>0 < p\<close> base_out[of n]
       
   639         unfolding s_eq enum_def[abs_def] f'_def[abs_def] upd_space
       
   640         by (intro image_cong) (simp_all add: u_eq b_def fun_eq_iff n_not_upd)
       
   641       finally show "s' \<union> {b} = f' ` {.. Suc n}" ..
       
   642     qed
       
   643     moreover have "b \<notin> s'"
       
   644       using * \<open>0 < p\<close> by (auto simp: b_def)
       
   645     ultimately show ?thesis by auto
       
   646   qed
       
   647 qed
       
   648 
       
   649 lemma ksimplex_replace_0:
       
   650   assumes s: "ksimplex p n s" and a: "a \<in> s"
       
   651   assumes j: "j < n" and p: "\<forall>x\<in>s - {a}. x j = 0"
       
   652   shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 1"
       
   653   using s
       
   654 proof cases
       
   655   case (ksimplex b_s u_s)
       
   656 
       
   657   { fix t b assume "ksimplex p n t"
       
   658     then obtain b_t u_t where "kuhn_simplex p n b_t u_t t"
       
   659       by (auto elim: ksimplex.cases)
       
   660     interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t
       
   661       by intro_locales fact+
       
   662 
       
   663     assume b: "b \<in> t" "t - {b} = s - {a}"
       
   664     with a j p s.replace_0[of _ a] t.replace_0[of _ b] have "s = t"
       
   665       by (intro ksimplex_eq_top[of a b]) auto }
       
   666   then have "{s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = {s}"
       
   667     using s \<open>a \<in> s\<close> by auto
       
   668   then show ?thesis
       
   669     by simp
       
   670 qed
       
   671 
       
   672 lemma ksimplex_replace_1:
       
   673   assumes s: "ksimplex p n s" and a: "a \<in> s"
       
   674   assumes j: "j < n" and p: "\<forall>x\<in>s - {a}. x j = p"
       
   675   shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 1"
       
   676   using s
       
   677 proof cases
       
   678   case (ksimplex b_s u_s)
       
   679 
       
   680   { fix t b assume "ksimplex p n t"
       
   681     then obtain b_t u_t where "kuhn_simplex p n b_t u_t t"
       
   682       by (auto elim: ksimplex.cases)
       
   683     interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t
       
   684       by intro_locales fact+
       
   685 
       
   686     assume b: "b \<in> t" "t - {b} = s - {a}"
       
   687     with a j p s.replace_1[of _ a] t.replace_1[of _ b] have "s = t"
       
   688       by (intro ksimplex_eq_bot[of a b]) auto }
       
   689   then have "{s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = {s}"
       
   690     using s \<open>a \<in> s\<close> by auto
       
   691   then show ?thesis
       
   692     by simp
       
   693 qed
       
   694 
       
   695 lemma card_2_exists: "card s = 2 \<longleftrightarrow> (\<exists>x\<in>s. \<exists>y\<in>s. x \<noteq> y \<and> (\<forall>z\<in>s. z = x \<or> z = y))"
       
   696   by (auto simp add: card_Suc_eq eval_nat_numeral)
       
   697 
       
   698 lemma ksimplex_replace_2:
       
   699   assumes s: "ksimplex p n s" and "a \<in> s" and "n \<noteq> 0"
       
   700     and lb: "\<forall>j<n. \<exists>x\<in>s - {a}. x j \<noteq> 0"
       
   701     and ub: "\<forall>j<n. \<exists>x\<in>s - {a}. x j \<noteq> p"
       
   702   shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 2"
       
   703   using s
       
   704 proof cases
       
   705   case (ksimplex base upd)
       
   706   then interpret kuhn_simplex p n base upd s .
       
   707 
       
   708   from \<open>a \<in> s\<close> obtain i where "i \<le> n" "a = enum i"
       
   709     unfolding s_eq by auto
       
   710 
       
   711   from \<open>i \<le> n\<close> have "i = 0 \<or> i = n \<or> (0 < i \<and> i < n)"
       
   712     by linarith
       
   713   then have "\<exists>!s'. s' \<noteq> s \<and> ksimplex p n s' \<and> (\<exists>b\<in>s'. s - {a} = s'- {b})"
       
   714   proof (elim disjE conjE)
       
   715     assume "i = 0"
       
   716     define rot where [abs_def]: "rot i = (if i + 1 = n then 0 else i + 1)" for i
       
   717     let ?upd = "upd \<circ> rot"
       
   718 
       
   719     have rot: "bij_betw rot {..< n} {..< n}"
       
   720       by (auto simp: bij_betw_def inj_on_def image_iff Ball_def rot_def)
       
   721          arith+
       
   722     from rot upd have "bij_betw ?upd {..<n} {..<n}"
       
   723       by (rule bij_betw_trans)
       
   724 
       
   725     define f' where [abs_def]: "f' i j =
       
   726       (if j \<in> ?upd`{..< i} then Suc (enum (Suc 0) j) else enum (Suc 0) j)" for i j
       
   727 
       
   728     interpret b: kuhn_simplex p n "enum (Suc 0)" "upd \<circ> rot" "f' ` {.. n}"
       
   729     proof
       
   730       from \<open>a = enum i\<close> ub \<open>n \<noteq> 0\<close> \<open>i = 0\<close>
       
   731       obtain i' where "i' \<le> n" "enum i' \<noteq> enum 0" "enum i' (upd 0) \<noteq> p"
       
   732         unfolding s_eq by (auto intro: upd_space simp: enum_inj)
       
   733       then have "enum 1 \<le> enum i'" "enum i' (upd 0) < p"
       
   734         using enum_le_p[of i' "upd 0"] by (auto simp add: enum_inj enum_mono upd_space)
       
   735       then have "enum 1 (upd 0) < p"
       
   736         by (auto simp add: le_fun_def intro: le_less_trans)
       
   737       then show "enum (Suc 0) \<in> {..<n} \<rightarrow> {..<p}"
       
   738         using base \<open>n \<noteq> 0\<close> by (auto simp add: enum_0 enum_Suc PiE_iff extensional_def upd_space)
       
   739 
       
   740       { fix i assume "n \<le> i" then show "enum (Suc 0) i = p"
       
   741         using \<open>n \<noteq> 0\<close> by (auto simp: enum_eq_p) }
       
   742       show "bij_betw ?upd {..<n} {..<n}" by fact
       
   743     qed (simp add: f'_def)
       
   744     have ks_f': "ksimplex p n (f' ` {.. n})"
       
   745       by rule unfold_locales
       
   746 
       
   747     have b_enum: "b.enum = f'" unfolding f'_def b.enum_def[abs_def] ..
       
   748     with b.inj_enum have inj_f': "inj_on f' {.. n}" by simp
       
   749 
       
   750     have [simp]: "\<And>j. j < n \<Longrightarrow> rot ` {..< j} = {0 <..< Suc j}"
       
   751       by (auto simp: rot_def image_iff Ball_def)
       
   752          arith
       
   753 
       
   754     { fix j assume j: "j < n"
       
   755       from j \<open>n \<noteq> 0\<close> have "f' j = enum (Suc j)"
       
   756         by (auto simp add: f'_def enum_def upd_inj in_upd_image image_comp[symmetric] fun_eq_iff) }
       
   757     note f'_eq_enum = this
       
   758     then have "enum ` Suc ` {..< n} = f' ` {..< n}"
       
   759       by (force simp: enum_inj)
       
   760     also have "Suc ` {..< n} = {.. n} - {0}"
       
   761       by (auto simp: image_iff Ball_def) arith
       
   762     also have "{..< n} = {.. n} - {n}"
       
   763       by auto
       
   764     finally have eq: "s - {a} = f' ` {.. n} - {f' n}"
       
   765       unfolding s_eq \<open>a = enum i\<close> \<open>i = 0\<close>
       
   766       by (simp add: Diff_subset inj_on_image_set_diff[OF inj_enum] inj_on_image_set_diff[OF inj_f'])
       
   767 
       
   768     have "enum 0 < f' 0"
       
   769       using \<open>n \<noteq> 0\<close> by (simp add: enum_strict_mono f'_eq_enum)
       
   770     also have "\<dots> < f' n"
       
   771       using \<open>n \<noteq> 0\<close> b.enum_strict_mono[of 0 n] unfolding b_enum by simp
       
   772     finally have "a \<noteq> f' n"
       
   773       using \<open>a = enum i\<close> \<open>i = 0\<close> by auto
       
   774 
       
   775     { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
       
   776       obtain b u where "kuhn_simplex p n b u t"
       
   777         using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases)
       
   778       then interpret t: kuhn_simplex p n b u t .
       
   779 
       
   780       { fix x assume "x \<in> s" "x \<noteq> a"
       
   781          then have "x (upd 0) = enum (Suc 0) (upd 0)"
       
   782            by (auto simp: \<open>a = enum i\<close> \<open>i = 0\<close> s_eq enum_def enum_inj) }
       
   783       then have eq_upd0: "\<forall>x\<in>t-{c}. x (upd 0) = enum (Suc 0) (upd 0)"
       
   784         unfolding eq_sma[symmetric] by auto
       
   785       then have "c (upd 0) \<noteq> enum (Suc 0) (upd 0)"
       
   786         using \<open>n \<noteq> 0\<close> by (intro t.one_step[OF \<open>c\<in>t\<close> ]) (auto simp: upd_space)
       
   787       then have "c (upd 0) < enum (Suc 0) (upd 0) \<or> c (upd 0) > enum (Suc 0) (upd 0)"
       
   788         by auto
       
   789       then have "t = s \<or> t = f' ` {..n}"
       
   790       proof (elim disjE conjE)
       
   791         assume *: "c (upd 0) < enum (Suc 0) (upd 0)"
       
   792         interpret st: kuhn_simplex_pair p n base upd s b u t ..
       
   793         { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "c \<le> x"
       
   794             by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) }
       
   795         note top = this
       
   796         have "s = t"
       
   797           using \<open>a = enum i\<close> \<open>i = 0\<close> \<open>c \<in> t\<close>
       
   798           by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq_sma])
       
   799              (auto simp: s_eq enum_mono t.s_eq t.enum_mono top)
       
   800         then show ?thesis by simp
       
   801       next
       
   802         assume *: "c (upd 0) > enum (Suc 0) (upd 0)"
       
   803         interpret st: kuhn_simplex_pair p n "enum (Suc 0)" "upd \<circ> rot" "f' ` {.. n}" b u t ..
       
   804         have eq: "f' ` {..n} - {f' n} = t - {c}"
       
   805           using eq_sma eq by simp
       
   806         { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "x \<le> c"
       
   807             by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) }
       
   808         note top = this
       
   809         have "f' ` {..n} = t"
       
   810           using \<open>a = enum i\<close> \<open>i = 0\<close> \<open>c \<in> t\<close>
       
   811           by (intro st.ksimplex_eq_top[OF _ _ _ _ eq])
       
   812              (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono b_enum[symmetric] top)
       
   813         then show ?thesis by simp
       
   814       qed }
       
   815     with ks_f' eq \<open>a \<noteq> f' n\<close> \<open>n \<noteq> 0\<close> show ?thesis
       
   816       apply (intro ex1I[of _ "f' ` {.. n}"])
       
   817       apply auto []
       
   818       apply metis
       
   819       done
       
   820   next
       
   821     assume "i = n"
       
   822     from \<open>n \<noteq> 0\<close> obtain n' where n': "n = Suc n'"
       
   823       by (cases n) auto
       
   824 
       
   825     define rot where "rot i = (case i of 0 \<Rightarrow> n' | Suc i \<Rightarrow> i)" for i
       
   826     let ?upd = "upd \<circ> rot"
       
   827 
       
   828     have rot: "bij_betw rot {..< n} {..< n}"
       
   829       by (auto simp: bij_betw_def inj_on_def image_iff Bex_def rot_def n' split: nat.splits)
       
   830          arith
       
   831     from rot upd have "bij_betw ?upd {..<n} {..<n}"
       
   832       by (rule bij_betw_trans)
       
   833 
       
   834     define b where "b = base (upd n' := base (upd n') - 1)"
       
   835     define f' where [abs_def]: "f' i j = (if j \<in> ?upd`{..< i} then Suc (b j) else b j)" for i j
       
   836 
       
   837     interpret b: kuhn_simplex p n b "upd \<circ> rot" "f' ` {.. n}"
       
   838     proof
       
   839       { fix i assume "n \<le> i" then show "b i = p"
       
   840           using base_out[of i] upd_space[of n'] by (auto simp: b_def n') }
       
   841       show "b \<in> {..<n} \<rightarrow> {..<p}"
       
   842         using base \<open>n \<noteq> 0\<close> upd_space[of n']
       
   843         by (auto simp: b_def PiE_def Pi_iff Ball_def upd_space extensional_def n')
       
   844 
       
   845       show "bij_betw ?upd {..<n} {..<n}" by fact
       
   846     qed (simp add: f'_def)
       
   847     have f': "b.enum = f'" unfolding f'_def b.enum_def[abs_def] ..
       
   848     have ks_f': "ksimplex p n (b.enum ` {.. n})"
       
   849       unfolding f' by rule unfold_locales
       
   850 
       
   851     have "0 < n"
       
   852       using \<open>n \<noteq> 0\<close> by auto
       
   853 
       
   854     { from \<open>a = enum i\<close> \<open>n \<noteq> 0\<close> \<open>i = n\<close> lb upd_space[of n']
       
   855       obtain i' where "i' \<le> n" "enum i' \<noteq> enum n" "0 < enum i' (upd n')"
       
   856         unfolding s_eq by (auto simp: enum_inj n')
       
   857       moreover have "enum i' (upd n') = base (upd n')"
       
   858         unfolding enum_def using \<open>i' \<le> n\<close> \<open>enum i' \<noteq> enum n\<close> by (auto simp: n' upd_inj enum_inj)
       
   859       ultimately have "0 < base (upd n')"
       
   860         by auto }
       
   861     then have benum1: "b.enum (Suc 0) = base"
       
   862       unfolding b.enum_Suc[OF \<open>0<n\<close>] b.enum_0 by (auto simp: b_def rot_def)
       
   863 
       
   864     have [simp]: "\<And>j. Suc j < n \<Longrightarrow> rot ` {..< Suc j} = {n'} \<union> {..< j}"
       
   865       by (auto simp: rot_def image_iff Ball_def split: nat.splits)
       
   866     have rot_simps: "\<And>j. rot (Suc j) = j" "rot 0 = n'"
       
   867       by (simp_all add: rot_def)
       
   868 
       
   869     { fix j assume j: "Suc j \<le> n" then have "b.enum (Suc j) = enum j"
       
   870         by (induct j) (auto simp add: benum1 enum_0 b.enum_Suc enum_Suc rot_simps) }
       
   871     note b_enum_eq_enum = this
       
   872     then have "enum ` {..< n} = b.enum ` Suc ` {..< n}"
       
   873       by (auto simp add: image_comp intro!: image_cong)
       
   874     also have "Suc ` {..< n} = {.. n} - {0}"
       
   875       by (auto simp: image_iff Ball_def) arith
       
   876     also have "{..< n} = {.. n} - {n}"
       
   877       by auto
       
   878     finally have eq: "s - {a} = b.enum ` {.. n} - {b.enum 0}"
       
   879       unfolding s_eq \<open>a = enum i\<close> \<open>i = n\<close>
       
   880       using inj_on_image_set_diff[OF inj_enum Diff_subset, of "{n}"]
       
   881             inj_on_image_set_diff[OF b.inj_enum Diff_subset, of "{0}"]
       
   882       by (simp add: comp_def )
       
   883 
       
   884     have "b.enum 0 \<le> b.enum n"
       
   885       by (simp add: b.enum_mono)
       
   886     also have "b.enum n < enum n"
       
   887       using \<open>n \<noteq> 0\<close> by (simp add: enum_strict_mono b_enum_eq_enum n')
       
   888     finally have "a \<noteq> b.enum 0"
       
   889       using \<open>a = enum i\<close> \<open>i = n\<close> by auto
       
   890 
       
   891     { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
       
   892       obtain b' u where "kuhn_simplex p n b' u t"
       
   893         using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases)
       
   894       then interpret t: kuhn_simplex p n b' u t .
       
   895 
       
   896       { fix x assume "x \<in> s" "x \<noteq> a"
       
   897          then have "x (upd n') = enum n' (upd n')"
       
   898            by (auto simp: \<open>a = enum i\<close> n' \<open>i = n\<close> s_eq enum_def enum_inj in_upd_image) }
       
   899       then have eq_upd0: "\<forall>x\<in>t-{c}. x (upd n') = enum n' (upd n')"
       
   900         unfolding eq_sma[symmetric] by auto
       
   901       then have "c (upd n') \<noteq> enum n' (upd n')"
       
   902         using \<open>n \<noteq> 0\<close> by (intro t.one_step[OF \<open>c\<in>t\<close> ]) (auto simp: n' upd_space[unfolded n'])
       
   903       then have "c (upd n') < enum n' (upd n') \<or> c (upd n') > enum n' (upd n')"
       
   904         by auto
       
   905       then have "t = s \<or> t = b.enum ` {..n}"
       
   906       proof (elim disjE conjE)
       
   907         assume *: "c (upd n') > enum n' (upd n')"
       
   908         interpret st: kuhn_simplex_pair p n base upd s b' u t ..
       
   909         { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "x \<le> c"
       
   910             by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) }
       
   911         note top = this
       
   912         have "s = t"
       
   913           using \<open>a = enum i\<close> \<open>i = n\<close> \<open>c \<in> t\<close>
       
   914           by (intro st.ksimplex_eq_top[OF _ _ _ _ eq_sma])
       
   915              (auto simp: s_eq enum_mono t.s_eq t.enum_mono top)
       
   916         then show ?thesis by simp
       
   917       next
       
   918         assume *: "c (upd n') < enum n' (upd n')"
       
   919         interpret st: kuhn_simplex_pair p n b "upd \<circ> rot" "f' ` {.. n}" b' u t ..
       
   920         have eq: "f' ` {..n} - {b.enum 0} = t - {c}"
       
   921           using eq_sma eq f' by simp
       
   922         { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "c \<le> x"
       
   923             by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) }
       
   924         note bot = this
       
   925         have "f' ` {..n} = t"
       
   926           using \<open>a = enum i\<close> \<open>i = n\<close> \<open>c \<in> t\<close>
       
   927           by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq])
       
   928              (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono bot)
       
   929         with f' show ?thesis by simp
       
   930       qed }
       
   931     with ks_f' eq \<open>a \<noteq> b.enum 0\<close> \<open>n \<noteq> 0\<close> show ?thesis
       
   932       apply (intro ex1I[of _ "b.enum ` {.. n}"])
       
   933       apply auto []
       
   934       apply metis
       
   935       done
       
   936   next
       
   937     assume i: "0 < i" "i < n"
       
   938     define i' where "i' = i - 1"
       
   939     with i have "Suc i' < n"
       
   940       by simp
       
   941     with i have Suc_i': "Suc i' = i"
       
   942       by (simp add: i'_def)
       
   943 
       
   944     let ?upd = "Fun.swap i' i upd"
       
   945     from i upd have "bij_betw ?upd {..< n} {..< n}"
       
   946       by (subst bij_betw_swap_iff) (auto simp: i'_def)
       
   947 
       
   948     define f' where [abs_def]: "f' i j = (if j \<in> ?upd`{..< i} then Suc (base j) else base j)"
       
   949       for i j
       
   950     interpret b: kuhn_simplex p n base ?upd "f' ` {.. n}"
       
   951     proof
       
   952       show "base \<in> {..<n} \<rightarrow> {..<p}" by fact
       
   953       { fix i assume "n \<le> i" then show "base i = p" by fact }
       
   954       show "bij_betw ?upd {..<n} {..<n}" by fact
       
   955     qed (simp add: f'_def)
       
   956     have f': "b.enum = f'" unfolding f'_def b.enum_def[abs_def] ..
       
   957     have ks_f': "ksimplex p n (b.enum ` {.. n})"
       
   958       unfolding f' by rule unfold_locales
       
   959 
       
   960     have "{i} \<subseteq> {..n}"
       
   961       using i by auto
       
   962     { fix j assume "j \<le> n"
       
   963       moreover have "j < i \<or> i = j \<or> i < j" by arith
       
   964       moreover note i
       
   965       ultimately have "enum j = b.enum j \<longleftrightarrow> j \<noteq> i"
       
   966         unfolding enum_def[abs_def] b.enum_def[abs_def]
       
   967         by (auto simp add: fun_eq_iff swap_image i'_def
       
   968                            in_upd_image inj_on_image_set_diff[OF inj_upd]) }
       
   969     note enum_eq_benum = this
       
   970     then have "enum ` ({.. n} - {i}) = b.enum ` ({.. n} - {i})"
       
   971       by (intro image_cong) auto
       
   972     then have eq: "s - {a} = b.enum ` {.. n} - {b.enum i}"
       
   973       unfolding s_eq \<open>a = enum i\<close>
       
   974       using inj_on_image_set_diff[OF inj_enum Diff_subset \<open>{i} \<subseteq> {..n}\<close>]
       
   975             inj_on_image_set_diff[OF b.inj_enum Diff_subset \<open>{i} \<subseteq> {..n}\<close>]
       
   976       by (simp add: comp_def)
       
   977 
       
   978     have "a \<noteq> b.enum i"
       
   979       using \<open>a = enum i\<close> enum_eq_benum i by auto
       
   980 
       
   981     { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
       
   982       obtain b' u where "kuhn_simplex p n b' u t"
       
   983         using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases)
       
   984       then interpret t: kuhn_simplex p n b' u t .
       
   985       have "enum i' \<in> s - {a}" "enum (i + 1) \<in> s - {a}"
       
   986         using \<open>a = enum i\<close> i enum_in by (auto simp: enum_inj i'_def)
       
   987       then obtain l k where
       
   988         l: "t.enum l = enum i'" "l \<le> n" "t.enum l \<noteq> c" and
       
   989         k: "t.enum k = enum (i + 1)" "k \<le> n" "t.enum k \<noteq> c"
       
   990         unfolding eq_sma by (auto simp: t.s_eq)
       
   991       with i have "t.enum l < t.enum k"
       
   992         by (simp add: enum_strict_mono i'_def)
       
   993       with \<open>l \<le> n\<close> \<open>k \<le> n\<close> have "l < k"
       
   994         by (simp add: t.enum_strict_mono)
       
   995       { assume "Suc l = k"
       
   996         have "enum (Suc (Suc i')) = t.enum (Suc l)"
       
   997           using i by (simp add: k \<open>Suc l = k\<close> i'_def)
       
   998         then have False
       
   999           using \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close>
       
  1000           by (auto simp: t.enum_Suc enum_Suc l upd_inj fun_eq_iff split: if_split_asm)
       
  1001              (metis Suc_lessD n_not_Suc_n upd_inj) }
       
  1002       with \<open>l < k\<close> have "Suc l < k"
       
  1003         by arith
       
  1004       have c_eq: "c = t.enum (Suc l)"
       
  1005       proof (rule ccontr)
       
  1006         assume "c \<noteq> t.enum (Suc l)"
       
  1007         then have "t.enum (Suc l) \<in> s - {a}"
       
  1008           using \<open>l < k\<close> \<open>k \<le> n\<close> by (simp add: t.s_eq eq_sma)
       
  1009         then obtain j where "t.enum (Suc l) = enum j" "j \<le> n" "enum j \<noteq> enum i"
       
  1010           unfolding s_eq \<open>a = enum i\<close> by auto
       
  1011         with i have "t.enum (Suc l) \<le> t.enum l \<or> t.enum k \<le> t.enum (Suc l)"
       
  1012           by (auto simp add: i'_def enum_mono enum_inj l k)
       
  1013         with \<open>Suc l < k\<close> \<open>k \<le> n\<close> show False
       
  1014           by (simp add: t.enum_mono)
       
  1015       qed
       
  1016 
       
  1017       { have "t.enum (Suc (Suc l)) \<in> s - {a}"
       
  1018           unfolding eq_sma c_eq t.s_eq using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (auto simp: t.enum_inj)
       
  1019         then obtain j where eq: "t.enum (Suc (Suc l)) = enum j" and "j \<le> n" "j \<noteq> i"
       
  1020           by (auto simp: s_eq \<open>a = enum i\<close>)
       
  1021         moreover have "enum i' < t.enum (Suc (Suc l))"
       
  1022           unfolding l(1)[symmetric] using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (auto simp: t.enum_strict_mono)
       
  1023         ultimately have "i' < j"
       
  1024           using i by (simp add: enum_strict_mono i'_def)
       
  1025         with \<open>j \<noteq> i\<close> \<open>j \<le> n\<close> have "t.enum k \<le> t.enum (Suc (Suc l))"
       
  1026           unfolding i'_def by (simp add: enum_mono k eq)
       
  1027         then have "k \<le> Suc (Suc l)"
       
  1028           using \<open>k \<le> n\<close> \<open>Suc l < k\<close> by (simp add: t.enum_mono) }
       
  1029       with \<open>Suc l < k\<close> have "Suc (Suc l) = k" by simp
       
  1030       then have "enum (Suc (Suc i')) = t.enum (Suc (Suc l))"
       
  1031         using i by (simp add: k i'_def)
       
  1032       also have "\<dots> = (enum i') (u l := Suc (enum i' (u l)), u (Suc l) := Suc (enum i' (u (Suc l))))"
       
  1033         using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (simp add: t.enum_Suc l t.upd_inj)
       
  1034       finally have "(u l = upd i' \<and> u (Suc l) = upd (Suc i')) \<or>
       
  1035         (u l = upd (Suc i') \<and> u (Suc l) = upd i')"
       
  1036         using \<open>Suc i' < n\<close> by (auto simp: enum_Suc fun_eq_iff split: if_split_asm)
       
  1037 
       
  1038       then have "t = s \<or> t = b.enum ` {..n}"
       
  1039       proof (elim disjE conjE)
       
  1040         assume u: "u l = upd i'"
       
  1041         have "c = t.enum (Suc l)" unfolding c_eq ..
       
  1042         also have "t.enum (Suc l) = enum (Suc i')"
       
  1043           using u \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close> by (simp add: enum_Suc t.enum_Suc l)
       
  1044         also have "\<dots> = a"
       
  1045           using \<open>a = enum i\<close> i by (simp add: i'_def)
       
  1046         finally show ?thesis
       
  1047           using eq_sma \<open>a \<in> s\<close> \<open>c \<in> t\<close> by auto
       
  1048       next
       
  1049         assume u: "u l = upd (Suc i')"
       
  1050         define B where "B = b.enum ` {..n}"
       
  1051         have "b.enum i' = enum i'"
       
  1052           using enum_eq_benum[of i'] i by (auto simp add: i'_def gr0_conv_Suc)
       
  1053         have "c = t.enum (Suc l)" unfolding c_eq ..
       
  1054         also have "t.enum (Suc l) = b.enum (Suc i')"
       
  1055           using u \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close>
       
  1056           by (simp_all add: enum_Suc t.enum_Suc l b.enum_Suc \<open>b.enum i' = enum i'\<close> swap_apply1)
       
  1057              (simp add: Suc_i')
       
  1058         also have "\<dots> = b.enum i"
       
  1059           using i by (simp add: i'_def)
       
  1060         finally have "c = b.enum i" .
       
  1061         then have "t - {c} = B - {c}" "c \<in> B"
       
  1062           unfolding eq_sma[symmetric] eq B_def using i by auto
       
  1063         with \<open>c \<in> t\<close> have "t = B"
       
  1064           by auto
       
  1065         then show ?thesis
       
  1066           by (simp add: B_def)
       
  1067       qed }
       
  1068     with ks_f' eq \<open>a \<noteq> b.enum i\<close> \<open>n \<noteq> 0\<close> \<open>i \<le> n\<close> show ?thesis
       
  1069       apply (intro ex1I[of _ "b.enum ` {.. n}"])
       
  1070       apply auto []
       
  1071       apply metis
       
  1072       done
       
  1073   qed
       
  1074   then show ?thesis
       
  1075     using s \<open>a \<in> s\<close> by (simp add: card_2_exists Ex1_def) metis
       
  1076 qed
       
  1077 
       
  1078 text \<open>Hence another step towards concreteness.\<close>
       
  1079 
       
  1080 lemma kuhn_simplex_lemma:
       
  1081   assumes "\<forall>s. ksimplex p (Suc n) s \<longrightarrow> rl ` s \<subseteq> {.. Suc n}"
       
  1082     and "odd (card {f. \<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> (f = s - {a}) \<and>
       
  1083       rl ` f = {..n} \<and> ((\<exists>j\<le>n. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>f. x j = p))})"
       
  1084   shows "odd (card {s. ksimplex p (Suc n) s \<and> rl ` s = {..Suc n}})"
       
  1085 proof (rule kuhn_complete_lemma[OF finite_ksimplexes refl, unfolded mem_Collect_eq,
       
  1086     where bnd="\<lambda>f. (\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = p)"],
       
  1087     safe del: notI)
       
  1088 
       
  1089   have *: "\<And>x y. x = y \<Longrightarrow> odd (card x) \<Longrightarrow> odd (card y)"
       
  1090     by auto
       
  1091   show "odd (card {f. (\<exists>s\<in>{s. ksimplex p (Suc n) s}. \<exists>a\<in>s. f = s - {a}) \<and>
       
  1092     rl ` f = {..n} \<and> ((\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = p))})"
       
  1093     apply (rule *[OF _ assms(2)])
       
  1094     apply (auto simp: atLeast0AtMost)
       
  1095     done
       
  1096 
       
  1097 next
       
  1098 
       
  1099   fix s assume s: "ksimplex p (Suc n) s"
       
  1100   then show "card s = n + 2"
       
  1101     by (simp add: ksimplex_card)
       
  1102 
       
  1103   fix a assume a: "a \<in> s" then show "rl a \<le> Suc n"
       
  1104     using assms(1) s by (auto simp: subset_eq)
       
  1105 
       
  1106   let ?S = "{t. ksimplex p (Suc n) t \<and> (\<exists>b\<in>t. s - {a} = t - {b})}"
       
  1107   { fix j assume j: "j \<le> n" "\<forall>x\<in>s - {a}. x j = 0"
       
  1108     with s a show "card ?S = 1"
       
  1109       using ksimplex_replace_0[of p "n + 1" s a j]
       
  1110       by (subst eq_commute) simp }
       
  1111 
       
  1112   { fix j assume j: "j \<le> n" "\<forall>x\<in>s - {a}. x j = p"
       
  1113     with s a show "card ?S = 1"
       
  1114       using ksimplex_replace_1[of p "n + 1" s a j]
       
  1115       by (subst eq_commute) simp }
       
  1116 
       
  1117   { assume "card ?S \<noteq> 2" "\<not> (\<exists>j\<in>{..n}. \<forall>x\<in>s - {a}. x j = p)"
       
  1118     with s a show "\<exists>j\<in>{..n}. \<forall>x\<in>s - {a}. x j = 0"
       
  1119       using ksimplex_replace_2[of p "n + 1" s a]
       
  1120       by (subst (asm) eq_commute) auto }
       
  1121 qed
       
  1122 
       
  1123 subsection \<open>Reduced labelling\<close>
       
  1124 
       
  1125 definition reduced :: "nat \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> nat" where "reduced n x = (LEAST k. k = n \<or> x k \<noteq> 0)"
       
  1126 
       
  1127 lemma reduced_labelling:
       
  1128   shows "reduced n x \<le> n"
       
  1129     and "\<forall>i<reduced n x. x i = 0"
       
  1130     and "reduced n x = n \<or> x (reduced n x) \<noteq> 0"
       
  1131 proof -
       
  1132   show "reduced n x \<le> n"
       
  1133     unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) auto
       
  1134   show "\<forall>i<reduced n x. x i = 0"
       
  1135     unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) fastforce+
       
  1136   show "reduced n x = n \<or> x (reduced n x) \<noteq> 0"
       
  1137     unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) fastforce+
       
  1138 qed
       
  1139 
       
  1140 lemma reduced_labelling_unique:
       
  1141   "r \<le> n \<Longrightarrow> \<forall>i<r. x i = 0 \<Longrightarrow> r = n \<or> x r \<noteq> 0 \<Longrightarrow> reduced n x = r"
       
  1142  unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) (metis le_less not_le)+
       
  1143 
       
  1144 lemma reduced_labelling_zero: "j < n \<Longrightarrow> x j = 0 \<Longrightarrow> reduced n x \<noteq> j"
       
  1145   using reduced_labelling[of n x] by auto
       
  1146 
       
  1147 lemma reduce_labelling_zero[simp]: "reduced 0 x = 0"
       
  1148   by (rule reduced_labelling_unique) auto
       
  1149 
       
  1150 lemma reduced_labelling_nonzero: "j < n \<Longrightarrow> x j \<noteq> 0 \<Longrightarrow> reduced n x \<le> j"
       
  1151   using reduced_labelling[of n x] by (elim allE[where x=j]) auto
       
  1152 
       
  1153 lemma reduced_labelling_Suc: "reduced (Suc n) x \<noteq> Suc n \<Longrightarrow> reduced (Suc n) x = reduced n x"
       
  1154   using reduced_labelling[of "Suc n" x]
       
  1155   by (intro reduced_labelling_unique[symmetric]) auto
       
  1156 
       
  1157 lemma complete_face_top:
       
  1158   assumes "\<forall>x\<in>f. \<forall>j\<le>n. x j = 0 \<longrightarrow> lab x j = 0"
       
  1159     and "\<forall>x\<in>f. \<forall>j\<le>n. x j = p \<longrightarrow> lab x j = 1"
       
  1160     and eq: "(reduced (Suc n) \<circ> lab) ` f = {..n}"
       
  1161   shows "((\<exists>j\<le>n. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>f. x j = p)) \<longleftrightarrow> (\<forall>x\<in>f. x n = p)"
       
  1162 proof (safe del: disjCI)
       
  1163   fix x j assume j: "j \<le> n" "\<forall>x\<in>f. x j = 0"
       
  1164   { fix x assume "x \<in> f" with assms j have "reduced (Suc n) (lab x) \<noteq> j"
       
  1165       by (intro reduced_labelling_zero) auto }
       
  1166   moreover have "j \<in> (reduced (Suc n) \<circ> lab) ` f"
       
  1167     using j eq by auto
       
  1168   ultimately show "x n = p"
       
  1169     by force
       
  1170 next
       
  1171   fix x j assume j: "j \<le> n" "\<forall>x\<in>f. x j = p" and x: "x \<in> f"
       
  1172   have "j = n"
       
  1173   proof (rule ccontr)
       
  1174     assume "\<not> ?thesis"
       
  1175     { fix x assume "x \<in> f"
       
  1176       with assms j have "reduced (Suc n) (lab x) \<le> j"
       
  1177         by (intro reduced_labelling_nonzero) auto
       
  1178       then have "reduced (Suc n) (lab x) \<noteq> n"
       
  1179         using \<open>j \<noteq> n\<close> \<open>j \<le> n\<close> by simp }
       
  1180     moreover
       
  1181     have "n \<in> (reduced (Suc n) \<circ> lab) ` f"
       
  1182       using eq by auto
       
  1183     ultimately show False
       
  1184       by force
       
  1185   qed
       
  1186   moreover have "j \<in> (reduced (Suc n) \<circ> lab) ` f"
       
  1187     using j eq by auto
       
  1188   ultimately show "x n = p"
       
  1189     using j x by auto
       
  1190 qed auto
       
  1191 
       
  1192 text \<open>Hence we get just about the nice induction.\<close>
       
  1193 
       
  1194 lemma kuhn_induction:
       
  1195   assumes "0 < p"
       
  1196     and lab_0: "\<forall>x. \<forall>j\<le>n. (\<forall>j. x j \<le> p) \<and> x j = 0 \<longrightarrow> lab x j = 0"
       
  1197     and lab_1: "\<forall>x. \<forall>j\<le>n. (\<forall>j. x j \<le> p) \<and> x j = p \<longrightarrow> lab x j = 1"
       
  1198     and odd: "odd (card {s. ksimplex p n s \<and> (reduced n\<circ>lab) ` s = {..n}})"
       
  1199   shows "odd (card {s. ksimplex p (Suc n) s \<and> (reduced (Suc n)\<circ>lab) ` s = {..Suc n}})"
       
  1200 proof -
       
  1201   let ?rl = "reduced (Suc n) \<circ> lab" and ?ext = "\<lambda>f v. \<exists>j\<le>n. \<forall>x\<in>f. x j = v"
       
  1202   let ?ext = "\<lambda>s. (\<exists>j\<le>n. \<forall>x\<in>s. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>s. x j = p)"
       
  1203   have "\<forall>s. ksimplex p (Suc n) s \<longrightarrow> ?rl ` s \<subseteq> {..Suc n}"
       
  1204     by (simp add: reduced_labelling subset_eq)
       
  1205   moreover
       
  1206   have "{s. ksimplex p n s \<and> (reduced n \<circ> lab) ` s = {..n}} =
       
  1207         {f. \<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> f = s - {a} \<and> ?rl ` f = {..n} \<and> ?ext f}"
       
  1208   proof (intro set_eqI, safe del: disjCI equalityI disjE)
       
  1209     fix s assume s: "ksimplex p n s" and rl: "(reduced n \<circ> lab) ` s = {..n}"
       
  1210     from s obtain u b where "kuhn_simplex p n u b s" by (auto elim: ksimplex.cases)
       
  1211     then interpret kuhn_simplex p n u b s .
       
  1212     have all_eq_p: "\<forall>x\<in>s. x n = p"
       
  1213       by (auto simp: out_eq_p)
       
  1214     moreover
       
  1215     { fix x assume "x \<in> s"
       
  1216       with lab_1[rule_format, of n x] all_eq_p s_le_p[of x]
       
  1217       have "?rl x \<le> n"
       
  1218         by (auto intro!: reduced_labelling_nonzero)
       
  1219       then have "?rl x = reduced n (lab x)"
       
  1220         by (auto intro!: reduced_labelling_Suc) }
       
  1221     then have "?rl ` s = {..n}"
       
  1222       using rl by (simp cong: image_cong)
       
  1223     moreover
       
  1224     obtain t a where "ksimplex p (Suc n) t" "a \<in> t" "s = t - {a}"
       
  1225       using s unfolding simplex_top_face[OF \<open>0 < p\<close> all_eq_p] by auto
       
  1226     ultimately
       
  1227     show "\<exists>t a. ksimplex p (Suc n) t \<and> a \<in> t \<and> s = t - {a} \<and> ?rl ` s = {..n} \<and> ?ext s"
       
  1228       by auto
       
  1229   next
       
  1230     fix x s a assume s: "ksimplex p (Suc n) s" and rl: "?rl ` (s - {a}) = {.. n}"
       
  1231       and a: "a \<in> s" and "?ext (s - {a})"
       
  1232     from s obtain u b where "kuhn_simplex p (Suc n) u b s" by (auto elim: ksimplex.cases)
       
  1233     then interpret kuhn_simplex p "Suc n" u b s .
       
  1234     have all_eq_p: "\<forall>x\<in>s. x (Suc n) = p"
       
  1235       by (auto simp: out_eq_p)
       
  1236 
       
  1237     { fix x assume "x \<in> s - {a}"
       
  1238       then have "?rl x \<in> ?rl ` (s - {a})"
       
  1239         by auto
       
  1240       then have "?rl x \<le> n"
       
  1241         unfolding rl by auto
       
  1242       then have "?rl x = reduced n (lab x)"
       
  1243         by (auto intro!: reduced_labelling_Suc) }
       
  1244     then show rl': "(reduced n\<circ>lab) ` (s - {a}) = {..n}"
       
  1245       unfolding rl[symmetric] by (intro image_cong) auto
       
  1246 
       
  1247     from \<open>?ext (s - {a})\<close>
       
  1248     have all_eq_p: "\<forall>x\<in>s - {a}. x n = p"
       
  1249     proof (elim disjE exE conjE)
       
  1250       fix j assume "j \<le> n" "\<forall>x\<in>s - {a}. x j = 0"
       
  1251       with lab_0[rule_format, of j] all_eq_p s_le_p
       
  1252       have "\<And>x. x \<in> s - {a} \<Longrightarrow> reduced (Suc n) (lab x) \<noteq> j"
       
  1253         by (intro reduced_labelling_zero) auto
       
  1254       moreover have "j \<in> ?rl ` (s - {a})"
       
  1255         using \<open>j \<le> n\<close> unfolding rl by auto
       
  1256       ultimately show ?thesis
       
  1257         by force
       
  1258     next
       
  1259       fix j assume "j \<le> n" and eq_p: "\<forall>x\<in>s - {a}. x j = p"
       
  1260       show ?thesis
       
  1261       proof cases
       
  1262         assume "j = n" with eq_p show ?thesis by simp
       
  1263       next
       
  1264         assume "j \<noteq> n"
       
  1265         { fix x assume x: "x \<in> s - {a}"
       
  1266           have "reduced n (lab x) \<le> j"
       
  1267           proof (rule reduced_labelling_nonzero)
       
  1268             show "lab x j \<noteq> 0"
       
  1269               using lab_1[rule_format, of j x] x s_le_p[of x] eq_p \<open>j \<le> n\<close> by auto
       
  1270             show "j < n"
       
  1271               using \<open>j \<le> n\<close> \<open>j \<noteq> n\<close> by simp
       
  1272           qed
       
  1273           then have "reduced n (lab x) \<noteq> n"
       
  1274             using \<open>j \<le> n\<close> \<open>j \<noteq> n\<close> by simp }
       
  1275         moreover have "n \<in> (reduced n\<circ>lab) ` (s - {a})"
       
  1276           unfolding rl' by auto
       
  1277         ultimately show ?thesis
       
  1278           by force
       
  1279       qed
       
  1280     qed
       
  1281     show "ksimplex p n (s - {a})"
       
  1282       unfolding simplex_top_face[OF \<open>0 < p\<close> all_eq_p] using s a by auto
       
  1283   qed
       
  1284   ultimately show ?thesis
       
  1285     using assms by (intro kuhn_simplex_lemma) auto
       
  1286 qed
       
  1287 
       
  1288 text \<open>And so we get the final combinatorial result.\<close>
       
  1289 
       
  1290 lemma ksimplex_0: "ksimplex p 0 s \<longleftrightarrow> s = {(\<lambda>x. p)}"
       
  1291 proof
       
  1292   assume "ksimplex p 0 s" then show "s = {(\<lambda>x. p)}"
       
  1293     by (blast dest: kuhn_simplex.ksimplex_0 elim: ksimplex.cases)
       
  1294 next
       
  1295   assume s: "s = {(\<lambda>x. p)}"
       
  1296   show "ksimplex p 0 s"
       
  1297   proof (intro ksimplex, unfold_locales)
       
  1298     show "(\<lambda>_. p) \<in> {..<0::nat} \<rightarrow> {..<p}" by auto
       
  1299     show "bij_betw id {..<0} {..<0}"
       
  1300       by simp
       
  1301   qed (auto simp: s)
       
  1302 qed
       
  1303 
       
  1304 lemma kuhn_combinatorial:
       
  1305   assumes "0 < p"
       
  1306     and "\<forall>x j. (\<forall>j. x j \<le> p) \<and> j < n \<and> x j = 0 \<longrightarrow> lab x j = 0"
       
  1307     and "\<forall>x j. (\<forall>j. x j \<le> p) \<and> j < n  \<and> x j = p \<longrightarrow> lab x j = 1"
       
  1308   shows "odd (card {s. ksimplex p n s \<and> (reduced n\<circ>lab) ` s = {..n}})"
       
  1309     (is "odd (card (?M n))")
       
  1310   using assms
       
  1311 proof (induct n)
       
  1312   case 0 then show ?case
       
  1313     by (simp add: ksimplex_0 cong: conj_cong)
       
  1314 next
       
  1315   case (Suc n)
       
  1316   then have "odd (card (?M n))"
       
  1317     by force
       
  1318   with Suc show ?case
       
  1319     using kuhn_induction[of p n] by (auto simp: comp_def)
       
  1320 qed
       
  1321 
       
  1322 lemma kuhn_lemma:
       
  1323   fixes n p :: nat
       
  1324   assumes "0 < p"
       
  1325     and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. label x i = (0::nat) \<or> label x i = 1)"
       
  1326     and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = 0 \<longrightarrow> label x i = 0)"
       
  1327     and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = p \<longrightarrow> label x i = 1)"
       
  1328   obtains q where "\<forall>i<n. q i < p"
       
  1329     and "\<forall>i<n. \<exists>r s. (\<forall>j<n. q j \<le> r j \<and> r j \<le> q j + 1) \<and> (\<forall>j<n. q j \<le> s j \<and> s j \<le> q j + 1) \<and> label r i \<noteq> label s i"
       
  1330 proof -
       
  1331   let ?rl = "reduced n \<circ> label"
       
  1332   let ?A = "{s. ksimplex p n s \<and> ?rl ` s = {..n}}"
       
  1333   have "odd (card ?A)"
       
  1334     using assms by (intro kuhn_combinatorial[of p n label]) auto
       
  1335   then have "?A \<noteq> {}"
       
  1336     by fastforce
       
  1337   then obtain s b u where "kuhn_simplex p n b u s" and rl: "?rl ` s = {..n}"
       
  1338     by (auto elim: ksimplex.cases)
       
  1339   interpret kuhn_simplex p n b u s by fact
       
  1340 
       
  1341   show ?thesis
       
  1342   proof (intro that[of b] allI impI)
       
  1343     fix i
       
  1344     assume "i < n"
       
  1345     then show "b i < p"
       
  1346       using base by auto
       
  1347   next
       
  1348     fix i
       
  1349     assume "i < n"
       
  1350     then have "i \<in> {.. n}" "Suc i \<in> {.. n}"
       
  1351       by auto
       
  1352     then obtain u v where u: "u \<in> s" "Suc i = ?rl u" and v: "v \<in> s" "i = ?rl v"
       
  1353       unfolding rl[symmetric] by blast
       
  1354 
       
  1355     have "label u i \<noteq> label v i"
       
  1356       using reduced_labelling [of n "label u"] reduced_labelling [of n "label v"]
       
  1357         u(2)[symmetric] v(2)[symmetric] \<open>i < n\<close>
       
  1358       by auto
       
  1359     moreover
       
  1360     have "b j \<le> u j" "u j \<le> b j + 1" "b j \<le> v j" "v j \<le> b j + 1" if "j < n" for j
       
  1361       using that base_le[OF \<open>u\<in>s\<close>] le_Suc_base[OF \<open>u\<in>s\<close>] base_le[OF \<open>v\<in>s\<close>] le_Suc_base[OF \<open>v\<in>s\<close>]
       
  1362       by auto
       
  1363     ultimately show "\<exists>r s. (\<forall>j<n. b j \<le> r j \<and> r j \<le> b j + 1) \<and>
       
  1364         (\<forall>j<n. b j \<le> s j \<and> s j \<le> b j + 1) \<and> label r i \<noteq> label s i"
       
  1365       by blast
       
  1366   qed
       
  1367 qed
       
  1368 
       
  1369 subsection \<open>The main result for the unit cube\<close>
       
  1370 
       
  1371 lemma kuhn_labelling_lemma':
       
  1372   assumes "(\<forall>x::nat\<Rightarrow>real. P x \<longrightarrow> P (f x))"
       
  1373     and "\<forall>x. P x \<longrightarrow> (\<forall>i::nat. Q i \<longrightarrow> 0 \<le> x i \<and> x i \<le> 1)"
       
  1374   shows "\<exists>l. (\<forall>x i. l x i \<le> (1::nat)) \<and>
       
  1375              (\<forall>x i. P x \<and> Q i \<and> x i = 0 \<longrightarrow> l x i = 0) \<and>
       
  1376              (\<forall>x i. P x \<and> Q i \<and> x i = 1 \<longrightarrow> l x i = 1) \<and>
       
  1377              (\<forall>x i. P x \<and> Q i \<and> l x i = 0 \<longrightarrow> x i \<le> f x i) \<and>
       
  1378              (\<forall>x i. P x \<and> Q i \<and> l x i = 1 \<longrightarrow> f x i \<le> x i)"
       
  1379 proof -
       
  1380   have and_forall_thm: "\<And>P Q. (\<forall>x. P x) \<and> (\<forall>x. Q x) \<longleftrightarrow> (\<forall>x. P x \<and> Q x)"
       
  1381     by auto
       
  1382   have *: "\<forall>x y::real. 0 \<le> x \<and> x \<le> 1 \<and> 0 \<le> y \<and> y \<le> 1 \<longrightarrow> x \<noteq> 1 \<and> x \<le> y \<or> x \<noteq> 0 \<and> y \<le> x"
       
  1383     by auto
       
  1384   show ?thesis
       
  1385     unfolding and_forall_thm
       
  1386     apply (subst choice_iff[symmetric])+
       
  1387     apply rule
       
  1388     apply rule
       
  1389   proof -
       
  1390     fix x x'
       
  1391     let ?R = "\<lambda>y::nat.
       
  1392       (P x \<and> Q x' \<and> x x' = 0 \<longrightarrow> y = 0) \<and>
       
  1393       (P x \<and> Q x' \<and> x x' = 1 \<longrightarrow> y = 1) \<and>
       
  1394       (P x \<and> Q x' \<and> y = 0 \<longrightarrow> x x' \<le> (f x) x') \<and>
       
  1395       (P x \<and> Q x' \<and> y = 1 \<longrightarrow> (f x) x' \<le> x x')"
       
  1396     have "0 \<le> f x x' \<and> f x x' \<le> 1" if "P x" "Q x'"
       
  1397       using assms(2)[rule_format,of "f x" x'] that
       
  1398       apply (drule_tac assms(1)[rule_format])
       
  1399       apply auto
       
  1400       done
       
  1401     then have "?R 0 \<or> ?R 1"
       
  1402       by auto
       
  1403     then show "\<exists>y\<le>1. ?R y"
       
  1404       by auto
       
  1405   qed
       
  1406 qed
       
  1407 
       
  1408 definition unit_cube :: "'a::euclidean_space set"
       
  1409   where "unit_cube = {x. \<forall>i\<in>Basis. 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1}"
       
  1410 
       
  1411 lemma mem_unit_cube: "x \<in> unit_cube \<longleftrightarrow> (\<forall>i\<in>Basis. 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)"
       
  1412   unfolding unit_cube_def by simp
       
  1413 
       
  1414 lemma bounded_unit_cube: "bounded unit_cube"
       
  1415   unfolding bounded_def
       
  1416 proof (intro exI ballI)
       
  1417   fix y :: 'a assume y: "y \<in> unit_cube"
       
  1418   have "dist 0 y = norm y" by (rule dist_0_norm)
       
  1419   also have "\<dots> = norm (\<Sum>i\<in>Basis. (y \<bullet> i) *\<^sub>R i)" unfolding euclidean_representation ..
       
  1420   also have "\<dots> \<le> (\<Sum>i\<in>Basis. norm ((y \<bullet> i) *\<^sub>R i))" by (rule norm_setsum)
       
  1421   also have "\<dots> \<le> (\<Sum>i::'a\<in>Basis. 1)"
       
  1422     by (rule setsum_mono, simp add: y [unfolded mem_unit_cube])
       
  1423   finally show "dist 0 y \<le> (\<Sum>i::'a\<in>Basis. 1)" .
       
  1424 qed
       
  1425 
       
  1426 lemma closed_unit_cube: "closed unit_cube"
       
  1427   unfolding unit_cube_def Collect_ball_eq Collect_conj_eq
       
  1428   by (rule closed_INT, auto intro!: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id)
       
  1429 
       
  1430 lemma compact_unit_cube: "compact unit_cube" (is "compact ?C")
       
  1431   unfolding compact_eq_seq_compact_metric
       
  1432   using bounded_unit_cube closed_unit_cube
       
  1433   by (rule bounded_closed_imp_seq_compact)
       
  1434 
       
  1435 lemma brouwer_cube:
       
  1436   fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
       
  1437   assumes "continuous_on unit_cube f"
       
  1438     and "f ` unit_cube \<subseteq> unit_cube"
       
  1439   shows "\<exists>x\<in>unit_cube. f x = x"
       
  1440 proof (rule ccontr)
       
  1441   define n where "n = DIM('a)"
       
  1442   have n: "1 \<le> n" "0 < n" "n \<noteq> 0"
       
  1443     unfolding n_def by (auto simp add: Suc_le_eq DIM_positive)
       
  1444   assume "\<not> ?thesis"
       
  1445   then have *: "\<not> (\<exists>x\<in>unit_cube. f x - x = 0)"
       
  1446     by auto
       
  1447   obtain d where
       
  1448       d: "d > 0" "\<And>x. x \<in> unit_cube \<Longrightarrow> d \<le> norm (f x - x)"
       
  1449     apply (rule brouwer_compactness_lemma[OF compact_unit_cube _ *])
       
  1450     apply (rule continuous_intros assms)+
       
  1451     apply blast
       
  1452     done
       
  1453   have *: "\<forall>x. x \<in> unit_cube \<longrightarrow> f x \<in> unit_cube"
       
  1454     "\<forall>x. x \<in> (unit_cube::'a set) \<longrightarrow> (\<forall>i\<in>Basis. True \<longrightarrow> 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)"
       
  1455     using assms(2)[unfolded image_subset_iff Ball_def]
       
  1456     unfolding mem_unit_cube
       
  1457     by auto
       
  1458   obtain label :: "'a \<Rightarrow> 'a \<Rightarrow> nat" where
       
  1459     "\<forall>x. \<forall>i\<in>Basis. label x i \<le> 1"
       
  1460     "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> x \<bullet> i = 0 \<longrightarrow> label x i = 0"
       
  1461     "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> x \<bullet> i = 1 \<longrightarrow> label x i = 1"
       
  1462     "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> label x i = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i"
       
  1463     "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> label x i = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i"
       
  1464     using kuhn_labelling_lemma[OF *] by blast
       
  1465   note label = this [rule_format]
       
  1466   have lem1: "\<forall>x\<in>unit_cube. \<forall>y\<in>unit_cube. \<forall>i\<in>Basis. label x i \<noteq> label y i \<longrightarrow>
       
  1467     \<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)"
       
  1468   proof safe
       
  1469     fix x y :: 'a
       
  1470     assume x: "x \<in> unit_cube"
       
  1471     assume y: "y \<in> unit_cube"
       
  1472     fix i
       
  1473     assume i: "label x i \<noteq> label y i" "i \<in> Basis"
       
  1474     have *: "\<And>x y fx fy :: real. x \<le> fx \<and> fy \<le> y \<or> fx \<le> x \<and> y \<le> fy \<Longrightarrow>
       
  1475       \<bar>fx - x\<bar> \<le> \<bar>fy - fx\<bar> + \<bar>y - x\<bar>" by auto
       
  1476     have "\<bar>(f x - x) \<bullet> i\<bar> \<le> \<bar>(f y - f x)\<bullet>i\<bar> + \<bar>(y - x)\<bullet>i\<bar>"
       
  1477       unfolding inner_simps
       
  1478       apply (rule *)
       
  1479       apply (cases "label x i = 0")
       
  1480       apply (rule disjI1)
       
  1481       apply rule
       
  1482       prefer 3
       
  1483       apply (rule disjI2)
       
  1484       apply rule
       
  1485     proof -
       
  1486       assume lx: "label x i = 0"
       
  1487       then have ly: "label y i = 1"
       
  1488         using i label(1)[of i y]
       
  1489         by auto
       
  1490       show "x \<bullet> i \<le> f x \<bullet> i"
       
  1491         apply (rule label(4)[rule_format])
       
  1492         using x y lx i(2)
       
  1493         apply auto
       
  1494         done
       
  1495       show "f y \<bullet> i \<le> y \<bullet> i"
       
  1496         apply (rule label(5)[rule_format])
       
  1497         using x y ly i(2)
       
  1498         apply auto
       
  1499         done
       
  1500     next
       
  1501       assume "label x i \<noteq> 0"
       
  1502       then have l: "label x i = 1" "label y i = 0"
       
  1503         using i label(1)[of i x] label(1)[of i y]
       
  1504         by auto
       
  1505       show "f x \<bullet> i \<le> x \<bullet> i"
       
  1506         apply (rule label(5)[rule_format])
       
  1507         using x y l i(2)
       
  1508         apply auto
       
  1509         done
       
  1510       show "y \<bullet> i \<le> f y \<bullet> i"
       
  1511         apply (rule label(4)[rule_format])
       
  1512         using x y l i(2)
       
  1513         apply auto
       
  1514         done
       
  1515     qed
       
  1516     also have "\<dots> \<le> norm (f y - f x) + norm (y - x)"
       
  1517       apply (rule add_mono)
       
  1518       apply (rule Basis_le_norm[OF i(2)])+
       
  1519       done
       
  1520     finally show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)"
       
  1521       unfolding inner_simps .
       
  1522   qed
       
  1523   have "\<exists>e>0. \<forall>x\<in>unit_cube. \<forall>y\<in>unit_cube. \<forall>z\<in>unit_cube. \<forall>i\<in>Basis.
       
  1524     norm (x - z) < e \<and> norm (y - z) < e \<and> label x i \<noteq> label y i \<longrightarrow>
       
  1525       \<bar>(f(z) - z)\<bullet>i\<bar> < d / (real n)"
       
  1526   proof -
       
  1527     have d': "d / real n / 8 > 0"
       
  1528       using d(1) by (simp add: n_def DIM_positive)
       
  1529     have *: "uniformly_continuous_on unit_cube f"
       
  1530       by (rule compact_uniformly_continuous[OF assms(1) compact_unit_cube])
       
  1531     obtain e where e:
       
  1532         "e > 0"
       
  1533         "\<And>x x'. x \<in> unit_cube \<Longrightarrow>
       
  1534           x' \<in> unit_cube \<Longrightarrow>
       
  1535           norm (x' - x) < e \<Longrightarrow>
       
  1536           norm (f x' - f x) < d / real n / 8"
       
  1537       using *[unfolded uniformly_continuous_on_def,rule_format,OF d']
       
  1538       unfolding dist_norm
       
  1539       by blast
       
  1540     show ?thesis
       
  1541       apply (rule_tac x="min (e/2) (d/real n/8)" in exI)
       
  1542       apply safe
       
  1543     proof -
       
  1544       show "0 < min (e / 2) (d / real n / 8)"
       
  1545         using d' e by auto
       
  1546       fix x y z i
       
  1547       assume as:
       
  1548         "x \<in> unit_cube" "y \<in> unit_cube" "z \<in> unit_cube"
       
  1549         "norm (x - z) < min (e / 2) (d / real n / 8)"
       
  1550         "norm (y - z) < min (e / 2) (d / real n / 8)"
       
  1551         "label x i \<noteq> label y i"
       
  1552       assume i: "i \<in> Basis"
       
  1553       have *: "\<And>z fz x fx n1 n2 n3 n4 d4 d :: real. \<bar>fx - x\<bar> \<le> n1 + n2 \<Longrightarrow>
       
  1554         \<bar>fx - fz\<bar> \<le> n3 \<Longrightarrow> \<bar>x - z\<bar> \<le> n4 \<Longrightarrow>
       
  1555         n1 < d4 \<Longrightarrow> n2 < 2 * d4 \<Longrightarrow> n3 < d4 \<Longrightarrow> n4 < d4 \<Longrightarrow>
       
  1556         (8 * d4 = d) \<Longrightarrow> \<bar>fz - z\<bar> < d"
       
  1557         by auto
       
  1558       show "\<bar>(f z - z) \<bullet> i\<bar> < d / real n"
       
  1559         unfolding inner_simps
       
  1560       proof (rule *)
       
  1561         show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y -f x) + norm (y - x)"
       
  1562           apply (rule lem1[rule_format])
       
  1563           using as i
       
  1564           apply auto
       
  1565           done
       
  1566         show "\<bar>f x \<bullet> i - f z \<bullet> i\<bar> \<le> norm (f x - f z)" "\<bar>x \<bullet> i - z \<bullet> i\<bar> \<le> norm (x - z)"
       
  1567           unfolding inner_diff_left[symmetric]
       
  1568           by (rule Basis_le_norm[OF i])+
       
  1569         have tria: "norm (y - x) \<le> norm (y - z) + norm (x - z)"
       
  1570           using dist_triangle[of y x z, unfolded dist_norm]
       
  1571           unfolding norm_minus_commute
       
  1572           by auto
       
  1573         also have "\<dots> < e / 2 + e / 2"
       
  1574           apply (rule add_strict_mono)
       
  1575           using as(4,5)
       
  1576           apply auto
       
  1577           done
       
  1578         finally show "norm (f y - f x) < d / real n / 8"
       
  1579           apply -
       
  1580           apply (rule e(2))
       
  1581           using as
       
  1582           apply auto
       
  1583           done
       
  1584         have "norm (y - z) + norm (x - z) < d / real n / 8 + d / real n / 8"
       
  1585           apply (rule add_strict_mono)
       
  1586           using as
       
  1587           apply auto
       
  1588           done
       
  1589         then show "norm (y - x) < 2 * (d / real n / 8)"
       
  1590           using tria
       
  1591           by auto
       
  1592         show "norm (f x - f z) < d / real n / 8"
       
  1593           apply (rule e(2))
       
  1594           using as e(1)
       
  1595           apply auto
       
  1596           done
       
  1597       qed (insert as, auto)
       
  1598     qed
       
  1599   qed
       
  1600   then
       
  1601   obtain e where e:
       
  1602     "e > 0"
       
  1603     "\<And>x y z i. x \<in> unit_cube \<Longrightarrow>
       
  1604       y \<in> unit_cube \<Longrightarrow>
       
  1605       z \<in> unit_cube \<Longrightarrow>
       
  1606       i \<in> Basis \<Longrightarrow>
       
  1607       norm (x - z) < e \<and> norm (y - z) < e \<and> label x i \<noteq> label y i \<Longrightarrow>
       
  1608       \<bar>(f z - z) \<bullet> i\<bar> < d / real n"
       
  1609     by blast
       
  1610   obtain p :: nat where p: "1 + real n / e \<le> real p"
       
  1611     using real_arch_simple ..
       
  1612   have "1 + real n / e > 0"
       
  1613     using e(1) n by (simp add: add_pos_pos)
       
  1614   then have "p > 0"
       
  1615     using p by auto
       
  1616 
       
  1617   obtain b :: "nat \<Rightarrow> 'a" where b: "bij_betw b {..< n} Basis"
       
  1618     by atomize_elim (auto simp: n_def intro!: finite_same_card_bij)
       
  1619   define b' where "b' = inv_into {..< n} b"
       
  1620   then have b': "bij_betw b' Basis {..< n}"
       
  1621     using bij_betw_inv_into[OF b] by auto
       
  1622   then have b'_Basis: "\<And>i. i \<in> Basis \<Longrightarrow> b' i \<in> {..< n}"
       
  1623     unfolding bij_betw_def by (auto simp: set_eq_iff)
       
  1624   have bb'[simp]:"\<And>i. i \<in> Basis \<Longrightarrow> b (b' i) = i"
       
  1625     unfolding b'_def
       
  1626     using b
       
  1627     by (auto simp: f_inv_into_f bij_betw_def)
       
  1628   have b'b[simp]:"\<And>i. i < n \<Longrightarrow> b' (b i) = i"
       
  1629     unfolding b'_def
       
  1630     using b
       
  1631     by (auto simp: inv_into_f_eq bij_betw_def)
       
  1632   have *: "\<And>x :: nat. x = 0 \<or> x = 1 \<longleftrightarrow> x \<le> 1"
       
  1633     by auto
       
  1634   have b'': "\<And>j. j < n \<Longrightarrow> b j \<in> Basis"
       
  1635     using b unfolding bij_betw_def by auto
       
  1636   have q1: "0 < p" "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow>
       
  1637     (\<forall>i<n. (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0 \<or>
       
  1638            (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)"
       
  1639     unfolding *
       
  1640     using \<open>p > 0\<close> \<open>n > 0\<close>
       
  1641     using label(1)[OF b'']
       
  1642     by auto
       
  1643   { fix x :: "nat \<Rightarrow> nat" and i assume "\<forall>i<n. x i \<le> p" "i < n" "x i = p \<or> x i = 0"
       
  1644     then have "(\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<in> (unit_cube::'a set)"
       
  1645       using b'_Basis
       
  1646       by (auto simp add: mem_unit_cube inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) }
       
  1647   note cube = this
       
  1648   have q2: "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = 0 \<longrightarrow>
       
  1649       (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0)"
       
  1650     unfolding o_def using cube \<open>p > 0\<close> by (intro allI impI label(2)) (auto simp add: b'')
       
  1651   have q3: "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = p \<longrightarrow>
       
  1652       (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)"
       
  1653     using cube \<open>p > 0\<close> unfolding o_def by (intro allI impI label(3)) (auto simp add: b'')
       
  1654   obtain q where q:
       
  1655       "\<forall>i<n. q i < p"
       
  1656       "\<forall>i<n.
       
  1657          \<exists>r s. (\<forall>j<n. q j \<le> r j \<and> r j \<le> q j + 1) \<and>
       
  1658                (\<forall>j<n. q j \<le> s j \<and> s j \<le> q j + 1) \<and>
       
  1659                (label (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i) \<circ> b) i \<noteq>
       
  1660                (label (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i) \<circ> b) i"
       
  1661     by (rule kuhn_lemma[OF q1 q2 q3])
       
  1662   define z :: 'a where "z = (\<Sum>i\<in>Basis. (real (q (b' i)) / real p) *\<^sub>R i)"
       
  1663   have "\<exists>i\<in>Basis. d / real n \<le> \<bar>(f z - z)\<bullet>i\<bar>"
       
  1664   proof (rule ccontr)
       
  1665     have "\<forall>i\<in>Basis. q (b' i) \<in> {0..p}"
       
  1666       using q(1) b'
       
  1667       by (auto intro: less_imp_le simp: bij_betw_def)
       
  1668     then have "z \<in> unit_cube"
       
  1669       unfolding z_def mem_unit_cube
       
  1670       using b'_Basis
       
  1671       by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1)
       
  1672     then have d_fz_z: "d \<le> norm (f z - z)"
       
  1673       by (rule d)
       
  1674     assume "\<not> ?thesis"
       
  1675     then have as: "\<forall>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar> < d / real n"
       
  1676       using \<open>n > 0\<close>
       
  1677       by (auto simp add: not_le inner_diff)
       
  1678     have "norm (f z - z) \<le> (\<Sum>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar>)"
       
  1679       unfolding inner_diff_left[symmetric]
       
  1680       by (rule norm_le_l1)
       
  1681     also have "\<dots> < (\<Sum>(i::'a) \<in> Basis. d / real n)"
       
  1682       apply (rule setsum_strict_mono)
       
  1683       using as
       
  1684       apply auto
       
  1685       done
       
  1686     also have "\<dots> = d"
       
  1687       using DIM_positive[where 'a='a]
       
  1688       by (auto simp: n_def)
       
  1689     finally show False
       
  1690       using d_fz_z by auto
       
  1691   qed
       
  1692   then obtain i where i: "i \<in> Basis" "d / real n \<le> \<bar>(f z - z) \<bullet> i\<bar>" ..
       
  1693   have *: "b' i < n"
       
  1694     using i and b'[unfolded bij_betw_def]
       
  1695     by auto
       
  1696   obtain r s where rs:
       
  1697     "\<And>j. j < n \<Longrightarrow> q j \<le> r j \<and> r j \<le> q j + 1"
       
  1698     "\<And>j. j < n \<Longrightarrow> q j \<le> s j \<and> s j \<le> q j + 1"
       
  1699     "(label (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i) \<circ> b) (b' i) \<noteq>
       
  1700       (label (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i) \<circ> b) (b' i)"
       
  1701     using q(2)[rule_format,OF *] by blast
       
  1702   have b'_im: "\<And>i. i \<in> Basis \<Longrightarrow>  b' i < n"
       
  1703     using b' unfolding bij_betw_def by auto
       
  1704   define r' ::'a where "r' = (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i)"
       
  1705   have "\<And>i. i \<in> Basis \<Longrightarrow> r (b' i) \<le> p"
       
  1706     apply (rule order_trans)
       
  1707     apply (rule rs(1)[OF b'_im,THEN conjunct2])
       
  1708     using q(1)[rule_format,OF b'_im]
       
  1709     apply (auto simp add: Suc_le_eq)
       
  1710     done
       
  1711   then have "r' \<in> unit_cube"
       
  1712     unfolding r'_def mem_unit_cube
       
  1713     using b'_Basis
       
  1714     by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1)
       
  1715   define s' :: 'a where "s' = (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i)"
       
  1716   have "\<And>i. i \<in> Basis \<Longrightarrow> s (b' i) \<le> p"
       
  1717     apply (rule order_trans)
       
  1718     apply (rule rs(2)[OF b'_im, THEN conjunct2])
       
  1719     using q(1)[rule_format,OF b'_im]
       
  1720     apply (auto simp add: Suc_le_eq)
       
  1721     done
       
  1722   then have "s' \<in> unit_cube"
       
  1723     unfolding s'_def mem_unit_cube
       
  1724     using b'_Basis
       
  1725     by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1)
       
  1726   have "z \<in> unit_cube"
       
  1727     unfolding z_def mem_unit_cube
       
  1728     using b'_Basis q(1)[rule_format,OF b'_im] \<open>p > 0\<close>
       
  1729     by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le)
       
  1730   have *: "\<And>x. 1 + real x = real (Suc x)"
       
  1731     by auto
       
  1732   {
       
  1733     have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)"
       
  1734       apply (rule setsum_mono)
       
  1735       using rs(1)[OF b'_im]
       
  1736       apply (auto simp add:* field_simps simp del: of_nat_Suc)
       
  1737       done
       
  1738     also have "\<dots> < e * real p"
       
  1739       using p \<open>e > 0\<close> \<open>p > 0\<close>
       
  1740       by (auto simp add: field_simps n_def)
       
  1741     finally have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) < e * real p" .
       
  1742   }
       
  1743   moreover
       
  1744   {
       
  1745     have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)"
       
  1746       apply (rule setsum_mono)
       
  1747       using rs(2)[OF b'_im]
       
  1748       apply (auto simp add:* field_simps simp del: of_nat_Suc)
       
  1749       done
       
  1750     also have "\<dots> < e * real p"
       
  1751       using p \<open>e > 0\<close> \<open>p > 0\<close>
       
  1752       by (auto simp add: field_simps n_def)
       
  1753     finally have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) < e * real p" .
       
  1754   }
       
  1755   ultimately
       
  1756   have "norm (r' - z) < e" and "norm (s' - z) < e"
       
  1757     unfolding r'_def s'_def z_def
       
  1758     using \<open>p > 0\<close>
       
  1759     apply (rule_tac[!] le_less_trans[OF norm_le_l1])
       
  1760     apply (auto simp add: field_simps setsum_divide_distrib[symmetric] inner_diff_left)
       
  1761     done
       
  1762   then have "\<bar>(f z - z) \<bullet> i\<bar> < d / real n"
       
  1763     using rs(3) i
       
  1764     unfolding r'_def[symmetric] s'_def[symmetric] o_def bb'
       
  1765     by (intro e(2)[OF \<open>r'\<in>unit_cube\<close> \<open>s'\<in>unit_cube\<close> \<open>z\<in>unit_cube\<close>]) auto
       
  1766   then show False
       
  1767     using i by auto
       
  1768 qed
       
  1769 
       
  1770 
       
  1771 subsection \<open>Retractions\<close>
       
  1772 
       
  1773 definition "retraction s t r \<longleftrightarrow> t \<subseteq> s \<and> continuous_on s r \<and> r ` s \<subseteq> t \<and> (\<forall>x\<in>t. r x = x)"
       
  1774 
       
  1775 definition retract_of (infixl "retract'_of" 50)
       
  1776   where "(t retract_of s) \<longleftrightarrow> (\<exists>r. retraction s t r)"
       
  1777 
       
  1778 lemma retraction_idempotent: "retraction s t r \<Longrightarrow> x \<in> s \<Longrightarrow>  r (r x) = r x"
       
  1779   unfolding retraction_def by auto
       
  1780 
       
  1781 subsection \<open>Preservation of fixpoints under (more general notion of) retraction\<close>
       
  1782 
       
  1783 lemma invertible_fixpoint_property:
       
  1784   fixes s :: "'a::euclidean_space set"
       
  1785     and t :: "'b::euclidean_space set"
       
  1786   assumes "continuous_on t i"
       
  1787     and "i ` t \<subseteq> s"
       
  1788     and "continuous_on s r"
       
  1789     and "r ` s \<subseteq> t"
       
  1790     and "\<forall>y\<in>t. r (i y) = y"
       
  1791     and "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)"
       
  1792     and "continuous_on t g"
       
  1793     and "g ` t \<subseteq> t"
       
  1794   obtains y where "y \<in> t" and "g y = y"
       
  1795 proof -
       
  1796   have "\<exists>x\<in>s. (i \<circ> g \<circ> r) x = x"
       
  1797     apply (rule assms(6)[rule_format])
       
  1798     apply rule
       
  1799     apply (rule continuous_on_compose assms)+
       
  1800     apply ((rule continuous_on_subset)?, rule assms)+
       
  1801     using assms(2,4,8)
       
  1802     apply auto
       
  1803     apply blast
       
  1804     done
       
  1805   then obtain x where x: "x \<in> s" "(i \<circ> g \<circ> r) x = x" ..
       
  1806   then have *: "g (r x) \<in> t"
       
  1807     using assms(4,8) by auto
       
  1808   have "r ((i \<circ> g \<circ> r) x) = r x"
       
  1809     using x by auto
       
  1810   then show ?thesis
       
  1811     apply (rule_tac that[of "r x"])
       
  1812     using x
       
  1813     unfolding o_def
       
  1814     unfolding assms(5)[rule_format,OF *]
       
  1815     using assms(4)
       
  1816     apply auto
       
  1817     done
       
  1818 qed
       
  1819 
       
  1820 lemma homeomorphic_fixpoint_property:
       
  1821   fixes s :: "'a::euclidean_space set"
       
  1822     and t :: "'b::euclidean_space set"
       
  1823   assumes "s homeomorphic t"
       
  1824   shows "(\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)) \<longleftrightarrow>
       
  1825     (\<forall>g. continuous_on t g \<and> g ` t \<subseteq> t \<longrightarrow> (\<exists>y\<in>t. g y = y))"
       
  1826 proof -
       
  1827   obtain r i where
       
  1828       "\<forall>x\<in>s. i (r x) = x"
       
  1829       "r ` s = t"
       
  1830       "continuous_on s r"
       
  1831       "\<forall>y\<in>t. r (i y) = y"
       
  1832       "i ` t = s"
       
  1833       "continuous_on t i"
       
  1834     using assms
       
  1835     unfolding homeomorphic_def homeomorphism_def
       
  1836     by blast
       
  1837   then show ?thesis
       
  1838     apply -
       
  1839     apply rule
       
  1840     apply (rule_tac[!] allI impI)+
       
  1841     apply (rule_tac g=g in invertible_fixpoint_property[of t i s r])
       
  1842     prefer 10
       
  1843     apply (rule_tac g=f in invertible_fixpoint_property[of s r t i])
       
  1844     apply auto
       
  1845     done
       
  1846 qed
       
  1847 
       
  1848 lemma retract_fixpoint_property:
       
  1849   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
  1850     and s :: "'a set"
       
  1851   assumes "t retract_of s"
       
  1852     and "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)"
       
  1853     and "continuous_on t g"
       
  1854     and "g ` t \<subseteq> t"
       
  1855   obtains y where "y \<in> t" and "g y = y"
       
  1856 proof -
       
  1857   obtain h where "retraction s t h"
       
  1858     using assms(1) unfolding retract_of_def ..
       
  1859   then show ?thesis
       
  1860     unfolding retraction_def
       
  1861     apply -
       
  1862     apply (rule invertible_fixpoint_property[OF continuous_on_id _ _ _ _ assms(2), of t h g])
       
  1863     prefer 7
       
  1864     apply (rule_tac y = y in that)
       
  1865     using assms
       
  1866     apply auto
       
  1867     done
       
  1868 qed
       
  1869 
       
  1870 
       
  1871 subsection \<open>The Brouwer theorem for any set with nonempty interior\<close>
       
  1872 
       
  1873 lemma convex_unit_cube: "convex unit_cube"
       
  1874   apply (rule is_interval_convex)
       
  1875   apply (clarsimp simp add: is_interval_def mem_unit_cube)
       
  1876   apply (drule (1) bspec)+
       
  1877   apply auto
       
  1878   done
       
  1879 
       
  1880 lemma brouwer_weak:
       
  1881   fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
       
  1882   assumes "compact s"
       
  1883     and "convex s"
       
  1884     and "interior s \<noteq> {}"
       
  1885     and "continuous_on s f"
       
  1886     and "f ` s \<subseteq> s"
       
  1887   obtains x where "x \<in> s" and "f x = x"
       
  1888 proof -
       
  1889   let ?U = "unit_cube :: 'a set"
       
  1890   have "\<Sum>Basis /\<^sub>R 2 \<in> interior ?U"
       
  1891   proof (rule interiorI)
       
  1892     let ?I = "(\<Inter>i\<in>Basis. {x::'a. 0 < x \<bullet> i} \<inter> {x. x \<bullet> i < 1})"
       
  1893     show "open ?I"
       
  1894       by (intro open_INT finite_Basis ballI open_Int, auto intro: open_Collect_less simp: continuous_on_inner continuous_on_const continuous_on_id)
       
  1895     show "\<Sum>Basis /\<^sub>R 2 \<in> ?I"
       
  1896       by simp
       
  1897     show "?I \<subseteq> unit_cube"
       
  1898       unfolding unit_cube_def by force
       
  1899   qed
       
  1900   then have *: "interior ?U \<noteq> {}" by fast
       
  1901   have *: "?U homeomorphic s"
       
  1902     using homeomorphic_convex_compact[OF convex_unit_cube compact_unit_cube * assms(2,1,3)] .
       
  1903   have "\<forall>f. continuous_on ?U f \<and> f ` ?U \<subseteq> ?U \<longrightarrow>
       
  1904     (\<exists>x\<in>?U. f x = x)"
       
  1905     using brouwer_cube by auto
       
  1906   then show ?thesis
       
  1907     unfolding homeomorphic_fixpoint_property[OF *]
       
  1908     using assms
       
  1909     by (auto simp: intro: that)
       
  1910 qed
       
  1911 
       
  1912 
       
  1913 text \<open>And in particular for a closed ball.\<close>
       
  1914 
       
  1915 lemma brouwer_ball:
       
  1916   fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
       
  1917   assumes "e > 0"
       
  1918     and "continuous_on (cball a e) f"
       
  1919     and "f ` cball a e \<subseteq> cball a e"
       
  1920   obtains x where "x \<in> cball a e" and "f x = x"
       
  1921   using brouwer_weak[OF compact_cball convex_cball, of a e f]
       
  1922   unfolding interior_cball ball_eq_empty
       
  1923   using assms by auto
       
  1924 
       
  1925 text \<open>Still more general form; could derive this directly without using the
       
  1926   rather involved \<open>HOMEOMORPHIC_CONVEX_COMPACT\<close> theorem, just using
       
  1927   a scaling and translation to put the set inside the unit cube.\<close>
       
  1928 
       
  1929 lemma brouwer:
       
  1930   fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
       
  1931   assumes "compact s"
       
  1932     and "convex s"
       
  1933     and "s \<noteq> {}"
       
  1934     and "continuous_on s f"
       
  1935     and "f ` s \<subseteq> s"
       
  1936   obtains x where "x \<in> s" and "f x = x"
       
  1937 proof -
       
  1938   have "\<exists>e>0. s \<subseteq> cball 0 e"
       
  1939     using compact_imp_bounded[OF assms(1)]
       
  1940     unfolding bounded_pos
       
  1941     apply (erule_tac exE)
       
  1942     apply (rule_tac x=b in exI)
       
  1943     apply (auto simp add: dist_norm)
       
  1944     done
       
  1945   then obtain e where e: "e > 0" "s \<subseteq> cball 0 e"
       
  1946     by blast
       
  1947   have "\<exists>x\<in> cball 0 e. (f \<circ> closest_point s) x = x"
       
  1948     apply (rule_tac brouwer_ball[OF e(1), of 0 "f \<circ> closest_point s"])
       
  1949     apply (rule continuous_on_compose )
       
  1950     apply (rule continuous_on_closest_point[OF assms(2) compact_imp_closed[OF assms(1)] assms(3)])
       
  1951     apply (rule continuous_on_subset[OF assms(4)])
       
  1952     apply (insert closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)])
       
  1953     using assms(5)[unfolded subset_eq]
       
  1954     using e(2)[unfolded subset_eq mem_cball]
       
  1955     apply (auto simp add: dist_norm)
       
  1956     done
       
  1957   then obtain x where x: "x \<in> cball 0 e" "(f \<circ> closest_point s) x = x" ..
       
  1958   have *: "closest_point s x = x"
       
  1959     apply (rule closest_point_self)
       
  1960     apply (rule assms(5)[unfolded subset_eq,THEN bspec[where x="x"], unfolded image_iff])
       
  1961     apply (rule_tac x="closest_point s x" in bexI)
       
  1962     using x
       
  1963     unfolding o_def
       
  1964     using closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3), of x]
       
  1965     apply auto
       
  1966     done
       
  1967   show thesis
       
  1968     apply (rule_tac x="closest_point s x" in that)
       
  1969     unfolding x(2)[unfolded o_def]
       
  1970     apply (rule closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)])
       
  1971     using *
       
  1972     apply auto
       
  1973     done
       
  1974 qed
       
  1975 
       
  1976 text \<open>So we get the no-retraction theorem.\<close>
       
  1977 
       
  1978 lemma no_retraction_cball:
       
  1979   fixes a :: "'a::euclidean_space"
       
  1980   assumes "e > 0"
       
  1981   shows "\<not> (frontier (cball a e) retract_of (cball a e))"
       
  1982 proof
       
  1983   assume *: "frontier (cball a e) retract_of (cball a e)"
       
  1984   have **: "\<And>xa. a - (2 *\<^sub>R a - xa) = - (a - xa)"
       
  1985     using scaleR_left_distrib[of 1 1 a] by auto
       
  1986   obtain x where x:
       
  1987       "x \<in> {x. norm (a - x) = e}"
       
  1988       "2 *\<^sub>R a - x = x"
       
  1989     apply (rule retract_fixpoint_property[OF *, of "\<lambda>x. scaleR 2 a - x"])
       
  1990     apply (blast intro: brouwer_ball[OF assms])
       
  1991     apply (intro continuous_intros)
       
  1992     unfolding frontier_cball subset_eq Ball_def image_iff dist_norm sphere_def
       
  1993     apply (auto simp add: ** norm_minus_commute)
       
  1994     done
       
  1995   then have "scaleR 2 a = scaleR 1 x + scaleR 1 x"
       
  1996     by (auto simp add: algebra_simps)
       
  1997   then have "a = x"
       
  1998     unfolding scaleR_left_distrib[symmetric]
       
  1999     by auto
       
  2000   then show False
       
  2001     using x assms by auto
       
  2002 qed
       
  2003 
       
  2004 subsection\<open>Retractions\<close>
       
  2005 
       
  2006 lemma retraction:
       
  2007    "retraction s t r \<longleftrightarrow>
       
  2008     t \<subseteq> s \<and> continuous_on s r \<and> r ` s = t \<and> (\<forall>x \<in> t. r x = x)"
       
  2009 by (force simp: retraction_def)
       
  2010 
       
  2011 lemma retract_of_imp_extensible:
       
  2012   assumes "s retract_of t" and "continuous_on s f" and "f ` s \<subseteq> u"
       
  2013   obtains g where "continuous_on t g" "g ` t \<subseteq> u" "\<And>x. x \<in> s \<Longrightarrow> g x = f x"
       
  2014 using assms
       
  2015 apply (clarsimp simp add: retract_of_def retraction)
       
  2016 apply (rule_tac g = "f o r" in that)
       
  2017 apply (auto simp: continuous_on_compose2)
       
  2018 done
       
  2019 
       
  2020 lemma idempotent_imp_retraction:
       
  2021   assumes "continuous_on s f" and "f ` s \<subseteq> s" and "\<And>x. x \<in> s \<Longrightarrow> f(f x) = f x"
       
  2022     shows "retraction s (f ` s) f"
       
  2023 by (simp add: assms retraction)
       
  2024 
       
  2025 lemma retraction_subset:
       
  2026   assumes "retraction s t r" and "t \<subseteq> s'" and "s' \<subseteq> s"
       
  2027     shows "retraction s' t r"
       
  2028 apply (simp add: retraction_def)
       
  2029 by (metis assms continuous_on_subset image_mono retraction)
       
  2030 
       
  2031 lemma retract_of_subset:
       
  2032   assumes "t retract_of s" and "t \<subseteq> s'" and "s' \<subseteq> s"
       
  2033     shows "t retract_of s'"
       
  2034 by (meson assms retract_of_def retraction_subset)
       
  2035 
       
  2036 lemma retraction_refl [simp]: "retraction s s (\<lambda>x. x)"
       
  2037 by (simp add: continuous_on_id retraction)
       
  2038 
       
  2039 lemma retract_of_refl [iff]: "s retract_of s"
       
  2040   using continuous_on_id retract_of_def retraction_def by fastforce
       
  2041 
       
  2042 lemma retract_of_imp_subset:
       
  2043    "s retract_of t \<Longrightarrow> s \<subseteq> t"
       
  2044 by (simp add: retract_of_def retraction_def)
       
  2045 
       
  2046 lemma retract_of_empty [simp]:
       
  2047      "({} retract_of s) \<longleftrightarrow> s = {}"  "(s retract_of {}) \<longleftrightarrow> s = {}"
       
  2048 by (auto simp: retract_of_def retraction_def)
       
  2049 
       
  2050 lemma retract_of_singleton [iff]: "({x} retract_of s) \<longleftrightarrow> x \<in> s"
       
  2051   using continuous_on_const
       
  2052   by (auto simp: retract_of_def retraction_def)
       
  2053 
       
  2054 lemma retraction_comp:
       
  2055    "\<lbrakk>retraction s t f; retraction t u g\<rbrakk>
       
  2056         \<Longrightarrow> retraction s u (g o f)"
       
  2057 apply (auto simp: retraction_def intro: continuous_on_compose2)
       
  2058 by blast
       
  2059 
       
  2060 lemma retract_of_trans [trans]:
       
  2061   assumes "s retract_of t" and "t retract_of u"
       
  2062     shows "s retract_of u"
       
  2063 using assms by (auto simp: retract_of_def intro: retraction_comp)
       
  2064 
       
  2065 lemma closedin_retract:
       
  2066   fixes s :: "'a :: real_normed_vector set"
       
  2067   assumes "s retract_of t"
       
  2068     shows "closedin (subtopology euclidean t) s"
       
  2069 proof -
       
  2070   obtain r where "s \<subseteq> t" "continuous_on t r" "r ` t \<subseteq> s" "\<And>x. x \<in> s \<Longrightarrow> r x = x"
       
  2071     using assms by (auto simp: retract_of_def retraction_def)
       
  2072   then have s: "s = {x \<in> t. (norm(r x - x)) = 0}" by auto
       
  2073   show ?thesis
       
  2074     apply (subst s)
       
  2075     apply (rule continuous_closedin_preimage_constant)
       
  2076     by (simp add: \<open>continuous_on t r\<close> continuous_on_diff continuous_on_id continuous_on_norm)
       
  2077 qed
       
  2078 
       
  2079 lemma closedin_self [simp]:
       
  2080     fixes S :: "'a :: real_normed_vector set"
       
  2081     shows "closedin (subtopology euclidean S) S"
       
  2082   by (simp add: closedin_retract)
       
  2083 
       
  2084 lemma retract_of_contractible:
       
  2085   assumes "contractible t" "s retract_of t"
       
  2086     shows "contractible s"
       
  2087 using assms
       
  2088 apply (clarsimp simp add: retract_of_def contractible_def retraction_def homotopic_with)
       
  2089 apply (rule_tac x="r a" in exI)
       
  2090 apply (rule_tac x="r o h" in exI)
       
  2091 apply (intro conjI continuous_intros continuous_on_compose)
       
  2092 apply (erule continuous_on_subset | force)+
       
  2093 done
       
  2094 
       
  2095 lemma retract_of_compact:
       
  2096      "\<lbrakk>compact t; s retract_of t\<rbrakk> \<Longrightarrow> compact s"
       
  2097   by (metis compact_continuous_image retract_of_def retraction)
       
  2098 
       
  2099 lemma retract_of_closed:
       
  2100     fixes s :: "'a :: real_normed_vector set"
       
  2101     shows "\<lbrakk>closed t; s retract_of t\<rbrakk> \<Longrightarrow> closed s"
       
  2102   by (metis closedin_retract closedin_closed_eq)
       
  2103 
       
  2104 lemma retract_of_connected:
       
  2105     "\<lbrakk>connected t; s retract_of t\<rbrakk> \<Longrightarrow> connected s"
       
  2106   by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction)
       
  2107 
       
  2108 lemma retract_of_path_connected:
       
  2109     "\<lbrakk>path_connected t; s retract_of t\<rbrakk> \<Longrightarrow> path_connected s"
       
  2110   by (metis path_connected_continuous_image retract_of_def retraction)
       
  2111 
       
  2112 lemma retract_of_simply_connected:
       
  2113     "\<lbrakk>simply_connected t; s retract_of t\<rbrakk> \<Longrightarrow> simply_connected s"
       
  2114 apply (simp add: retract_of_def retraction_def, clarify)
       
  2115 apply (rule simply_connected_retraction_gen)
       
  2116 apply (force simp: continuous_on_id elim!: continuous_on_subset)+
       
  2117 done
       
  2118 
       
  2119 lemma retract_of_homotopically_trivial:
       
  2120   assumes ts: "t retract_of s"
       
  2121       and hom: "\<And>f g. \<lbrakk>continuous_on u f; f ` u \<subseteq> s;
       
  2122                        continuous_on u g; g ` u \<subseteq> s\<rbrakk>
       
  2123                        \<Longrightarrow> homotopic_with (\<lambda>x. True) u s f g"
       
  2124       and "continuous_on u f" "f ` u \<subseteq> t"
       
  2125       and "continuous_on u g" "g ` u \<subseteq> t"
       
  2126     shows "homotopic_with (\<lambda>x. True) u t f g"
       
  2127 proof -
       
  2128   obtain r where "r ` s \<subseteq> s" "continuous_on s r" "\<forall>x\<in>s. r (r x) = r x" "t = r ` s"
       
  2129     using ts by (auto simp: retract_of_def retraction)
       
  2130   then obtain k where "Retracts s r t k"
       
  2131     unfolding Retracts_def
       
  2132     by (metis continuous_on_subset dual_order.trans image_iff image_mono)
       
  2133   then show ?thesis
       
  2134     apply (rule Retracts.homotopically_trivial_retraction_gen)
       
  2135     using assms
       
  2136     apply (force simp: hom)+
       
  2137     done
       
  2138 qed
       
  2139 
       
  2140 lemma retract_of_homotopically_trivial_null:
       
  2141   assumes ts: "t retract_of s"
       
  2142       and hom: "\<And>f. \<lbrakk>continuous_on u f; f ` u \<subseteq> s\<rbrakk>
       
  2143                      \<Longrightarrow> \<exists>c. homotopic_with (\<lambda>x. True) u s f (\<lambda>x. c)"
       
  2144       and "continuous_on u f" "f ` u \<subseteq> t"
       
  2145   obtains c where "homotopic_with (\<lambda>x. True) u t f (\<lambda>x. c)"
       
  2146 proof -
       
  2147   obtain r where "r ` s \<subseteq> s" "continuous_on s r" "\<forall>x\<in>s. r (r x) = r x" "t = r ` s"
       
  2148     using ts by (auto simp: retract_of_def retraction)
       
  2149   then obtain k where "Retracts s r t k"
       
  2150     unfolding Retracts_def
       
  2151     by (metis continuous_on_subset dual_order.trans image_iff image_mono)
       
  2152   then show ?thesis
       
  2153     apply (rule Retracts.homotopically_trivial_retraction_null_gen)
       
  2154     apply (rule TrueI refl assms that | assumption)+
       
  2155     done
       
  2156 qed
       
  2157 
       
  2158 lemma retraction_imp_quotient_map:
       
  2159    "retraction s t r
       
  2160     \<Longrightarrow> u \<subseteq> t
       
  2161             \<Longrightarrow> (openin (subtopology euclidean s) {x. x \<in> s \<and> r x \<in> u} \<longleftrightarrow>
       
  2162                  openin (subtopology euclidean t) u)"
       
  2163 apply (clarsimp simp add: retraction)
       
  2164 apply (rule continuous_right_inverse_imp_quotient_map [where g=r])
       
  2165 apply (auto simp: elim: continuous_on_subset)
       
  2166 done
       
  2167 
       
  2168 lemma retract_of_locally_compact:
       
  2169     fixes s :: "'a :: {heine_borel,real_normed_vector} set"
       
  2170     shows  "\<lbrakk> locally compact s; t retract_of s\<rbrakk> \<Longrightarrow> locally compact t"
       
  2171   by (metis locally_compact_closedin closedin_retract)
       
  2172 
       
  2173 lemma retract_of_Times:
       
  2174    "\<lbrakk>s retract_of s'; t retract_of t'\<rbrakk> \<Longrightarrow> (s \<times> t) retract_of (s' \<times> t')"
       
  2175 apply (simp add: retract_of_def retraction_def Sigma_mono, clarify)
       
  2176 apply (rename_tac f g)
       
  2177 apply (rule_tac x="\<lambda>z. ((f o fst) z, (g o snd) z)" in exI)
       
  2178 apply (rule conjI continuous_intros | erule continuous_on_subset | force)+
       
  2179 done
       
  2180 
       
  2181 lemma homotopic_into_retract:
       
  2182    "\<lbrakk>f ` s \<subseteq> t; g ` s \<subseteq> t; t retract_of u;
       
  2183         homotopic_with (\<lambda>x. True) s u f g\<rbrakk>
       
  2184         \<Longrightarrow> homotopic_with (\<lambda>x. True) s t f g"
       
  2185 apply (subst (asm) homotopic_with_def)
       
  2186 apply (simp add: homotopic_with retract_of_def retraction_def, clarify)
       
  2187 apply (rule_tac x="r o h" in exI)
       
  2188 apply (rule conjI continuous_intros | erule continuous_on_subset | force simp: image_subset_iff)+
       
  2189 done
       
  2190 
       
  2191 lemma retract_of_locally_connected:
       
  2192   assumes "locally connected T" "S retract_of T"
       
  2193     shows "locally connected S"
       
  2194   using assms
       
  2195   by (auto simp: retract_of_def retraction intro!: retraction_imp_quotient_map elim!: locally_connected_quotient_image)
       
  2196 
       
  2197 lemma retract_of_locally_path_connected:
       
  2198   assumes "locally path_connected T" "S retract_of T"
       
  2199     shows "locally path_connected S"
       
  2200   using assms
       
  2201   by (auto simp: retract_of_def retraction intro!: retraction_imp_quotient_map elim!: locally_path_connected_quotient_image)
       
  2202 
       
  2203 subsection\<open>Punctured affine hulls, etc.\<close>
       
  2204 
       
  2205 lemma continuous_on_compact_surface_projection_aux:
       
  2206   fixes S :: "'a::t2_space set"
       
  2207   assumes "compact S" "S \<subseteq> T" "image q T \<subseteq> S"
       
  2208       and contp: "continuous_on T p"
       
  2209       and "\<And>x. x \<in> S \<Longrightarrow> q x = x"
       
  2210       and [simp]: "\<And>x. x \<in> T \<Longrightarrow> q(p x) = q x"
       
  2211       and "\<And>x. x \<in> T \<Longrightarrow> p(q x) = p x"
       
  2212     shows "continuous_on T q"
       
  2213 proof -
       
  2214   have *: "image p T = image p S"
       
  2215     using assms by auto (metis imageI subset_iff)
       
  2216   have contp': "continuous_on S p"
       
  2217     by (rule continuous_on_subset [OF contp \<open>S \<subseteq> T\<close>])
       
  2218   have "continuous_on T (q \<circ> p)"
       
  2219     apply (rule continuous_on_compose [OF contp])
       
  2220     apply (simp add: *)
       
  2221     apply (rule continuous_on_inv [OF contp' \<open>compact S\<close>])
       
  2222     using assms by auto
       
  2223   then show ?thesis
       
  2224     apply (rule continuous_on_eq [of _ "q o p"])
       
  2225     apply (simp add: o_def)
       
  2226     done
       
  2227 qed
       
  2228 
       
  2229 lemma continuous_on_compact_surface_projection:
       
  2230   fixes S :: "'a::real_normed_vector set"
       
  2231   assumes "compact S"
       
  2232       and S: "S \<subseteq> V - {0}" and "cone V"
       
  2233       and iff: "\<And>x k. x \<in> V - {0} \<Longrightarrow> 0 < k \<and> (k *\<^sub>R x) \<in> S \<longleftrightarrow> d x = k"
       
  2234   shows "continuous_on (V - {0}) (\<lambda>x. d x *\<^sub>R x)"
       
  2235 proof (rule continuous_on_compact_surface_projection_aux [OF \<open>compact S\<close> S])
       
  2236   show "(\<lambda>x. d x *\<^sub>R x) ` (V - {0}) \<subseteq> S"
       
  2237     using iff by auto
       
  2238   show "continuous_on (V - {0}) (\<lambda>x. inverse(norm x) *\<^sub>R x)"
       
  2239     by (intro continuous_intros) force
       
  2240   show "\<And>x. x \<in> S \<Longrightarrow> d x *\<^sub>R x = x"
       
  2241     by (metis S zero_less_one local.iff scaleR_one subset_eq)
       
  2242   show "d (x /\<^sub>R norm x) *\<^sub>R (x /\<^sub>R norm x) = d x *\<^sub>R x" if "x \<in> V - {0}" for x
       
  2243     using iff [of "inverse(norm x) *\<^sub>R x" "norm x * d x", symmetric] iff that \<open>cone V\<close>
       
  2244     by (simp add: field_simps cone_def zero_less_mult_iff)
       
  2245   show "d x *\<^sub>R x /\<^sub>R norm (d x *\<^sub>R x) = x /\<^sub>R norm x" if "x \<in> V - {0}" for x
       
  2246   proof -
       
  2247     have "0 < d x"
       
  2248       using local.iff that by blast
       
  2249     then show ?thesis
       
  2250       by simp
       
  2251   qed
       
  2252 qed
       
  2253 
       
  2254 proposition rel_frontier_deformation_retract_of_punctured_convex:
       
  2255   fixes S :: "'a::euclidean_space set"
       
  2256   assumes "convex S" "convex T" "bounded S"
       
  2257       and arelS: "a \<in> rel_interior S"
       
  2258       and relS: "rel_frontier S \<subseteq> T"
       
  2259       and affS: "T \<subseteq> affine hull S"
       
  2260   obtains r where "homotopic_with (\<lambda>x. True) (T - {a}) (T - {a}) id r"
       
  2261                   "retraction (T - {a}) (rel_frontier S) r"
       
  2262 proof -
       
  2263   have "\<exists>d. 0 < d \<and> (a + d *\<^sub>R l) \<in> rel_frontier S \<and>
       
  2264             (\<forall>e. 0 \<le> e \<and> e < d \<longrightarrow> (a + e *\<^sub>R l) \<in> rel_interior S)"
       
  2265        if "(a + l) \<in> affine hull S" "l \<noteq> 0" for l
       
  2266     apply (rule ray_to_rel_frontier [OF \<open>bounded S\<close> arelS])
       
  2267     apply (rule that)+
       
  2268     by metis
       
  2269   then obtain dd
       
  2270     where dd1: "\<And>l. \<lbrakk>(a + l) \<in> affine hull S; l \<noteq> 0\<rbrakk> \<Longrightarrow> 0 < dd l \<and> (a + dd l *\<^sub>R l) \<in> rel_frontier S"
       
  2271       and dd2: "\<And>l e. \<lbrakk>(a + l) \<in> affine hull S; e < dd l; 0 \<le> e; l \<noteq> 0\<rbrakk>
       
  2272                       \<Longrightarrow> (a + e *\<^sub>R l) \<in> rel_interior S"
       
  2273     by metis+
       
  2274   have aaffS: "a \<in> affine hull S"
       
  2275     by (meson arelS subsetD hull_inc rel_interior_subset)
       
  2276   have "((\<lambda>z. z - a) ` (affine hull S - {a})) = ((\<lambda>z. z - a) ` (affine hull S)) - {0}"
       
  2277     by (auto simp: )
       
  2278   moreover have "continuous_on (((\<lambda>z. z - a) ` (affine hull S)) - {0}) (\<lambda>x. dd x *\<^sub>R x)"
       
  2279   proof (rule continuous_on_compact_surface_projection)
       
  2280     show "compact (rel_frontier ((\<lambda>z. z - a) ` S))"
       
  2281       by (simp add: \<open>bounded S\<close> bounded_translation_minus compact_rel_frontier_bounded)
       
  2282     have releq: "rel_frontier ((\<lambda>z. z - a) ` S) = (\<lambda>z. z - a) ` rel_frontier S"
       
  2283       using rel_frontier_translation [of "-a"] add.commute by simp
       
  2284     also have "... \<subseteq> (\<lambda>z. z - a) ` (affine hull S) - {0}"
       
  2285       using rel_frontier_affine_hull arelS rel_frontier_def by fastforce
       
  2286     finally show "rel_frontier ((\<lambda>z. z - a) ` S) \<subseteq> (\<lambda>z. z - a) ` (affine hull S) - {0}" .
       
  2287     show "cone ((\<lambda>z. z - a) ` (affine hull S))"
       
  2288       apply (rule subspace_imp_cone)
       
  2289       using aaffS
       
  2290       apply (simp add: subspace_affine image_comp o_def affine_translation_aux [of a])
       
  2291       done
       
  2292     show "(0 < k \<and> k *\<^sub>R x \<in> rel_frontier ((\<lambda>z. z - a) ` S)) \<longleftrightarrow> (dd x = k)"
       
  2293          if x: "x \<in> (\<lambda>z. z - a) ` (affine hull S) - {0}" for k x
       
  2294     proof
       
  2295       show "dd x = k \<Longrightarrow> 0 < k \<and> k *\<^sub>R x \<in> rel_frontier ((\<lambda>z. z - a) ` S)"
       
  2296       using dd1 [of x] that image_iff by (fastforce simp add: releq)
       
  2297     next
       
  2298       assume k: "0 < k \<and> k *\<^sub>R x \<in> rel_frontier ((\<lambda>z. z - a) ` S)"
       
  2299       have False if "dd x < k"
       
  2300       proof -
       
  2301         have "k \<noteq> 0" "a + k *\<^sub>R x \<in> closure S"
       
  2302           using k closure_translation [of "-a"]
       
  2303           by (auto simp: rel_frontier_def)
       
  2304         then have segsub: "open_segment a (a + k *\<^sub>R x) \<subseteq> rel_interior S"
       
  2305           by (metis rel_interior_closure_convex_segment [OF \<open>convex S\<close> arelS])
       
  2306         have "x \<noteq> 0" and xaffS: "a + x \<in> affine hull S"
       
  2307           using x by (auto simp: )
       
  2308         then have "0 < dd x" and inS: "a + dd x *\<^sub>R x \<in> rel_frontier S"
       
  2309           using dd1 by auto
       
  2310         moreover have "a + dd x *\<^sub>R x \<in> open_segment a (a + k *\<^sub>R x)"
       
  2311           using k \<open>x \<noteq> 0\<close> \<open>0 < dd x\<close>
       
  2312           apply (simp add: in_segment)
       
  2313           apply (rule_tac x = "dd x / k" in exI)
       
  2314           apply (simp add: field_simps that)
       
  2315           apply (simp add: vector_add_divide_simps algebra_simps)
       
  2316           apply (metis (no_types) \<open>k \<noteq> 0\<close> divide_inverse_commute inverse_eq_divide mult.left_commute right_inverse)
       
  2317           done
       
  2318         ultimately show ?thesis
       
  2319           using segsub by (auto simp add: rel_frontier_def)
       
  2320       qed
       
  2321       moreover have False if "k < dd x"
       
  2322         using x k that rel_frontier_def
       
  2323         by (fastforce simp: algebra_simps releq dest!: dd2)
       
  2324       ultimately show "dd x = k"
       
  2325         by fastforce
       
  2326     qed
       
  2327   qed
       
  2328   ultimately have *: "continuous_on ((\<lambda>z. z - a) ` (affine hull S - {a})) (\<lambda>x. dd x *\<^sub>R x)"
       
  2329     by auto
       
  2330   have "continuous_on (affine hull S - {a}) ((\<lambda>x. a + dd x *\<^sub>R x) \<circ> (\<lambda>z. z - a))"
       
  2331     by (intro * continuous_intros continuous_on_compose)
       
  2332   with affS have contdd: "continuous_on (T - {a}) ((\<lambda>x. a + dd x *\<^sub>R x) \<circ> (\<lambda>z. z - a))"
       
  2333     by (blast intro: continuous_on_subset elim: )
       
  2334   show ?thesis
       
  2335   proof
       
  2336     show "homotopic_with (\<lambda>x. True) (T - {a}) (T - {a}) id (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))"
       
  2337     proof (rule homotopic_with_linear)
       
  2338       show "continuous_on (T - {a}) id"
       
  2339         by (intro continuous_intros continuous_on_compose)
       
  2340       show "continuous_on (T - {a}) (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))"
       
  2341         using contdd by (simp add: o_def)
       
  2342       show "closed_segment (id x) (a + dd (x - a) *\<^sub>R (x - a)) \<subseteq> T - {a}"
       
  2343            if "x \<in> T - {a}" for x
       
  2344       proof (clarsimp simp: in_segment, intro conjI)
       
  2345         fix u::real assume u: "0 \<le> u" "u \<le> 1"
       
  2346         show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \<in> T"
       
  2347           apply (rule convexD [OF \<open>convex T\<close>])
       
  2348           using that u apply (auto simp add: )
       
  2349           apply (metis add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 relS subsetD)
       
  2350           done
       
  2351         have iff: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + d *\<^sub>R (x - a)) = a \<longleftrightarrow>
       
  2352                   (1 - u + u * d) *\<^sub>R (x - a) = 0" for d
       
  2353           by (auto simp: algebra_simps)
       
  2354         have "x \<in> T" "x \<noteq> a" using that by auto
       
  2355         then have axa: "a + (x - a) \<in> affine hull S"
       
  2356            by (metis (no_types) add.commute affS diff_add_cancel set_rev_mp)
       
  2357         then have "\<not> dd (x - a) \<le> 0 \<and> a + dd (x - a) *\<^sub>R (x - a) \<in> rel_frontier S"
       
  2358           using \<open>x \<noteq> a\<close> dd1 by fastforce
       
  2359         with \<open>x \<noteq> a\<close> show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \<noteq> a"
       
  2360           apply (auto simp: iff)
       
  2361           using less_eq_real_def mult_le_0_iff not_less u by fastforce
       
  2362       qed
       
  2363     qed
       
  2364     show "retraction (T - {a}) (rel_frontier S) (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))"
       
  2365     proof (simp add: retraction_def, intro conjI ballI)
       
  2366       show "rel_frontier S \<subseteq> T - {a}"
       
  2367         using arelS relS rel_frontier_def by fastforce
       
  2368       show "continuous_on (T - {a}) (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))"
       
  2369         using contdd by (simp add: o_def)
       
  2370       show "(\<lambda>x. a + dd (x - a) *\<^sub>R (x - a)) ` (T - {a}) \<subseteq> rel_frontier S"
       
  2371         apply (auto simp: rel_frontier_def)
       
  2372         apply (metis Diff_subset add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 rel_frontier_def subset_iff)
       
  2373         by (metis DiffE add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 rel_frontier_def rev_subsetD)
       
  2374       show "a + dd (x - a) *\<^sub>R (x - a) = x" if x: "x \<in> rel_frontier S" for x
       
  2375       proof -
       
  2376         have "x \<noteq> a"
       
  2377           using that arelS by (auto simp add: rel_frontier_def)
       
  2378         have False if "dd (x - a) < 1"
       
  2379         proof -
       
  2380           have "x \<in> closure S"
       
  2381             using x by (auto simp: rel_frontier_def)
       
  2382           then have segsub: "open_segment a x \<subseteq> rel_interior S"
       
  2383             by (metis rel_interior_closure_convex_segment [OF \<open>convex S\<close> arelS])
       
  2384           have  xaffS: "x \<in> affine hull S"
       
  2385             using affS relS x by auto
       
  2386           then have "0 < dd (x - a)" and inS: "a + dd (x - a) *\<^sub>R (x - a) \<in> rel_frontier S"
       
  2387             using dd1 by (auto simp add: \<open>x \<noteq> a\<close>)
       
  2388           moreover have "a + dd (x - a) *\<^sub>R (x - a) \<in> open_segment a x"
       
  2389             using  \<open>x \<noteq> a\<close> \<open>0 < dd (x - a)\<close>
       
  2390             apply (simp add: in_segment)
       
  2391             apply (rule_tac x = "dd (x - a)" in exI)
       
  2392             apply (simp add: algebra_simps that)
       
  2393             done
       
  2394           ultimately show ?thesis
       
  2395             using segsub by (auto simp add: rel_frontier_def)
       
  2396         qed
       
  2397         moreover have False if "1 < dd (x - a)"
       
  2398           using x that dd2 [of "x - a" 1] \<open>x \<noteq> a\<close> closure_affine_hull
       
  2399           by (auto simp: rel_frontier_def)
       
  2400         ultimately have "dd (x - a) = 1" --\<open>similar to another proof above\<close>
       
  2401           by fastforce
       
  2402         with that show ?thesis
       
  2403           by (simp add: rel_frontier_def)
       
  2404       qed
       
  2405     qed
       
  2406   qed
       
  2407 qed
       
  2408 
       
  2409 corollary rel_frontier_retract_of_punctured_affine_hull:
       
  2410   fixes S :: "'a::euclidean_space set"
       
  2411   assumes "bounded S" "convex S" "a \<in> rel_interior S"
       
  2412     shows "rel_frontier S retract_of (affine hull S - {a})"
       
  2413 apply (rule rel_frontier_deformation_retract_of_punctured_convex [of S "affine hull S" a])
       
  2414 apply (auto simp add: affine_imp_convex rel_frontier_affine_hull retract_of_def assms)
       
  2415 done
       
  2416 
       
  2417 corollary rel_boundary_retract_of_punctured_affine_hull:
       
  2418   fixes S :: "'a::euclidean_space set"
       
  2419   assumes "compact S" "convex S" "a \<in> rel_interior S"
       
  2420     shows "(S - rel_interior S) retract_of (affine hull S - {a})"
       
  2421 by (metis assms closure_closed compact_eq_bounded_closed rel_frontier_def
       
  2422           rel_frontier_retract_of_punctured_affine_hull)
       
  2423 
       
  2424 subsection\<open>Borsuk-style characterization of separation\<close>
       
  2425 
       
  2426 lemma continuous_on_Borsuk_map:
       
  2427    "a \<notin> s \<Longrightarrow>  continuous_on s (\<lambda>x. inverse(norm (x - a)) *\<^sub>R (x - a))"
       
  2428 by (rule continuous_intros | force)+
       
  2429 
       
  2430 lemma Borsuk_map_into_sphere:
       
  2431    "(\<lambda>x. inverse(norm (x - a)) *\<^sub>R (x - a)) ` s \<subseteq> sphere 0 1 \<longleftrightarrow> (a \<notin> s)"
       
  2432   by auto (metis eq_iff_diff_eq_0 left_inverse norm_eq_zero)
       
  2433 
       
  2434 lemma Borsuk_maps_homotopic_in_path_component:
       
  2435   assumes "path_component (- s) a b"
       
  2436     shows "homotopic_with (\<lambda>x. True) s (sphere 0 1)
       
  2437                    (\<lambda>x. inverse(norm(x - a)) *\<^sub>R (x - a))
       
  2438                    (\<lambda>x. inverse(norm(x - b)) *\<^sub>R (x - b))"
       
  2439 proof -
       
  2440   obtain g where "path g" "path_image g \<subseteq> -s" "pathstart g = a" "pathfinish g = b"
       
  2441     using assms by (auto simp: path_component_def)
       
  2442   then show ?thesis
       
  2443     apply (simp add: path_def path_image_def pathstart_def pathfinish_def homotopic_with_def)
       
  2444     apply (rule_tac x = "\<lambda>z. inverse(norm(snd z - (g o fst)z)) *\<^sub>R (snd z - (g o fst)z)" in exI)
       
  2445     apply (intro conjI continuous_intros)
       
  2446     apply (rule continuous_intros | erule continuous_on_subset | fastforce simp: divide_simps sphere_def)+
       
  2447     done
       
  2448 qed
       
  2449 
       
  2450 lemma non_extensible_Borsuk_map:
       
  2451   fixes a :: "'a :: euclidean_space"
       
  2452   assumes "compact s" and cin: "c \<in> components(- s)" and boc: "bounded c" and "a \<in> c"
       
  2453     shows "~ (\<exists>g. continuous_on (s \<union> c) g \<and>
       
  2454                   g ` (s \<union> c) \<subseteq> sphere 0 1 \<and>
       
  2455                   (\<forall>x \<in> s. g x = inverse(norm(x - a)) *\<^sub>R (x - a)))"
       
  2456 proof -
       
  2457   have "closed s" using assms by (simp add: compact_imp_closed)
       
  2458   have "c \<subseteq> -s"
       
  2459     using assms by (simp add: in_components_subset)
       
  2460   with \<open>a \<in> c\<close> have "a \<notin> s" by blast
       
  2461   then have ceq: "c = connected_component_set (- s) a"
       
  2462     by (metis \<open>a \<in> c\<close> cin components_iff connected_component_eq)
       
  2463   then have "bounded (s \<union> connected_component_set (- s) a)"
       
  2464     using \<open>compact s\<close> boc compact_imp_bounded by auto
       
  2465   with bounded_subset_ballD obtain r where "0 < r" and r: "(s \<union> connected_component_set (- s) a) \<subseteq> ball a r"
       
  2466     by blast
       
  2467   { fix g
       
  2468     assume "continuous_on (s \<union> c) g"
       
  2469             "g ` (s \<union> c) \<subseteq> sphere 0 1"
       
  2470        and [simp]: "\<And>x. x \<in> s \<Longrightarrow> g x = (x - a) /\<^sub>R norm (x - a)"
       
  2471     then have [simp]: "\<And>x. x \<in> s \<union> c \<Longrightarrow> norm (g x) = 1"
       
  2472       by force
       
  2473     have cb_eq: "cball a r = (s \<union> connected_component_set (- s) a) \<union>
       
  2474                       (cball a r - connected_component_set (- s) a)"
       
  2475       using ball_subset_cball [of a r] r by auto
       
  2476     have cont1: "continuous_on (s \<union> connected_component_set (- s) a)
       
  2477                      (\<lambda>x. a + r *\<^sub>R g x)"
       
  2478       apply (rule continuous_intros)+
       
  2479       using \<open>continuous_on (s \<union> c) g\<close> ceq by blast
       
  2480     have cont2: "continuous_on (cball a r - connected_component_set (- s) a)
       
  2481             (\<lambda>x. a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))"
       
  2482       by (rule continuous_intros | force simp: \<open>a \<notin> s\<close>)+
       
  2483     have 1: "continuous_on (cball a r)
       
  2484              (\<lambda>x. if connected_component (- s) a x
       
  2485                   then a + r *\<^sub>R g x
       
  2486                   else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))"
       
  2487       apply (subst cb_eq)
       
  2488       apply (rule continuous_on_cases [OF _ _ cont1 cont2])
       
  2489         using ceq cin
       
  2490       apply (auto intro: closed_Un_complement_component
       
  2491                   simp: \<open>closed s\<close> open_Compl open_connected_component)
       
  2492       done
       
  2493     have 2: "(\<lambda>x. a + r *\<^sub>R g x) ` (cball a r \<inter> connected_component_set (- s) a)
       
  2494              \<subseteq> sphere a r "
       
  2495       using \<open>0 < r\<close> by (force simp: dist_norm ceq)
       
  2496     have "retraction (cball a r) (sphere a r)
       
  2497             (\<lambda>x. if x \<in> connected_component_set (- s) a
       
  2498                  then a + r *\<^sub>R g x
       
  2499                  else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))"
       
  2500       using  \<open>0 < r\<close>
       
  2501       apply (simp add: retraction_def dist_norm 1 2, safe)
       
  2502       apply (force simp: dist_norm abs_if mult_less_0_iff divide_simps \<open>a \<notin> s\<close>)
       
  2503       using r
       
  2504       by (auto simp: dist_norm norm_minus_commute)
       
  2505     then have False
       
  2506       using no_retraction_cball
       
  2507              [OF \<open>0 < r\<close>, of a, unfolded retract_of_def, simplified, rule_format,
       
  2508               of "\<lambda>x. if x \<in> connected_component_set (- s) a
       
  2509                       then a + r *\<^sub>R g x
       
  2510                       else a + r *\<^sub>R inverse(norm(x - a)) *\<^sub>R (x - a)"]
       
  2511       by blast
       
  2512   }
       
  2513   then show ?thesis
       
  2514     by blast
       
  2515 qed
       
  2516 
       
  2517 subsection\<open>Absolute retracts, Etc.\<close>
       
  2518 
       
  2519 text\<open>Absolute retracts (AR), absolute neighbourhood retracts (ANR) and also
       
  2520  Euclidean neighbourhood retracts (ENR). We define AR and ANR by
       
  2521  specializing the standard definitions for a set to embedding in
       
  2522 spaces of higher dimension. \<close>
       
  2523 
       
  2524 (*This turns out to be sufficient (since any set in
       
  2525 R^n can be embedded as a closed subset of a convex subset of R^{n+1}) to
       
  2526 derive the usual definitions, but we need to split them into two
       
  2527 implications because of the lack of type quantifiers. Then ENR turns out
       
  2528 to be equivalent to ANR plus local compactness. -- JRH*)
       
  2529 
       
  2530 definition AR :: "'a::topological_space set => bool"
       
  2531   where
       
  2532    "AR S \<equiv> \<forall>U. \<forall>S'::('a * real) set. S homeomorphic S' \<and> closedin (subtopology euclidean U) S'
       
  2533                 \<longrightarrow> S' retract_of U"
       
  2534 
       
  2535 definition ANR :: "'a::topological_space set => bool"
       
  2536   where
       
  2537    "ANR S \<equiv> \<forall>U. \<forall>S'::('a * real) set. S homeomorphic S' \<and> closedin (subtopology euclidean U) S'
       
  2538                 \<longrightarrow> (\<exists>T. openin (subtopology euclidean U) T \<and>
       
  2539                         S' retract_of T)"
       
  2540 
       
  2541 definition ENR :: "'a::topological_space set => bool"
       
  2542   where "ENR S \<equiv> \<exists>U. open U \<and> S retract_of U"
       
  2543 
       
  2544 text\<open> First, show that we do indeed get the "usual" properties of ARs and ANRs.\<close>
       
  2545 
       
  2546 proposition AR_imp_absolute_extensor:
       
  2547   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
  2548   assumes "AR S" and contf: "continuous_on T f" and "f ` T \<subseteq> S"
       
  2549       and cloUT: "closedin (subtopology euclidean U) T"
       
  2550   obtains g where "continuous_on U g" "g ` U \<subseteq> S" "\<And>x. x \<in> T \<Longrightarrow> g x = f x"
       
  2551 proof -
       
  2552   have "aff_dim S < int (DIM('b \<times> real))"
       
  2553     using aff_dim_le_DIM [of S] by simp
       
  2554   then obtain C and S' :: "('b * real) set"
       
  2555           where C: "convex C" "C \<noteq> {}"
       
  2556             and cloCS: "closedin (subtopology euclidean C) S'"
       
  2557             and hom: "S homeomorphic S'"
       
  2558     by (metis that homeomorphic_closedin_convex)
       
  2559   then have "S' retract_of C"
       
  2560     using \<open>AR S\<close> by (simp add: AR_def)
       
  2561   then obtain r where "S' \<subseteq> C" and contr: "continuous_on C r"
       
  2562                   and "r ` C \<subseteq> S'" and rid: "\<And>x. x\<in>S' \<Longrightarrow> r x = x"
       
  2563     by (auto simp: retraction_def retract_of_def)
       
  2564   obtain g h where "homeomorphism S S' g h"
       
  2565     using hom by (force simp: homeomorphic_def)
       
  2566   then have "continuous_on (f ` T) g"
       
  2567     by (meson \<open>f ` T \<subseteq> S\<close> continuous_on_subset homeomorphism_def)
       
  2568   then have contgf: "continuous_on T (g o f)"
       
  2569     by (metis continuous_on_compose contf)
       
  2570   have gfTC: "(g \<circ> f) ` T \<subseteq> C"
       
  2571   proof -
       
  2572     have "g ` S = S'"
       
  2573       by (metis (no_types) \<open>homeomorphism S S' g h\<close> homeomorphism_def)
       
  2574     with \<open>S' \<subseteq> C\<close> \<open>f ` T \<subseteq> S\<close> show ?thesis by force
       
  2575   qed
       
  2576   obtain f' where f': "continuous_on U f'"  "f' ` U \<subseteq> C"
       
  2577                       "\<And>x. x \<in> T \<Longrightarrow> f' x = (g \<circ> f) x"
       
  2578     by (metis Dugundji [OF C cloUT contgf gfTC])
       
  2579   show ?thesis
       
  2580   proof (rule_tac g = "h o r o f'" in that)
       
  2581     show "continuous_on U (h \<circ> r \<circ> f')"
       
  2582       apply (intro continuous_on_compose f')
       
  2583        using continuous_on_subset contr f' apply blast
       
  2584       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)
       
  2585     show "(h \<circ> r \<circ> f') ` U \<subseteq> S"
       
  2586       using \<open>homeomorphism S S' g h\<close> \<open>r ` C \<subseteq> S'\<close> \<open>f' ` U \<subseteq> C\<close>
       
  2587       by (fastforce simp: homeomorphism_def)
       
  2588     show "\<And>x. x \<in> T \<Longrightarrow> (h \<circ> r \<circ> f') x = f x"
       
  2589       using \<open>homeomorphism S S' g h\<close> \<open>f ` T \<subseteq> S\<close> f'
       
  2590       by (auto simp: rid homeomorphism_def)
       
  2591   qed
       
  2592 qed
       
  2593 
       
  2594 lemma AR_imp_absolute_retract:
       
  2595   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
  2596   assumes "AR S" "S homeomorphic S'"
       
  2597       and clo: "closedin (subtopology euclidean U) S'"
       
  2598     shows "S' retract_of U"
       
  2599 proof -
       
  2600   obtain g h where hom: "homeomorphism S S' g h"
       
  2601     using assms by (force simp: homeomorphic_def)
       
  2602   have h: "continuous_on S' h" " h ` S' \<subseteq> S"
       
  2603     using hom homeomorphism_def apply blast
       
  2604     apply (metis hom equalityE homeomorphism_def)
       
  2605     done
       
  2606   obtain h' where h': "continuous_on U h'" "h' ` U \<subseteq> S"
       
  2607               and h'h: "\<And>x. x \<in> S' \<Longrightarrow> h' x = h x"
       
  2608     by (blast intro: AR_imp_absolute_extensor [OF \<open>AR S\<close> h clo])
       
  2609   have [simp]: "S' \<subseteq> U" using clo closedin_limpt by blast
       
  2610   show ?thesis
       
  2611   proof (simp add: retraction_def retract_of_def, intro exI conjI)
       
  2612     show "continuous_on U (g o h')"
       
  2613       apply (intro continuous_on_compose h')
       
  2614       apply (meson hom continuous_on_subset h' homeomorphism_cont1)
       
  2615       done
       
  2616     show "(g \<circ> h') ` U \<subseteq> S'"
       
  2617       using h'  by clarsimp (metis hom subsetD homeomorphism_def imageI)
       
  2618     show "\<forall>x\<in>S'. (g \<circ> h') x = x"
       
  2619       by clarsimp (metis h'h hom homeomorphism_def)
       
  2620   qed
       
  2621 qed
       
  2622 
       
  2623 lemma AR_imp_absolute_retract_UNIV:
       
  2624   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
  2625   assumes "AR S" and hom: "S homeomorphic S'"
       
  2626       and clo: "closed S'"
       
  2627     shows "S' retract_of UNIV"
       
  2628 apply (rule AR_imp_absolute_retract [OF \<open>AR S\<close> hom])
       
  2629 using clo closed_closedin by auto
       
  2630 
       
  2631 lemma absolute_extensor_imp_AR:
       
  2632   fixes S :: "'a::euclidean_space set"
       
  2633   assumes "\<And>f :: 'a * real \<Rightarrow> 'a.
       
  2634            \<And>U T. \<lbrakk>continuous_on T f;  f ` T \<subseteq> S;
       
  2635                   closedin (subtopology euclidean U) T\<rbrakk>
       
  2636                  \<Longrightarrow> \<exists>g. continuous_on U g \<and> g ` U \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)"
       
  2637   shows "AR S"
       
  2638 proof (clarsimp simp: AR_def)
       
  2639   fix U and T :: "('a * real) set"
       
  2640   assume "S homeomorphic T" and clo: "closedin (subtopology euclidean U) T"
       
  2641   then obtain g h where hom: "homeomorphism S T g h"
       
  2642     by (force simp: homeomorphic_def)
       
  2643   have h: "continuous_on T h" " h ` T \<subseteq> S"
       
  2644     using hom homeomorphism_def apply blast
       
  2645     apply (metis hom equalityE homeomorphism_def)
       
  2646     done
       
  2647   obtain h' where h': "continuous_on U h'" "h' ` U \<subseteq> S"
       
  2648               and h'h: "\<forall>x\<in>T. h' x = h x"
       
  2649     using assms [OF h clo] by blast
       
  2650   have [simp]: "T \<subseteq> U"
       
  2651     using clo closedin_imp_subset by auto
       
  2652   show "T retract_of U"
       
  2653   proof (simp add: retraction_def retract_of_def, intro exI conjI)
       
  2654     show "continuous_on U (g o h')"
       
  2655       apply (intro continuous_on_compose h')
       
  2656       apply (meson hom continuous_on_subset h' homeomorphism_cont1)
       
  2657       done
       
  2658     show "(g \<circ> h') ` U \<subseteq> T"
       
  2659       using h'  by clarsimp (metis hom subsetD homeomorphism_def imageI)
       
  2660     show "\<forall>x\<in>T. (g \<circ> h') x = x"
       
  2661       by clarsimp (metis h'h hom homeomorphism_def)
       
  2662   qed
       
  2663 qed
       
  2664 
       
  2665 lemma AR_eq_absolute_extensor:
       
  2666   fixes S :: "'a::euclidean_space set"
       
  2667   shows "AR S \<longleftrightarrow>
       
  2668        (\<forall>f :: 'a * real \<Rightarrow> 'a.
       
  2669         \<forall>U T. continuous_on T f \<longrightarrow> f ` T \<subseteq> S \<longrightarrow>
       
  2670                closedin (subtopology euclidean U) T \<longrightarrow>
       
  2671                 (\<exists>g. continuous_on U g \<and> g ` U \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)))"
       
  2672 apply (rule iffI)
       
  2673  apply (metis AR_imp_absolute_extensor)
       
  2674 apply (simp add: absolute_extensor_imp_AR)
       
  2675 done
       
  2676 
       
  2677 lemma AR_imp_retract:
       
  2678   fixes S :: "'a::euclidean_space set"
       
  2679   assumes "AR S \<and> closedin (subtopology euclidean U) S"
       
  2680     shows "S retract_of U"
       
  2681 using AR_imp_absolute_retract assms homeomorphic_refl by blast
       
  2682 
       
  2683 lemma AR_homeomorphic_AR:
       
  2684   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  2685   assumes "AR T" "S homeomorphic T"
       
  2686     shows "AR S"
       
  2687 unfolding AR_def
       
  2688 by (metis assms AR_imp_absolute_retract homeomorphic_trans [of _ S] homeomorphic_sym)
       
  2689 
       
  2690 lemma homeomorphic_AR_iff_AR:
       
  2691   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  2692   shows "S homeomorphic T \<Longrightarrow> AR S \<longleftrightarrow> AR T"
       
  2693 by (metis AR_homeomorphic_AR homeomorphic_sym)
       
  2694 
       
  2695 
       
  2696 proposition ANR_imp_absolute_neighbourhood_extensor:
       
  2697   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
  2698   assumes "ANR S" and contf: "continuous_on T f" and "f ` T \<subseteq> S"
       
  2699       and cloUT: "closedin (subtopology euclidean U) T"
       
  2700   obtains V g where "T \<subseteq> V" "openin (subtopology euclidean U) V"
       
  2701                     "continuous_on V g"
       
  2702                     "g ` V \<subseteq> S" "\<And>x. x \<in> T \<Longrightarrow> g x = f x"
       
  2703 proof -
       
  2704   have "aff_dim S < int (DIM('b \<times> real))"
       
  2705     using aff_dim_le_DIM [of S] by simp
       
  2706   then obtain C and S' :: "('b * real) set"
       
  2707           where C: "convex C" "C \<noteq> {}"
       
  2708             and cloCS: "closedin (subtopology euclidean C) S'"
       
  2709             and hom: "S homeomorphic S'"
       
  2710     by (metis that homeomorphic_closedin_convex)
       
  2711   then obtain D where opD: "openin (subtopology euclidean C) D" and "S' retract_of D"
       
  2712     using \<open>ANR S\<close> by (auto simp: ANR_def)
       
  2713   then obtain r where "S' \<subseteq> D" and contr: "continuous_on D r"
       
  2714                   and "r ` D \<subseteq> S'" and rid: "\<And>x. x \<in> S' \<Longrightarrow> r x = x"
       
  2715     by (auto simp: retraction_def retract_of_def)
       
  2716   obtain g h where homgh: "homeomorphism S S' g h"
       
  2717     using hom by (force simp: homeomorphic_def)
       
  2718   have "continuous_on (f ` T) g"
       
  2719     by (meson \<open>f ` T \<subseteq> S\<close> continuous_on_subset homeomorphism_def homgh)
       
  2720   then have contgf: "continuous_on T (g o f)"
       
  2721     by (intro continuous_on_compose contf)
       
  2722   have gfTC: "(g \<circ> f) ` T \<subseteq> C"
       
  2723   proof -
       
  2724     have "g ` S = S'"
       
  2725       by (metis (no_types) homeomorphism_def homgh)
       
  2726     then show ?thesis
       
  2727       by (metis (no_types) assms(3) cloCS closedin_def image_comp image_mono order.trans topspace_euclidean_subtopology)
       
  2728   qed
       
  2729   obtain f' where contf': "continuous_on U f'"
       
  2730               and "f' ` U \<subseteq> C"
       
  2731               and eq: "\<And>x. x \<in> T \<Longrightarrow> f' x = (g \<circ> f) x"
       
  2732     by (metis Dugundji [OF C cloUT contgf gfTC])
       
  2733   show ?thesis
       
  2734   proof (rule_tac V = "{x \<in> U. f' x \<in> D}" and g = "h o r o f'" in that)
       
  2735     show "T \<subseteq> {x \<in> U. f' x \<in> D}"
       
  2736       using cloUT closedin_imp_subset \<open>S' \<subseteq> D\<close> \<open>f ` T \<subseteq> S\<close> eq homeomorphism_image1 homgh
       
  2737       by fastforce
       
  2738     show ope: "openin (subtopology euclidean U) {x \<in> U. f' x \<in> D}"
       
  2739       using  \<open>f' ` U \<subseteq> C\<close> by (auto simp: opD contf' continuous_openin_preimage)
       
  2740     have conth: "continuous_on (r ` f' ` {x \<in> U. f' x \<in> D}) h"
       
  2741       apply (rule continuous_on_subset [of S'])
       
  2742       using homeomorphism_def homgh apply blast
       
  2743       using \<open>r ` D \<subseteq> S'\<close> by blast
       
  2744     show "continuous_on {x \<in> U. f' x \<in> D} (h \<circ> r \<circ> f')"
       
  2745       apply (intro continuous_on_compose conth
       
  2746                    continuous_on_subset [OF contr] continuous_on_subset [OF contf'], auto)
       
  2747       done
       
  2748     show "(h \<circ> r \<circ> f') ` {x \<in> U. f' x \<in> D} \<subseteq> S"
       
  2749       using \<open>homeomorphism S S' g h\<close>  \<open>f' ` U \<subseteq> C\<close>  \<open>r ` D \<subseteq> S'\<close>
       
  2750       by (auto simp: homeomorphism_def)
       
  2751     show "\<And>x. x \<in> T \<Longrightarrow> (h \<circ> r \<circ> f') x = f x"
       
  2752       using \<open>homeomorphism S S' g h\<close> \<open>f ` T \<subseteq> S\<close> eq
       
  2753       by (auto simp: rid homeomorphism_def)
       
  2754   qed
       
  2755 qed
       
  2756 
       
  2757 
       
  2758 corollary ANR_imp_absolute_neighbourhood_retract:
       
  2759   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
  2760   assumes "ANR S" "S homeomorphic S'"
       
  2761       and clo: "closedin (subtopology euclidean U) S'"
       
  2762   obtains V where "openin (subtopology euclidean U) V" "S' retract_of V"
       
  2763 proof -
       
  2764   obtain g h where hom: "homeomorphism S S' g h"
       
  2765     using assms by (force simp: homeomorphic_def)
       
  2766   have h: "continuous_on S' h" " h ` S' \<subseteq> S"
       
  2767     using hom homeomorphism_def apply blast
       
  2768     apply (metis hom equalityE homeomorphism_def)
       
  2769     done
       
  2770     from ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> h clo]
       
  2771   obtain V h' where "S' \<subseteq> V" and opUV: "openin (subtopology euclidean U) V"
       
  2772                 and h': "continuous_on V h'" "h' ` V \<subseteq> S"
       
  2773                 and h'h:"\<And>x. x \<in> S' \<Longrightarrow> h' x = h x"
       
  2774     by (blast intro: ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> h clo])
       
  2775   have "S' retract_of V"
       
  2776   proof (simp add: retraction_def retract_of_def, intro exI conjI \<open>S' \<subseteq> V\<close>)
       
  2777     show "continuous_on V (g o h')"
       
  2778       apply (intro continuous_on_compose h')
       
  2779       apply (meson hom continuous_on_subset h' homeomorphism_cont1)
       
  2780       done
       
  2781     show "(g \<circ> h') ` V \<subseteq> S'"
       
  2782       using h'  by clarsimp (metis hom subsetD homeomorphism_def imageI)
       
  2783     show "\<forall>x\<in>S'. (g \<circ> h') x = x"
       
  2784       by clarsimp (metis h'h hom homeomorphism_def)
       
  2785   qed
       
  2786   then show ?thesis
       
  2787     by (rule that [OF opUV])
       
  2788 qed
       
  2789 
       
  2790 corollary ANR_imp_absolute_neighbourhood_retract_UNIV:
       
  2791   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
  2792   assumes "ANR S" and hom: "S homeomorphic S'" and clo: "closed S'"
       
  2793   obtains V where "open V" "S' retract_of V"
       
  2794   using ANR_imp_absolute_neighbourhood_retract [OF \<open>ANR S\<close> hom]
       
  2795 by (metis clo closed_closedin open_openin subtopology_UNIV)
       
  2796 
       
  2797 lemma absolute_neighbourhood_extensor_imp_ANR:
       
  2798   fixes S :: "'a::euclidean_space set"
       
  2799   assumes "\<And>f :: 'a * real \<Rightarrow> 'a.
       
  2800            \<And>U T. \<lbrakk>continuous_on T f;  f ` T \<subseteq> S;
       
  2801                   closedin (subtopology euclidean U) T\<rbrakk>
       
  2802                  \<Longrightarrow> \<exists>V g. T \<subseteq> V \<and> openin (subtopology euclidean U) V \<and>
       
  2803                        continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)"
       
  2804   shows "ANR S"
       
  2805 proof (clarsimp simp: ANR_def)
       
  2806   fix U and T :: "('a * real) set"
       
  2807   assume "S homeomorphic T" and clo: "closedin (subtopology euclidean U) T"
       
  2808   then obtain g h where hom: "homeomorphism S T g h"
       
  2809     by (force simp: homeomorphic_def)
       
  2810   have h: "continuous_on T h" " h ` T \<subseteq> S"
       
  2811     using hom homeomorphism_def apply blast
       
  2812     apply (metis hom equalityE homeomorphism_def)
       
  2813     done
       
  2814   obtain V h' where "T \<subseteq> V" and opV: "openin (subtopology euclidean U) V"
       
  2815                 and h': "continuous_on V h'" "h' ` V \<subseteq> S"
       
  2816               and h'h: "\<forall>x\<in>T. h' x = h x"
       
  2817     using assms [OF h clo] by blast
       
  2818   have [simp]: "T \<subseteq> U"
       
  2819     using clo closedin_imp_subset by auto
       
  2820   have "T retract_of V"
       
  2821   proof (simp add: retraction_def retract_of_def, intro exI conjI \<open>T \<subseteq> V\<close>)
       
  2822     show "continuous_on V (g o h')"
       
  2823       apply (intro continuous_on_compose h')
       
  2824       apply (meson hom continuous_on_subset h' homeomorphism_cont1)
       
  2825       done
       
  2826     show "(g \<circ> h') ` V \<subseteq> T"
       
  2827       using h'  by clarsimp (metis hom subsetD homeomorphism_def imageI)
       
  2828     show "\<forall>x\<in>T. (g \<circ> h') x = x"
       
  2829       by clarsimp (metis h'h hom homeomorphism_def)
       
  2830   qed
       
  2831   then show "\<exists>V. openin (subtopology euclidean U) V \<and> T retract_of V"
       
  2832     using opV by blast
       
  2833 qed
       
  2834 
       
  2835 lemma ANR_eq_absolute_neighbourhood_extensor:
       
  2836   fixes S :: "'a::euclidean_space set"
       
  2837   shows "ANR S \<longleftrightarrow>
       
  2838          (\<forall>f :: 'a * real \<Rightarrow> 'a.
       
  2839           \<forall>U T. continuous_on T f \<longrightarrow> f ` T \<subseteq> S \<longrightarrow>
       
  2840                 closedin (subtopology euclidean U) T \<longrightarrow>
       
  2841                (\<exists>V g. T \<subseteq> V \<and> openin (subtopology euclidean U) V \<and>
       
  2842                        continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)))"
       
  2843 apply (rule iffI)
       
  2844  apply (metis ANR_imp_absolute_neighbourhood_extensor)
       
  2845 apply (simp add: absolute_neighbourhood_extensor_imp_ANR)
       
  2846 done
       
  2847 
       
  2848 lemma ANR_imp_neighbourhood_retract:
       
  2849   fixes S :: "'a::euclidean_space set"
       
  2850   assumes "ANR S" "closedin (subtopology euclidean U) S"
       
  2851   obtains V where "openin (subtopology euclidean U) V" "S retract_of V"
       
  2852 using ANR_imp_absolute_neighbourhood_retract assms homeomorphic_refl by blast
       
  2853 
       
  2854 lemma ANR_imp_absolute_closed_neighbourhood_retract:
       
  2855   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
  2856   assumes "ANR S" "S homeomorphic S'" and US': "closedin (subtopology euclidean U) S'"
       
  2857   obtains V W
       
  2858     where "openin (subtopology euclidean U) V"
       
  2859           "closedin (subtopology euclidean U) W"
       
  2860           "S' \<subseteq> V" "V \<subseteq> W" "S' retract_of W"
       
  2861 proof -
       
  2862   obtain Z where "openin (subtopology euclidean U) Z" and S'Z: "S' retract_of Z"
       
  2863     by (blast intro: assms ANR_imp_absolute_neighbourhood_retract)
       
  2864   then have UUZ: "closedin (subtopology euclidean U) (U - Z)"
       
  2865     by auto
       
  2866   have "S' \<inter> (U - Z) = {}"
       
  2867     using \<open>S' retract_of Z\<close> closedin_retract closedin_subtopology by fastforce
       
  2868   then obtain V W
       
  2869       where "openin (subtopology euclidean U) V"
       
  2870         and "openin (subtopology euclidean U) W"
       
  2871         and "S' \<subseteq> V" "U - Z \<subseteq> W" "V \<inter> W = {}"
       
  2872       using separation_normal_local [OF US' UUZ]  by auto
       
  2873   moreover have "S' retract_of U - W"
       
  2874     apply (rule retract_of_subset [OF S'Z])
       
  2875     using US' \<open>S' \<subseteq> V\<close> \<open>V \<inter> W = {}\<close> closedin_subset apply fastforce
       
  2876     using Diff_subset_conv \<open>U - Z \<subseteq> W\<close> by blast
       
  2877   ultimately show ?thesis
       
  2878     apply (rule_tac V=V and W = "U-W" in that)
       
  2879     using openin_imp_subset apply (force simp:)+
       
  2880     done
       
  2881 qed
       
  2882 
       
  2883 lemma ANR_imp_closed_neighbourhood_retract:
       
  2884   fixes S :: "'a::euclidean_space set"
       
  2885   assumes "ANR S" "closedin (subtopology euclidean U) S"
       
  2886   obtains V W where "openin (subtopology euclidean U) V"
       
  2887                     "closedin (subtopology euclidean U) W"
       
  2888                     "S \<subseteq> V" "V \<subseteq> W" "S retract_of W"
       
  2889 by (meson ANR_imp_absolute_closed_neighbourhood_retract assms homeomorphic_refl)
       
  2890 
       
  2891 lemma ANR_homeomorphic_ANR:
       
  2892   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  2893   assumes "ANR T" "S homeomorphic T"
       
  2894     shows "ANR S"
       
  2895 unfolding ANR_def
       
  2896 by (metis assms ANR_imp_absolute_neighbourhood_retract homeomorphic_trans [of _ S] homeomorphic_sym)
       
  2897 
       
  2898 lemma homeomorphic_ANR_iff_ANR:
       
  2899   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  2900   shows "S homeomorphic T \<Longrightarrow> ANR S \<longleftrightarrow> ANR T"
       
  2901 by (metis ANR_homeomorphic_ANR homeomorphic_sym)
       
  2902 
       
  2903 subsection\<open> Analogous properties of ENRs.\<close>
       
  2904 
       
  2905 proposition ENR_imp_absolute_neighbourhood_retract:
       
  2906   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
  2907   assumes "ENR S" and hom: "S homeomorphic S'"
       
  2908       and "S' \<subseteq> U"
       
  2909   obtains V where "openin (subtopology euclidean U) V" "S' retract_of V"
       
  2910 proof -
       
  2911   obtain X where "open X" "S retract_of X"
       
  2912     using \<open>ENR S\<close> by (auto simp: ENR_def)
       
  2913   then obtain r where "retraction X S r"
       
  2914     by (auto simp: retract_of_def)
       
  2915   have "locally compact S'"
       
  2916     using retract_of_locally_compact open_imp_locally_compact
       
  2917           homeomorphic_local_compactness \<open>S retract_of X\<close> \<open>open X\<close> hom by blast
       
  2918   then obtain W where UW: "openin (subtopology euclidean U) W"
       
  2919                   and WS': "closedin (subtopology euclidean W) S'"
       
  2920     apply (rule locally_compact_closedin_open)
       
  2921     apply (rename_tac W)
       
  2922     apply (rule_tac W = "U \<inter> W" in that, blast)
       
  2923     by (simp add: \<open>S' \<subseteq> U\<close> closedin_limpt)
       
  2924   obtain f g where hom: "homeomorphism S S' f g"
       
  2925     using assms by (force simp: homeomorphic_def)
       
  2926   have contg: "continuous_on S' g"
       
  2927     using hom homeomorphism_def by blast
       
  2928   moreover have "g ` S' \<subseteq> S" by (metis hom equalityE homeomorphism_def)
       
  2929   ultimately obtain h where conth: "continuous_on W h" and hg: "\<And>x. x \<in> S' \<Longrightarrow> h x = g x"
       
  2930     using Tietze_unbounded [of S' g W] WS' by blast
       
  2931   have "W \<subseteq> U" using UW openin_open by auto
       
  2932   have "S' \<subseteq> W" using WS' closedin_closed by auto
       
  2933   have him: "\<And>x. x \<in> S' \<Longrightarrow> h x \<in> X"
       
  2934     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)
       
  2935   have "S' retract_of {x \<in> W. h x \<in> X}"
       
  2936   proof (simp add: retraction_def retract_of_def, intro exI conjI)
       
  2937     show "S' \<subseteq> {x \<in> W. h x \<in> X}"
       
  2938       using him WS' closedin_imp_subset by blast
       
  2939     show "continuous_on {x \<in> W. h x \<in> X} (f o r o h)"
       
  2940     proof (intro continuous_on_compose)
       
  2941       show "continuous_on {x \<in> W. h x \<in> X} h"
       
  2942         by (metis (no_types) Collect_restrict conth continuous_on_subset)
       
  2943       show "continuous_on (h ` {x \<in> W. h x \<in> X}) r"
       
  2944       proof -
       
  2945         have "h ` {b \<in> W. h b \<in> X} \<subseteq> X"
       
  2946           by blast
       
  2947         then show "continuous_on (h ` {b \<in> W. h b \<in> X}) r"
       
  2948           by (meson \<open>retraction X S r\<close> continuous_on_subset retraction)
       
  2949       qed
       
  2950       show "continuous_on (r ` h ` {x \<in> W. h x \<in> X}) f"
       
  2951         apply (rule continuous_on_subset [of S])
       
  2952          using hom homeomorphism_def apply blast
       
  2953         apply clarify
       
  2954         apply (meson \<open>retraction X S r\<close> subsetD imageI retraction_def)
       
  2955         done
       
  2956     qed
       
  2957     show "(f \<circ> r \<circ> h) ` {x \<in> W. h x \<in> X} \<subseteq> S'"
       
  2958       using \<open>retraction X S r\<close> hom
       
  2959       by (auto simp: retraction_def homeomorphism_def)
       
  2960     show "\<forall>x\<in>S'. (f \<circ> r \<circ> h) x = x"
       
  2961       using \<open>retraction X S r\<close> hom by (auto simp: retraction_def homeomorphism_def hg)
       
  2962   qed
       
  2963   then show ?thesis
       
  2964     apply (rule_tac V = "{x. x \<in> W \<and> h x \<in> X}" in that)
       
  2965      apply (rule openin_trans [OF _ UW])
       
  2966      using \<open>continuous_on W h\<close> \<open>open X\<close> continuous_openin_preimage_eq apply blast+
       
  2967      done
       
  2968 qed
       
  2969 
       
  2970 corollary ENR_imp_absolute_neighbourhood_retract_UNIV:
       
  2971   fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
       
  2972   assumes "ENR S" "S homeomorphic S'"
       
  2973   obtains T' where "open T'" "S' retract_of T'"
       
  2974 by (metis ENR_imp_absolute_neighbourhood_retract UNIV_I assms(1) assms(2) open_openin subsetI subtopology_UNIV)
       
  2975 
       
  2976 lemma ENR_homeomorphic_ENR:
       
  2977   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  2978   assumes "ENR T" "S homeomorphic T"
       
  2979     shows "ENR S"
       
  2980 unfolding ENR_def
       
  2981 by (meson ENR_imp_absolute_neighbourhood_retract_UNIV assms homeomorphic_sym)
       
  2982 
       
  2983 lemma homeomorphic_ENR_iff_ENR:
       
  2984   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  2985   assumes "S homeomorphic T"
       
  2986     shows "ENR S \<longleftrightarrow> ENR T"
       
  2987 by (meson ENR_homeomorphic_ENR assms homeomorphic_sym)
       
  2988 
       
  2989 lemma ENR_translation:
       
  2990   fixes S :: "'a::euclidean_space set"
       
  2991   shows "ENR(image (\<lambda>x. a + x) S) \<longleftrightarrow> ENR S"
       
  2992 by (meson homeomorphic_sym homeomorphic_translation homeomorphic_ENR_iff_ENR)
       
  2993 
       
  2994 lemma ENR_linear_image_eq:
       
  2995   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
  2996   assumes "linear f" "inj f"
       
  2997   shows "ENR (image f S) \<longleftrightarrow> ENR S"
       
  2998 apply (rule homeomorphic_ENR_iff_ENR)
       
  2999 using assms homeomorphic_sym linear_homeomorphic_image by auto
       
  3000 
       
  3001 subsection\<open>Some relations among the concepts\<close>
       
  3002 
       
  3003 text\<open>We also relate AR to being a retract of UNIV, which is often a more convenient proxy in the closed case.\<close>
       
  3004 
       
  3005 lemma AR_imp_ANR: "AR S \<Longrightarrow> ANR S"
       
  3006   using ANR_def AR_def by fastforce
       
  3007 
       
  3008 lemma ENR_imp_ANR:
       
  3009   fixes S :: "'a::euclidean_space set"
       
  3010   shows "ENR S \<Longrightarrow> ANR S"
       
  3011 apply (simp add: ANR_def)
       
  3012 by (metis ENR_imp_absolute_neighbourhood_retract closedin_imp_subset)
       
  3013 
       
  3014 lemma ENR_ANR:
       
  3015   fixes S :: "'a::euclidean_space set"
       
  3016   shows "ENR S \<longleftrightarrow> ANR S \<and> locally compact S"
       
  3017 proof
       
  3018   assume "ENR S"
       
  3019   then have "locally compact S"
       
  3020     using ENR_def open_imp_locally_compact retract_of_locally_compact by auto
       
  3021   then show "ANR S \<and> locally compact S"
       
  3022     using ENR_imp_ANR \<open>ENR S\<close> by blast
       
  3023 next
       
  3024   assume "ANR S \<and> locally compact S"
       
  3025   then have "ANR S" "locally compact S" by auto
       
  3026   then obtain T :: "('a * real) set" where "closed T" "S homeomorphic T"
       
  3027     using locally_compact_homeomorphic_closed
       
  3028     by (metis DIM_prod DIM_real Suc_eq_plus1 lessI)
       
  3029   then show "ENR S"
       
  3030     using \<open>ANR S\<close>
       
  3031     apply (simp add: ANR_def)
       
  3032     apply (drule_tac x=UNIV in spec)
       
  3033     apply (drule_tac x=T in spec)
       
  3034     apply (auto simp: closed_closedin)
       
  3035     apply (meson ENR_def ENR_homeomorphic_ENR open_openin)
       
  3036     done
       
  3037 qed
       
  3038 
       
  3039 
       
  3040 proposition AR_ANR:
       
  3041   fixes S :: "'a::euclidean_space set"
       
  3042   shows "AR S \<longleftrightarrow> ANR S \<and> contractible S \<and> S \<noteq> {}"
       
  3043         (is "?lhs = ?rhs")
       
  3044 proof
       
  3045   assume ?lhs
       
  3046   obtain C and S' :: "('a * real) set"
       
  3047     where "convex C" "C \<noteq> {}" "closedin (subtopology euclidean C) S'" "S homeomorphic S'"
       
  3048       apply (rule homeomorphic_closedin_convex [of S, where 'n = "'a * real"])
       
  3049       using aff_dim_le_DIM [of S] by auto
       
  3050   with \<open>AR S\<close> have "contractible S"
       
  3051     apply (simp add: AR_def)
       
  3052     apply (drule_tac x=C in spec)
       
  3053     apply (drule_tac x="S'" in spec, simp)
       
  3054     using convex_imp_contractible homeomorphic_contractible_eq retract_of_contractible by fastforce
       
  3055   with \<open>AR S\<close> show ?rhs
       
  3056     apply (auto simp: AR_imp_ANR)
       
  3057     apply (force simp: AR_def)
       
  3058     done
       
  3059 next
       
  3060   assume ?rhs
       
  3061   then obtain a and h:: "real \<times> 'a \<Rightarrow> 'a"
       
  3062       where conth: "continuous_on ({0..1} \<times> S) h"
       
  3063         and hS: "h ` ({0..1} \<times> S) \<subseteq> S"
       
  3064         and [simp]: "\<And>x. h(0, x) = x"
       
  3065         and [simp]: "\<And>x. h(1, x) = a"
       
  3066         and "ANR S" "S \<noteq> {}"
       
  3067     by (auto simp: contractible_def homotopic_with_def)
       
  3068   then have "a \<in> S"
       
  3069     by (metis all_not_in_conv atLeastAtMost_iff image_subset_iff mem_Sigma_iff order_refl zero_le_one)
       
  3070   have "\<exists>g. continuous_on W g \<and> g ` W \<subseteq> S \<and> (\<forall>x\<in>T. g x = f x)"
       
  3071          if      f: "continuous_on T f" "f ` T \<subseteq> S"
       
  3072             and WT: "closedin (subtopology euclidean W) T"
       
  3073          for W T and f :: "'a \<times> real \<Rightarrow> 'a"
       
  3074   proof -
       
  3075     obtain U g
       
  3076       where "T \<subseteq> U" and WU: "openin (subtopology euclidean W) U"
       
  3077         and contg: "continuous_on U g"
       
  3078         and "g ` U \<subseteq> S" and gf: "\<And>x. x \<in> T \<Longrightarrow> g x = f x"
       
  3079       using iffD1 [OF ANR_eq_absolute_neighbourhood_extensor \<open>ANR S\<close>, rule_format, OF f WT]
       
  3080       by auto
       
  3081     have WWU: "closedin (subtopology euclidean W) (W - U)"
       
  3082       using WU closedin_diff by fastforce
       
  3083     moreover have "(W - U) \<inter> T = {}"
       
  3084       using \<open>T \<subseteq> U\<close> by auto
       
  3085     ultimately obtain V V'
       
  3086       where WV': "openin (subtopology euclidean W) V'"
       
  3087         and WV: "openin (subtopology euclidean W) V"
       
  3088         and "W - U \<subseteq> V'" "T \<subseteq> V" "V' \<inter> V = {}"
       
  3089       using separation_normal_local [of W "W-U" T] WT by blast
       
  3090     then have WVT: "T \<inter> (W - V) = {}"
       
  3091       by auto
       
  3092     have WWV: "closedin (subtopology euclidean W) (W - V)"
       
  3093       using WV closedin_diff by fastforce
       
  3094     obtain j :: " 'a \<times> real \<Rightarrow> real"
       
  3095       where contj: "continuous_on W j"
       
  3096         and j:  "\<And>x. x \<in> W \<Longrightarrow> j x \<in> {0..1}"
       
  3097         and j0: "\<And>x. x \<in> W - V \<Longrightarrow> j x = 1"
       
  3098         and j1: "\<And>x. x \<in> T \<Longrightarrow> j x = 0"
       
  3099       by (rule Urysohn_local [OF WT WWV WVT, of 0 "1::real"]) (auto simp: in_segment)
       
  3100     have Weq: "W = (W - V) \<union> (W - V')"
       
  3101       using \<open>V' \<inter> V = {}\<close> by force
       
  3102     show ?thesis
       
  3103     proof (intro conjI exI)
       
  3104       have *: "continuous_on (W - V') (\<lambda>x. h (j x, g x))"
       
  3105         apply (rule continuous_on_compose2 [OF conth continuous_on_Pair])
       
  3106           apply (rule continuous_on_subset [OF contj Diff_subset])
       
  3107          apply (rule continuous_on_subset [OF contg])
       
  3108          apply (metis Diff_subset_conv Un_commute \<open>W - U \<subseteq> V'\<close>)
       
  3109         using j \<open>g ` U \<subseteq> S\<close> \<open>W - U \<subseteq> V'\<close> apply fastforce
       
  3110         done
       
  3111       show "continuous_on W (\<lambda>x. if x \<in> W - V then a else h (j x, g x))"
       
  3112         apply (subst Weq)
       
  3113         apply (rule continuous_on_cases_local)
       
  3114             apply (simp_all add: Weq [symmetric] WWV continuous_on_const *)
       
  3115           using WV' closedin_diff apply fastforce
       
  3116          apply (auto simp: j0 j1)
       
  3117         done
       
  3118     next
       
  3119       have "h (j (x, y), g (x, y)) \<in> S" if "(x, y) \<in> W" "(x, y) \<in> V" for x y
       
  3120       proof -
       
  3121         have "j(x, y) \<in> {0..1}"
       
  3122           using j that by blast
       
  3123         moreover have "g(x, y) \<in> S"
       
  3124           using \<open>V' \<inter> V = {}\<close> \<open>W - U \<subseteq> V'\<close> \<open>g ` U \<subseteq> S\<close> that by fastforce
       
  3125         ultimately show ?thesis
       
  3126           using hS by blast
       
  3127       qed
       
  3128       with \<open>a \<in> S\<close> \<open>g ` U \<subseteq> S\<close>
       
  3129       show "(\<lambda>x. if x \<in> W - V then a else h (j x, g x)) ` W \<subseteq> S"
       
  3130         by auto
       
  3131     next
       
  3132       show "\<forall>x\<in>T. (if x \<in> W - V then a else h (j x, g x)) = f x"
       
  3133         using \<open>T \<subseteq> V\<close> by (auto simp: j0 j1 gf)
       
  3134     qed
       
  3135   qed
       
  3136   then show ?lhs
       
  3137     by (simp add: AR_eq_absolute_extensor)
       
  3138 qed
       
  3139 
       
  3140 
       
  3141 lemma ANR_retract_of_ANR:
       
  3142   fixes S :: "'a::euclidean_space set"
       
  3143   assumes "ANR T" "S retract_of T"
       
  3144   shows "ANR S"
       
  3145 using assms
       
  3146 apply (simp add: ANR_eq_absolute_neighbourhood_extensor retract_of_def retraction_def)
       
  3147 apply (clarsimp elim!: all_forward)
       
  3148 apply (erule impCE, metis subset_trans)
       
  3149 apply (clarsimp elim!: ex_forward)
       
  3150 apply (rule_tac x="r o g" in exI)
       
  3151 by (metis comp_apply continuous_on_compose continuous_on_subset subsetD imageI image_comp image_mono subset_trans)
       
  3152 
       
  3153 lemma AR_retract_of_AR:
       
  3154   fixes S :: "'a::euclidean_space set"
       
  3155   shows "\<lbrakk>AR T; S retract_of T\<rbrakk> \<Longrightarrow> AR S"
       
  3156 using ANR_retract_of_ANR AR_ANR retract_of_contractible by fastforce
       
  3157 
       
  3158 lemma ENR_retract_of_ENR:
       
  3159    "\<lbrakk>ENR T; S retract_of T\<rbrakk> \<Longrightarrow> ENR S"
       
  3160 by (meson ENR_def retract_of_trans)
       
  3161 
       
  3162 lemma retract_of_UNIV:
       
  3163   fixes S :: "'a::euclidean_space set"
       
  3164   shows "S retract_of UNIV \<longleftrightarrow> AR S \<and> closed S"
       
  3165 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)
       
  3166 
       
  3167 lemma compact_AR [simp]:
       
  3168   fixes S :: "'a::euclidean_space set"
       
  3169   shows "compact S \<and> AR S \<longleftrightarrow> compact S \<and> S retract_of UNIV"
       
  3170 using compact_imp_closed retract_of_UNIV by blast
       
  3171 
       
  3172 subsection\<open>More properties of ARs, ANRs and ENRs\<close>
       
  3173 
       
  3174 lemma not_AR_empty [simp]: "~ AR({})"
       
  3175   by (auto simp: AR_def)
       
  3176 
       
  3177 lemma ENR_empty [simp]: "ENR {}"
       
  3178   by (simp add: ENR_def)
       
  3179 
       
  3180 lemma ANR_empty [simp]: "ANR ({} :: 'a::euclidean_space set)"
       
  3181   by (simp add: ENR_imp_ANR)
       
  3182 
       
  3183 lemma convex_imp_AR:
       
  3184   fixes S :: "'a::euclidean_space set"
       
  3185   shows "\<lbrakk>convex S; S \<noteq> {}\<rbrakk> \<Longrightarrow> AR S"
       
  3186 apply (rule absolute_extensor_imp_AR)
       
  3187 apply (rule Dugundji, assumption+)
       
  3188 by blast
       
  3189 
       
  3190 lemma convex_imp_ANR:
       
  3191   fixes S :: "'a::euclidean_space set"
       
  3192   shows "convex S \<Longrightarrow> ANR S"
       
  3193 using ANR_empty AR_imp_ANR convex_imp_AR by blast
       
  3194 
       
  3195 lemma ENR_convex_closed:
       
  3196   fixes S :: "'a::euclidean_space set"
       
  3197   shows "\<lbrakk>closed S; convex S\<rbrakk> \<Longrightarrow> ENR S"
       
  3198 using ENR_def ENR_empty convex_imp_AR retract_of_UNIV by blast
       
  3199 
       
  3200 lemma AR_UNIV [simp]: "AR (UNIV :: 'a::euclidean_space set)"
       
  3201   using retract_of_UNIV by auto
       
  3202 
       
  3203 lemma ANR_UNIV [simp]: "ANR (UNIV :: 'a::euclidean_space set)"
       
  3204   by (simp add: AR_imp_ANR)
       
  3205 
       
  3206 lemma ENR_UNIV [simp]:"ENR UNIV"
       
  3207   using ENR_def by blast
       
  3208 
       
  3209 lemma AR_singleton:
       
  3210     fixes a :: "'a::euclidean_space"
       
  3211     shows "AR {a}"
       
  3212   using retract_of_UNIV by blast
       
  3213 
       
  3214 lemma ANR_singleton:
       
  3215     fixes a :: "'a::euclidean_space"
       
  3216     shows "ANR {a}"
       
  3217   by (simp add: AR_imp_ANR AR_singleton)
       
  3218 
       
  3219 lemma ENR_singleton: "ENR {a}"
       
  3220   using ENR_def by blast
       
  3221 
       
  3222 subsection\<open>ARs closed under union\<close>
       
  3223 
       
  3224 lemma AR_closed_Un_local_aux:
       
  3225   fixes U :: "'a::euclidean_space set"
       
  3226   assumes "closedin (subtopology euclidean U) S"
       
  3227           "closedin (subtopology euclidean U) T"
       
  3228           "AR S" "AR T" "AR(S \<inter> T)"
       
  3229   shows "(S \<union> T) retract_of U"
       
  3230 proof -
       
  3231   have "S \<inter> T \<noteq> {}"
       
  3232     using assms AR_def by fastforce
       
  3233   have "S \<subseteq> U" "T \<subseteq> U"
       
  3234     using assms by (auto simp: closedin_imp_subset)
       
  3235   define S' where "S' \<equiv> {x \<in> U. setdist {x} S \<le> setdist {x} T}"
       
  3236   define T' where "T' \<equiv> {x \<in> U. setdist {x} T \<le> setdist {x} S}"
       
  3237   define W  where "W \<equiv> {x \<in> U. setdist {x} S = setdist {x} T}"
       
  3238   have US': "closedin (subtopology euclidean U) S'"
       
  3239     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} S - setdist {x} T" "{..0}"]
       
  3240     by (simp add: S'_def continuous_intros)
       
  3241   have UT': "closedin (subtopology euclidean U) T'"
       
  3242     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} T - setdist {x} S" "{..0}"]
       
  3243     by (simp add: T'_def continuous_intros)
       
  3244   have "S \<subseteq> S'"
       
  3245     using S'_def \<open>S \<subseteq> U\<close> setdist_sing_in_set by fastforce
       
  3246   have "T \<subseteq> T'"
       
  3247     using T'_def \<open>T \<subseteq> U\<close> setdist_sing_in_set by fastforce
       
  3248   have "S \<inter> T \<subseteq> W" "W \<subseteq> U"
       
  3249     using \<open>S \<subseteq> U\<close> by (auto simp: W_def setdist_sing_in_set)
       
  3250   have "(S \<inter> T) retract_of W"
       
  3251     apply (rule AR_imp_absolute_retract [OF \<open>AR(S \<inter> T)\<close>])
       
  3252      apply (simp add: homeomorphic_refl)
       
  3253     apply (rule closedin_subset_trans [of U])
       
  3254     apply (simp_all add: assms closedin_Int \<open>S \<inter> T \<subseteq> W\<close> \<open>W \<subseteq> U\<close>)
       
  3255     done
       
  3256   then obtain r0
       
  3257     where "S \<inter> T \<subseteq> W" and contr0: "continuous_on W r0"
       
  3258       and "r0 ` W \<subseteq> S \<inter> T"
       
  3259       and r0 [simp]: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> r0 x = x"
       
  3260       by (auto simp: retract_of_def retraction_def)
       
  3261   have ST: "x \<in> W \<Longrightarrow> x \<in> S \<longleftrightarrow> x \<in> T" for x
       
  3262     using setdist_eq_0_closedin \<open>S \<inter> T \<noteq> {}\<close> assms
       
  3263     by (force simp: W_def setdist_sing_in_set)
       
  3264   have "S' \<inter> T' = W"
       
  3265     by (auto simp: S'_def T'_def W_def)
       
  3266   then have cloUW: "closedin (subtopology euclidean U) W"
       
  3267     using closedin_Int US' UT' by blast
       
  3268   define r where "r \<equiv> \<lambda>x. if x \<in> W then r0 x else x"
       
  3269   have "r ` (W \<union> S) \<subseteq> S" "r ` (W \<union> T) \<subseteq> T"
       
  3270     using \<open>r0 ` W \<subseteq> S \<inter> T\<close> r_def by auto
       
  3271   have contr: "continuous_on (W \<union> (S \<union> T)) r"
       
  3272   unfolding r_def
       
  3273   proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id])
       
  3274     show "closedin (subtopology euclidean (W \<union> (S \<union> T))) W"
       
  3275       using \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> \<open>W \<subseteq> U\<close> \<open>closedin (subtopology euclidean U) W\<close> closedin_subset_trans by fastforce
       
  3276     show "closedin (subtopology euclidean (W \<union> (S \<union> T))) (S \<union> T)"
       
  3277       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)
       
  3278     show "\<And>x. x \<in> W \<and> x \<notin> W \<or> x \<in> S \<union> T \<and> x \<in> W \<Longrightarrow> r0 x = x"
       
  3279       by (auto simp: ST)
       
  3280   qed
       
  3281   have cloUWS: "closedin (subtopology euclidean U) (W \<union> S)"
       
  3282     by (simp add: cloUW assms closedin_Un)
       
  3283   obtain g where contg: "continuous_on U g"
       
  3284              and "g ` U \<subseteq> S" and geqr: "\<And>x. x \<in> W \<union> S \<Longrightarrow> g x = r x"
       
  3285     apply (rule AR_imp_absolute_extensor [OF \<open>AR S\<close> _ _ cloUWS])
       
  3286       apply (rule continuous_on_subset [OF contr])
       
  3287       using \<open>r ` (W \<union> S) \<subseteq> S\<close> apply auto
       
  3288     done
       
  3289   have cloUWT: "closedin (subtopology euclidean U) (W \<union> T)"
       
  3290     by (simp add: cloUW assms closedin_Un)
       
  3291   obtain h where conth: "continuous_on U h"
       
  3292              and "h ` U \<subseteq> T" and heqr: "\<And>x. x \<in> W \<union> T \<Longrightarrow> h x = r x"
       
  3293     apply (rule AR_imp_absolute_extensor [OF \<open>AR T\<close> _ _ cloUWT])
       
  3294       apply (rule continuous_on_subset [OF contr])
       
  3295       using \<open>r ` (W \<union> T) \<subseteq> T\<close> apply auto
       
  3296     done
       
  3297   have "U = S' \<union> T'"
       
  3298     by (force simp: S'_def T'_def)
       
  3299   then have cont: "continuous_on U (\<lambda>x. if x \<in> S' then g x else h x)"
       
  3300     apply (rule ssubst)
       
  3301     apply (rule continuous_on_cases_local)
       
  3302     using US' UT' \<open>S' \<inter> T' = W\<close> \<open>U = S' \<union> T'\<close>
       
  3303           contg conth continuous_on_subset geqr heqr apply auto
       
  3304     done
       
  3305   have UST: "(\<lambda>x. if x \<in> S' then g x else h x) ` U \<subseteq> S \<union> T"
       
  3306     using \<open>g ` U \<subseteq> S\<close> \<open>h ` U \<subseteq> T\<close> by auto
       
  3307   show ?thesis
       
  3308     apply (simp add: retract_of_def retraction_def \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close>)
       
  3309     apply (rule_tac x="\<lambda>x. if x \<in> S' then g x else h x" in exI)
       
  3310     apply (intro conjI cont UST)
       
  3311     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)
       
  3312 qed
       
  3313 
       
  3314 
       
  3315 proposition AR_closed_Un_local:
       
  3316   fixes S :: "'a::euclidean_space set"
       
  3317   assumes STS: "closedin (subtopology euclidean (S \<union> T)) S"
       
  3318       and STT: "closedin (subtopology euclidean (S \<union> T)) T"
       
  3319       and "AR S" "AR T" "AR(S \<inter> T)"
       
  3320     shows "AR(S \<union> T)"
       
  3321 proof -
       
  3322   have "C retract_of U"
       
  3323        if hom: "S \<union> T homeomorphic C" and UC: "closedin (subtopology euclidean U) C"
       
  3324        for U and C :: "('a * real) set"
       
  3325   proof -
       
  3326     obtain f g where hom: "homeomorphism (S \<union> T) C f g"
       
  3327       using hom by (force simp: homeomorphic_def)
       
  3328     have US: "closedin (subtopology euclidean U) {x \<in> C. g x \<in> S}"
       
  3329       apply (rule closedin_trans [OF _ UC])
       
  3330       apply (rule continuous_closedin_preimage_gen [OF _ _ STS])
       
  3331       using hom homeomorphism_def apply blast
       
  3332       apply (metis hom homeomorphism_def set_eq_subset)
       
  3333       done
       
  3334     have UT: "closedin (subtopology euclidean U) {x \<in> C. g x \<in> T}"
       
  3335       apply (rule closedin_trans [OF _ UC])
       
  3336       apply (rule continuous_closedin_preimage_gen [OF _ _ STT])
       
  3337       using hom homeomorphism_def apply blast
       
  3338       apply (metis hom homeomorphism_def set_eq_subset)
       
  3339       done
       
  3340     have ARS: "AR {x \<in> C. g x \<in> S}"
       
  3341       apply (rule AR_homeomorphic_AR [OF \<open>AR S\<close>])
       
  3342       apply (simp add: homeomorphic_def)
       
  3343       apply (rule_tac x=g in exI)
       
  3344       apply (rule_tac x=f in exI)
       
  3345       using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset)
       
  3346       apply (rule_tac x="f x" in image_eqI, auto)
       
  3347       done
       
  3348     have ART: "AR {x \<in> C. g x \<in> T}"
       
  3349       apply (rule AR_homeomorphic_AR [OF \<open>AR T\<close>])
       
  3350       apply (simp add: homeomorphic_def)
       
  3351       apply (rule_tac x=g in exI)
       
  3352       apply (rule_tac x=f in exI)
       
  3353       using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset)
       
  3354       apply (rule_tac x="f x" in image_eqI, auto)
       
  3355       done
       
  3356     have ARI: "AR ({x \<in> C. g x \<in> S} \<inter> {x \<in> C. g x \<in> T})"
       
  3357       apply (rule AR_homeomorphic_AR [OF \<open>AR (S \<inter> T)\<close>])
       
  3358       apply (simp add: homeomorphic_def)
       
  3359       apply (rule_tac x=g in exI)
       
  3360       apply (rule_tac x=f in exI)
       
  3361       using hom
       
  3362       apply (auto simp: homeomorphism_def elim!: continuous_on_subset)
       
  3363       apply (rule_tac x="f x" in image_eqI, auto)
       
  3364       done
       
  3365     have "C = {x \<in> C. g x \<in> S} \<union> {x \<in> C. g x \<in> T}"
       
  3366       using hom  by (auto simp: homeomorphism_def)
       
  3367     then show ?thesis
       
  3368       by (metis AR_closed_Un_local_aux [OF US UT ARS ART ARI])
       
  3369   qed
       
  3370   then show ?thesis
       
  3371     by (force simp: AR_def)
       
  3372 qed
       
  3373 
       
  3374 corollary AR_closed_Un:
       
  3375   fixes S :: "'a::euclidean_space set"
       
  3376   shows "\<lbrakk>closed S; closed T; AR S; AR T; AR (S \<inter> T)\<rbrakk> \<Longrightarrow> AR (S \<union> T)"
       
  3377 by (metis AR_closed_Un_local_aux closed_closedin retract_of_UNIV subtopology_UNIV)
       
  3378 
       
  3379 subsection\<open>ANRs closed under union\<close>
       
  3380 
       
  3381 lemma ANR_closed_Un_local_aux:
       
  3382   fixes U :: "'a::euclidean_space set"
       
  3383   assumes US: "closedin (subtopology euclidean U) S"
       
  3384       and UT: "closedin (subtopology euclidean U) T"
       
  3385       and "ANR S" "ANR T" "ANR(S \<inter> T)"
       
  3386   obtains V where "openin (subtopology euclidean U) V" "(S \<union> T) retract_of V"
       
  3387 proof (cases "S = {} \<or> T = {}")
       
  3388   case True with assms that show ?thesis
       
  3389     by (auto simp: intro: ANR_imp_neighbourhood_retract)
       
  3390 next
       
  3391   case False
       
  3392   then have [simp]: "S \<noteq> {}" "T \<noteq> {}" by auto
       
  3393   have "S \<subseteq> U" "T \<subseteq> U"
       
  3394     using assms by (auto simp: closedin_imp_subset)
       
  3395   define S' where "S' \<equiv> {x \<in> U. setdist {x} S \<le> setdist {x} T}"
       
  3396   define T' where "T' \<equiv> {x \<in> U. setdist {x} T \<le> setdist {x} S}"
       
  3397   define W  where "W \<equiv> {x \<in> U. setdist {x} S = setdist {x} T}"
       
  3398   have cloUS': "closedin (subtopology euclidean U) S'"
       
  3399     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} S - setdist {x} T" "{..0}"]
       
  3400     by (simp add: S'_def continuous_intros)
       
  3401   have cloUT': "closedin (subtopology euclidean U) T'"
       
  3402     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} T - setdist {x} S" "{..0}"]
       
  3403     by (simp add: T'_def continuous_intros)
       
  3404   have "S \<subseteq> S'"
       
  3405     using S'_def \<open>S \<subseteq> U\<close> setdist_sing_in_set by fastforce
       
  3406   have "T \<subseteq> T'"
       
  3407     using T'_def \<open>T \<subseteq> U\<close> setdist_sing_in_set by fastforce
       
  3408   have "S' \<union> T' = U"
       
  3409     by (auto simp: S'_def T'_def)
       
  3410   have "W \<subseteq> S'"
       
  3411     by (simp add: Collect_mono S'_def W_def)
       
  3412   have "W \<subseteq> T'"
       
  3413     by (simp add: Collect_mono T'_def W_def)
       
  3414   have ST_W: "S \<inter> T \<subseteq> W" and "W \<subseteq> U"
       
  3415     using \<open>S \<subseteq> U\<close> by (force simp: W_def setdist_sing_in_set)+
       
  3416   have "S' \<inter> T' = W"
       
  3417     by (auto simp: S'_def T'_def W_def)
       
  3418   then have cloUW: "closedin (subtopology euclidean U) W"
       
  3419     using closedin_Int cloUS' cloUT' by blast
       
  3420   obtain W' W0 where "openin (subtopology euclidean W) W'"
       
  3421                  and cloWW0: "closedin (subtopology euclidean W) W0"
       
  3422                  and "S \<inter> T \<subseteq> W'" "W' \<subseteq> W0"
       
  3423                  and ret: "(S \<inter> T) retract_of W0"
       
  3424     apply (rule ANR_imp_closed_neighbourhood_retract [OF \<open>ANR(S \<inter> T)\<close>])
       
  3425     apply (rule closedin_subset_trans [of U, OF _ ST_W \<open>W \<subseteq> U\<close>])
       
  3426     apply (blast intro: assms)+
       
  3427     done
       
  3428   then obtain U0 where opeUU0: "openin (subtopology euclidean U) U0"
       
  3429                    and U0: "S \<inter> T \<subseteq> U0" "U0 \<inter> W \<subseteq> W0"
       
  3430     unfolding openin_open  using \<open>W \<subseteq> U\<close> by blast
       
  3431   have "W0 \<subseteq> U"
       
  3432     using \<open>W \<subseteq> U\<close> cloWW0 closedin_subset by fastforce
       
  3433   obtain r0
       
  3434     where "S \<inter> T \<subseteq> W0" and contr0: "continuous_on W0 r0" and "r0 ` W0 \<subseteq> S \<inter> T"
       
  3435       and r0 [simp]: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> r0 x = x"
       
  3436     using ret  by (force simp add: retract_of_def retraction_def)
       
  3437   have ST: "x \<in> W \<Longrightarrow> x \<in> S \<longleftrightarrow> x \<in> T" for x
       
  3438     using assms by (auto simp: W_def setdist_sing_in_set dest!: setdist_eq_0_closedin)
       
  3439   define r where "r \<equiv> \<lambda>x. if x \<in> W0 then r0 x else x"
       
  3440   have "r ` (W0 \<union> S) \<subseteq> S" "r ` (W0 \<union> T) \<subseteq> T"
       
  3441     using \<open>r0 ` W0 \<subseteq> S \<inter> T\<close> r_def by auto
       
  3442   have contr: "continuous_on (W0 \<union> (S \<union> T)) r"
       
  3443   unfolding r_def
       
  3444   proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id])
       
  3445     show "closedin (subtopology euclidean (W0 \<union> (S \<union> T))) W0"
       
  3446       apply (rule closedin_subset_trans [of U])
       
  3447       using cloWW0 cloUW closedin_trans \<open>W0 \<subseteq> U\<close> \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> apply blast+
       
  3448       done
       
  3449     show "closedin (subtopology euclidean (W0 \<union> (S \<union> T))) (S \<union> T)"
       
  3450       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)
       
  3451     show "\<And>x. x \<in> W0 \<and> x \<notin> W0 \<or> x \<in> S \<union> T \<and> x \<in> W0 \<Longrightarrow> r0 x = x"
       
  3452       using ST cloWW0 closedin_subset by fastforce
       
  3453   qed
       
  3454   have cloS'WS: "closedin (subtopology euclidean S') (W0 \<union> S)"
       
  3455     by (meson closedin_subset_trans US cloUS' \<open>S \<subseteq> S'\<close> \<open>W \<subseteq> S'\<close> cloUW cloWW0 
       
  3456               closedin_Un closedin_imp_subset closedin_trans)
       
  3457   obtain W1 g where "W0 \<union> S \<subseteq> W1" and contg: "continuous_on W1 g"
       
  3458                 and opeSW1: "openin (subtopology euclidean S') W1"
       
  3459                 and "g ` W1 \<subseteq> S" and geqr: "\<And>x. x \<in> W0 \<union> S \<Longrightarrow> g x = r x"
       
  3460     apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> _ \<open>r ` (W0 \<union> S) \<subseteq> S\<close> cloS'WS])
       
  3461      apply (rule continuous_on_subset [OF contr])
       
  3462     apply (blast intro:  elim: )+
       
  3463     done
       
  3464   have cloT'WT: "closedin (subtopology euclidean T') (W0 \<union> T)"
       
  3465     by (meson closedin_subset_trans UT cloUT' \<open>T \<subseteq> T'\<close> \<open>W \<subseteq> T'\<close> cloUW cloWW0 
       
  3466               closedin_Un closedin_imp_subset closedin_trans)
       
  3467   obtain W2 h where "W0 \<union> T \<subseteq> W2" and conth: "continuous_on W2 h"
       
  3468                 and opeSW2: "openin (subtopology euclidean T') W2"
       
  3469                 and "h ` W2 \<subseteq> T" and heqr: "\<And>x. x \<in> W0 \<union> T \<Longrightarrow> h x = r x"
       
  3470     apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> _ \<open>r ` (W0 \<union> T) \<subseteq> T\<close> cloT'WT])
       
  3471      apply (rule continuous_on_subset [OF contr])
       
  3472     apply (blast intro:  elim: )+
       
  3473     done
       
  3474   have "S' \<inter> T' = W"
       
  3475     by (force simp: S'_def T'_def W_def)
       
  3476   obtain O1 O2 where "open O1" "W1 = S' \<inter> O1" "open O2" "W2 = T' \<inter> O2"
       
  3477     using opeSW1 opeSW2 by (force simp add: openin_open)
       
  3478   show ?thesis
       
  3479   proof
       
  3480     have eq: "W1 - (W - U0) \<union> (W2 - (W - U0)) =
       
  3481          ((U - T') \<inter> O1 \<union> (U - S') \<inter> O2 \<union> U \<inter> O1 \<inter> O2) - (W - U0)"
       
  3482      using \<open>U0 \<inter> W \<subseteq> W0\<close> \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W0 \<union> T \<subseteq> W2\<close>
       
  3483       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>)
       
  3484     show "openin (subtopology euclidean U) (W1 - (W - U0) \<union> (W2 - (W - U0)))"
       
  3485       apply (subst eq)
       
  3486       apply (intro openin_Un openin_Int_open openin_diff closedin_diff cloUW opeUU0 cloUS' cloUT' \<open>open O1\<close> \<open>open O2\<close>)
       
  3487       apply simp_all
       
  3488       done
       
  3489     have cloW1: "closedin (subtopology euclidean (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W1 - (W - U0))"
       
  3490       using cloUS' apply (simp add: closedin_closed)
       
  3491       apply (erule ex_forward)
       
  3492       using U0 \<open>W0 \<union> S \<subseteq> W1\<close>
       
  3493       apply (auto simp add: \<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])
       
  3494       done
       
  3495     have cloW2: "closedin (subtopology euclidean (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W2 - (W - U0))"
       
  3496       using cloUT' apply (simp add: closedin_closed)
       
  3497       apply (erule ex_forward)
       
  3498       using U0 \<open>W0 \<union> T \<subseteq> W2\<close>
       
  3499       apply (auto simp add: \<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])
       
  3500       done
       
  3501     have *: "\<forall>x\<in>S \<union> T. (if x \<in> S' then g x else h x) = x"
       
  3502       using ST \<open>S' \<inter> T' = W\<close> cloT'WT closedin_subset geqr heqr 
       
  3503       apply (auto simp: r_def)
       
  3504        apply fastforce
       
  3505       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
       
  3506     have "\<exists>r. continuous_on (W1 - (W - U0) \<union> (W2 - (W - U0))) r \<and>
       
  3507               r ` (W1 - (W - U0) \<union> (W2 - (W - U0))) \<subseteq> S \<union> T \<and> 
       
  3508               (\<forall>x\<in>S \<union> T. r x = x)"
       
  3509       apply (rule_tac x = "\<lambda>x. if  x \<in> S' then g x else h x" in exI)
       
  3510       apply (intro conjI *)
       
  3511       apply (rule continuous_on_cases_local 
       
  3512                   [OF cloW1 cloW2 continuous_on_subset [OF contg] continuous_on_subset [OF conth]])
       
  3513       using \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<inter> T' = W\<close>
       
  3514             \<open>g ` W1 \<subseteq> S\<close> \<open>h ` W2 \<subseteq> T\<close> apply auto
       
  3515       using \<open>U0 \<inter> W \<subseteq> W0\<close> \<open>W0 \<union> S \<subseteq> W1\<close> apply (fastforce simp add: geqr heqr)+
       
  3516       done
       
  3517     then show "S \<union> T retract_of W1 - (W - U0) \<union> (W2 - (W - U0))"
       
  3518       using  \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W0 \<union> T \<subseteq> W2\<close> ST opeUU0 U0
       
  3519       by (auto simp add: retract_of_def retraction_def)
       
  3520   qed
       
  3521 qed
       
  3522 
       
  3523 
       
  3524 proposition ANR_closed_Un_local:
       
  3525   fixes S :: "'a::euclidean_space set"
       
  3526   assumes STS: "closedin (subtopology euclidean (S \<union> T)) S"
       
  3527       and STT: "closedin (subtopology euclidean (S \<union> T)) T"
       
  3528       and "ANR S" "ANR T" "ANR(S \<inter> T)" 
       
  3529     shows "ANR(S \<union> T)"
       
  3530 proof -
       
  3531   have "\<exists>T. openin (subtopology euclidean U) T \<and> C retract_of T"
       
  3532        if hom: "S \<union> T homeomorphic C" and UC: "closedin (subtopology euclidean U) C"
       
  3533        for U and C :: "('a * real) set"
       
  3534   proof -
       
  3535     obtain f g where hom: "homeomorphism (S \<union> T) C f g"
       
  3536       using hom by (force simp: homeomorphic_def)
       
  3537     have US: "closedin (subtopology euclidean U) {x \<in> C. g x \<in> S}"
       
  3538       apply (rule closedin_trans [OF _ UC])
       
  3539       apply (rule continuous_closedin_preimage_gen [OF _ _ STS])
       
  3540       using hom [unfolded homeomorphism_def] apply blast
       
  3541       apply (metis hom homeomorphism_def set_eq_subset)
       
  3542       done
       
  3543     have UT: "closedin (subtopology euclidean U) {x \<in> C. g x \<in> T}"
       
  3544       apply (rule closedin_trans [OF _ UC])
       
  3545       apply (rule continuous_closedin_preimage_gen [OF _ _ STT])
       
  3546       using hom [unfolded homeomorphism_def] apply blast
       
  3547       apply (metis hom homeomorphism_def set_eq_subset)
       
  3548       done
       
  3549     have ANRS: "ANR {x \<in> C. g x \<in> S}"
       
  3550       apply (rule ANR_homeomorphic_ANR [OF \<open>ANR S\<close>])
       
  3551       apply (simp add: homeomorphic_def)
       
  3552       apply (rule_tac x=g in exI)
       
  3553       apply (rule_tac x=f in exI)
       
  3554       using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset)
       
  3555       apply (rule_tac x="f x" in image_eqI, auto)
       
  3556       done
       
  3557     have ANRT: "ANR {x \<in> C. g x \<in> T}"
       
  3558       apply (rule ANR_homeomorphic_ANR [OF \<open>ANR T\<close>])
       
  3559       apply (simp add: homeomorphic_def)
       
  3560       apply (rule_tac x=g in exI)
       
  3561       apply (rule_tac x=f in exI)
       
  3562       using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset)
       
  3563       apply (rule_tac x="f x" in image_eqI, auto)
       
  3564       done
       
  3565     have ANRI: "ANR ({x \<in> C. g x \<in> S} \<inter> {x \<in> C. g x \<in> T})"
       
  3566       apply (rule ANR_homeomorphic_ANR [OF \<open>ANR (S \<inter> T)\<close>])
       
  3567       apply (simp add: homeomorphic_def)
       
  3568       apply (rule_tac x=g in exI)
       
  3569       apply (rule_tac x=f in exI)
       
  3570       using hom
       
  3571       apply (auto simp: homeomorphism_def elim!: continuous_on_subset)
       
  3572       apply (rule_tac x="f x" in image_eqI, auto)
       
  3573       done
       
  3574     have "C = {x. x \<in> C \<and> g x \<in> S} \<union> {x. x \<in> C \<and> g x \<in> T}"
       
  3575       by auto (metis Un_iff hom homeomorphism_def imageI)
       
  3576     then show ?thesis
       
  3577       by (metis ANR_closed_Un_local_aux [OF US UT ANRS ANRT ANRI])
       
  3578   qed
       
  3579   then show ?thesis
       
  3580     by (auto simp: ANR_def)
       
  3581 qed    
       
  3582 
       
  3583 corollary ANR_closed_Un:
       
  3584   fixes S :: "'a::euclidean_space set"
       
  3585   shows "\<lbrakk>closed S; closed T; ANR S; ANR T; ANR (S \<inter> T)\<rbrakk> \<Longrightarrow> ANR (S \<union> T)"
       
  3586 by (simp add: ANR_closed_Un_local closedin_def diff_eq open_Compl openin_open_Int)
       
  3587 
       
  3588 lemma ANR_openin:
       
  3589   fixes S :: "'a::euclidean_space set"
       
  3590   assumes "ANR T" and opeTS: "openin (subtopology euclidean T) S"
       
  3591   shows "ANR S"
       
  3592 proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor)
       
  3593   fix f :: "'a \<times> real \<Rightarrow> 'a" and U C
       
  3594   assume contf: "continuous_on C f" and fim: "f ` C \<subseteq> S"
       
  3595      and cloUC: "closedin (subtopology euclidean U) C"
       
  3596   have "f ` C \<subseteq> T"
       
  3597     using fim opeTS openin_imp_subset by blast
       
  3598   obtain W g where "C \<subseteq> W"
       
  3599                and UW: "openin (subtopology euclidean U) W"
       
  3600                and contg: "continuous_on W g"
       
  3601                and gim: "g ` W \<subseteq> T"
       
  3602                and geq: "\<And>x. x \<in> C \<Longrightarrow> g x = f x"
       
  3603     apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> contf \<open>f ` C \<subseteq> T\<close> cloUC])
       
  3604     using fim by auto
       
  3605   show "\<exists>V g. C \<subseteq> V \<and> openin (subtopology euclidean U) V \<and> continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x\<in>C. g x = f x)"
       
  3606   proof (intro exI conjI)
       
  3607     show "C \<subseteq> {x \<in> W. g x \<in> S}"
       
  3608       using \<open>C \<subseteq> W\<close> fim geq by blast
       
  3609     show "openin (subtopology euclidean U) {x \<in> W. g x \<in> S}"
       
  3610       by (metis (mono_tags, lifting) UW contg continuous_openin_preimage gim opeTS openin_trans)
       
  3611     show "continuous_on {x \<in> W. g x \<in> S} g"
       
  3612       by (blast intro: continuous_on_subset [OF contg])
       
  3613     show "g ` {x \<in> W. g x \<in> S} \<subseteq> S"
       
  3614       using gim by blast
       
  3615     show "\<forall>x\<in>C. g x = f x"
       
  3616       using geq by blast
       
  3617   qed
       
  3618 qed
       
  3619 
       
  3620 lemma ENR_openin:
       
  3621     fixes S :: "'a::euclidean_space set"
       
  3622     assumes "ENR T" and opeTS: "openin (subtopology euclidean T) S"
       
  3623     shows "ENR S"
       
  3624   using assms apply (simp add: ENR_ANR)
       
  3625   using ANR_openin locally_open_subset by blast
       
  3626 
       
  3627 lemma ANR_neighborhood_retract:
       
  3628     fixes S :: "'a::euclidean_space set"
       
  3629     assumes "ANR U" "S retract_of T" "openin (subtopology euclidean U) T"
       
  3630     shows "ANR S"
       
  3631   using ANR_openin ANR_retract_of_ANR assms by blast
       
  3632 
       
  3633 lemma ENR_neighborhood_retract:
       
  3634     fixes S :: "'a::euclidean_space set"
       
  3635     assumes "ENR U" "S retract_of T" "openin (subtopology euclidean U) T"
       
  3636     shows "ENR S"
       
  3637   using ENR_openin ENR_retract_of_ENR assms by blast
       
  3638 
       
  3639 lemma ANR_rel_interior:
       
  3640   fixes S :: "'a::euclidean_space set"
       
  3641   shows "ANR S \<Longrightarrow> ANR(rel_interior S)"
       
  3642    by (blast intro: ANR_openin openin_set_rel_interior)
       
  3643 
       
  3644 lemma ANR_delete:
       
  3645   fixes S :: "'a::euclidean_space set"
       
  3646   shows "ANR S \<Longrightarrow> ANR(S - {a})"
       
  3647    by (blast intro: ANR_openin openin_delete openin_subtopology_self)
       
  3648 
       
  3649 lemma ENR_rel_interior:
       
  3650   fixes S :: "'a::euclidean_space set"
       
  3651   shows "ENR S \<Longrightarrow> ENR(rel_interior S)"
       
  3652    by (blast intro: ENR_openin openin_set_rel_interior)
       
  3653 
       
  3654 lemma ENR_delete:
       
  3655   fixes S :: "'a::euclidean_space set"
       
  3656   shows "ENR S \<Longrightarrow> ENR(S - {a})"
       
  3657    by (blast intro: ENR_openin openin_delete openin_subtopology_self)
       
  3658 
       
  3659 lemma open_imp_ENR: "open S \<Longrightarrow> ENR S"
       
  3660     using ENR_def by blast
       
  3661 
       
  3662 lemma open_imp_ANR:
       
  3663     fixes S :: "'a::euclidean_space set"
       
  3664     shows "open S \<Longrightarrow> ANR S"
       
  3665   by (simp add: ENR_imp_ANR open_imp_ENR)
       
  3666 
       
  3667 lemma ANR_ball [iff]:
       
  3668     fixes a :: "'a::euclidean_space"
       
  3669     shows "ANR(ball a r)"
       
  3670   by (simp add: convex_imp_ANR)
       
  3671 
       
  3672 lemma ENR_ball [iff]: "ENR(ball a r)"
       
  3673   by (simp add: open_imp_ENR)
       
  3674 
       
  3675 lemma AR_ball [simp]:
       
  3676     fixes a :: "'a::euclidean_space"
       
  3677     shows "AR(ball a r) \<longleftrightarrow> 0 < r"
       
  3678   by (auto simp: AR_ANR convex_imp_contractible)
       
  3679 
       
  3680 lemma ANR_cball [iff]:
       
  3681     fixes a :: "'a::euclidean_space"
       
  3682     shows "ANR(cball a r)"
       
  3683   by (simp add: convex_imp_ANR)
       
  3684 
       
  3685 lemma ENR_cball:
       
  3686     fixes a :: "'a::euclidean_space"
       
  3687     shows "ENR(cball a r)"
       
  3688   using ENR_convex_closed by blast
       
  3689 
       
  3690 lemma AR_cball [simp]:
       
  3691     fixes a :: "'a::euclidean_space"
       
  3692     shows "AR(cball a r) \<longleftrightarrow> 0 \<le> r"
       
  3693   by (auto simp: AR_ANR convex_imp_contractible)
       
  3694 
       
  3695 lemma ANR_box [iff]:
       
  3696     fixes a :: "'a::euclidean_space"
       
  3697     shows "ANR(cbox a b)" "ANR(box a b)"
       
  3698   by (auto simp: convex_imp_ANR open_imp_ANR)
       
  3699 
       
  3700 lemma ENR_box [iff]:
       
  3701     fixes a :: "'a::euclidean_space"
       
  3702     shows "ENR(cbox a b)" "ENR(box a b)"
       
  3703 apply (simp add: ENR_convex_closed closed_cbox)
       
  3704 by (simp add: open_box open_imp_ENR)
       
  3705 
       
  3706 lemma AR_box [simp]:
       
  3707     "AR(cbox a b) \<longleftrightarrow> cbox a b \<noteq> {}" "AR(box a b) \<longleftrightarrow> box a b \<noteq> {}"
       
  3708   by (auto simp: AR_ANR convex_imp_contractible)
       
  3709 
       
  3710 lemma ANR_interior:
       
  3711      fixes S :: "'a::euclidean_space set"
       
  3712      shows "ANR(interior S)"
       
  3713   by (simp add: open_imp_ANR)
       
  3714 
       
  3715 lemma ENR_interior:
       
  3716      fixes S :: "'a::euclidean_space set"
       
  3717      shows "ENR(interior S)"
       
  3718   by (simp add: open_imp_ENR)
       
  3719 
       
  3720 lemma AR_imp_contractible:
       
  3721     fixes S :: "'a::euclidean_space set"
       
  3722     shows "AR S \<Longrightarrow> contractible S"
       
  3723   by (simp add: AR_ANR)
       
  3724 
       
  3725 lemma ENR_imp_locally_compact:
       
  3726     fixes S :: "'a::euclidean_space set"
       
  3727     shows "ENR S \<Longrightarrow> locally compact S"
       
  3728   by (simp add: ENR_ANR)
       
  3729 
       
  3730 lemma ANR_imp_locally_path_connected:
       
  3731   fixes S :: "'a::euclidean_space set"
       
  3732   assumes "ANR S"
       
  3733     shows "locally path_connected S"
       
  3734 proof -
       
  3735   obtain U and T :: "('a \<times> real) set"
       
  3736      where "convex U" "U \<noteq> {}"
       
  3737        and UT: "closedin (subtopology euclidean U) T"
       
  3738        and "S homeomorphic T"
       
  3739     apply (rule homeomorphic_closedin_convex [of S])
       
  3740     using aff_dim_le_DIM [of S] apply auto
       
  3741     done
       
  3742   have "locally path_connected T"
       
  3743     by (meson ANR_imp_absolute_neighbourhood_retract \<open>S homeomorphic T\<close> \<open>closedin (subtopology euclidean U) T\<close> \<open>convex U\<close> assms convex_imp_locally_path_connected locally_open_subset retract_of_locally_path_connected)
       
  3744   then have S: "locally path_connected S"
       
  3745       if "openin (subtopology euclidean U) V" "T retract_of V" "U \<noteq> {}" for V
       
  3746     using \<open>S homeomorphic T\<close> homeomorphic_locally homeomorphic_path_connectedness by blast
       
  3747   show ?thesis
       
  3748     using assms
       
  3749     apply (clarsimp simp: ANR_def)
       
  3750     apply (drule_tac x=U in spec)
       
  3751     apply (drule_tac x=T in spec)
       
  3752     using \<open>S homeomorphic T\<close> \<open>U \<noteq> {}\<close> UT  apply (blast intro: S)
       
  3753     done
       
  3754 qed
       
  3755 
       
  3756 lemma ANR_imp_locally_connected:
       
  3757   fixes S :: "'a::euclidean_space set"
       
  3758   assumes "ANR S"
       
  3759     shows "locally connected S"
       
  3760 using locally_path_connected_imp_locally_connected ANR_imp_locally_path_connected assms by auto
       
  3761 
       
  3762 lemma AR_imp_locally_path_connected:
       
  3763   fixes S :: "'a::euclidean_space set"
       
  3764   assumes "AR S"
       
  3765     shows "locally path_connected S"
       
  3766 by (simp add: ANR_imp_locally_path_connected AR_imp_ANR assms)
       
  3767 
       
  3768 lemma AR_imp_locally_connected:
       
  3769   fixes S :: "'a::euclidean_space set"
       
  3770   assumes "AR S"
       
  3771     shows "locally connected S"
       
  3772 using ANR_imp_locally_connected AR_ANR assms by blast
       
  3773 
       
  3774 lemma ENR_imp_locally_path_connected:
       
  3775   fixes S :: "'a::euclidean_space set"
       
  3776   assumes "ENR S"
       
  3777     shows "locally path_connected S"
       
  3778 by (simp add: ANR_imp_locally_path_connected ENR_imp_ANR assms)
       
  3779 
       
  3780 lemma ENR_imp_locally_connected:
       
  3781   fixes S :: "'a::euclidean_space set"
       
  3782   assumes "ENR S"
       
  3783     shows "locally connected S"
       
  3784 using ANR_imp_locally_connected ENR_ANR assms by blast
       
  3785 
       
  3786 lemma ANR_Times:
       
  3787   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  3788   assumes "ANR S" "ANR T" shows "ANR(S \<times> T)"
       
  3789 proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor)
       
  3790   fix f :: " ('a \<times> 'b) \<times> real \<Rightarrow> 'a \<times> 'b" and U C
       
  3791   assume "continuous_on C f" and fim: "f ` C \<subseteq> S \<times> T"
       
  3792      and cloUC: "closedin (subtopology euclidean U) C"
       
  3793   have contf1: "continuous_on C (fst \<circ> f)"
       
  3794     by (simp add: \<open>continuous_on C f\<close> continuous_on_fst)
       
  3795   obtain W1 g where "C \<subseteq> W1"
       
  3796                and UW1: "openin (subtopology euclidean U) W1"
       
  3797                and contg: "continuous_on W1 g"
       
  3798                and gim: "g ` W1 \<subseteq> S"
       
  3799                and geq: "\<And>x. x \<in> C \<Longrightarrow> g x = (fst \<circ> f) x"
       
  3800     apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> contf1 _ cloUC])
       
  3801     using fim apply auto
       
  3802     done
       
  3803   have contf2: "continuous_on C (snd \<circ> f)"
       
  3804     by (simp add: \<open>continuous_on C f\<close> continuous_on_snd)
       
  3805   obtain W2 h where "C \<subseteq> W2"
       
  3806                and UW2: "openin (subtopology euclidean U) W2"
       
  3807                and conth: "continuous_on W2 h"
       
  3808                and him: "h ` W2 \<subseteq> T"
       
  3809                and heq: "\<And>x. x \<in> C \<Longrightarrow> h x = (snd \<circ> f) x"
       
  3810     apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> contf2 _ cloUC])
       
  3811     using fim apply auto
       
  3812     done
       
  3813   show "\<exists>V g. C \<subseteq> V \<and>
       
  3814                openin (subtopology euclidean U) V \<and>
       
  3815                continuous_on V g \<and> g ` V \<subseteq> S \<times> T \<and> (\<forall>x\<in>C. g x = f x)"
       
  3816   proof (intro exI conjI)
       
  3817     show "C \<subseteq> W1 \<inter> W2"
       
  3818       by (simp add: \<open>C \<subseteq> W1\<close> \<open>C \<subseteq> W2\<close>)
       
  3819     show "openin (subtopology euclidean U) (W1 \<inter> W2)"
       
  3820       by (simp add: UW1 UW2 openin_Int)
       
  3821     show  "continuous_on (W1 \<inter> W2) (\<lambda>x. (g x, h x))"
       
  3822       by (metis (no_types) contg conth continuous_on_Pair continuous_on_subset inf_commute inf_le1)
       
  3823     show  "(\<lambda>x. (g x, h x)) ` (W1 \<inter> W2) \<subseteq> S \<times> T"
       
  3824       using gim him by blast
       
  3825     show  "(\<forall>x\<in>C. (g x, h x) = f x)"
       
  3826       using geq heq by auto
       
  3827   qed
       
  3828 qed
       
  3829 
       
  3830 lemma AR_Times:
       
  3831   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  3832   assumes "AR S" "AR T" shows "AR(S \<times> T)"
       
  3833 using assms by (simp add: AR_ANR ANR_Times contractible_Times)
       
  3834 
       
  3835 
       
  3836 lemma ENR_rel_frontier_convex:
       
  3837   fixes S :: "'a::euclidean_space set"
       
  3838   assumes "bounded S" "convex S"
       
  3839     shows "ENR(rel_frontier S)"
       
  3840 proof (cases "S = {}")
       
  3841   case True then show ?thesis
       
  3842     by simp
       
  3843 next
       
  3844   case False
       
  3845   with assms have "rel_interior S \<noteq> {}"
       
  3846     by (simp add: rel_interior_eq_empty)
       
  3847   then obtain a where a: "a \<in> rel_interior S"
       
  3848     by auto
       
  3849   have ahS: "affine hull S - {a} \<subseteq> {x. closest_point (affine hull S) x \<noteq> a}"
       
  3850     by (auto simp: closest_point_self)
       
  3851   have "rel_frontier S retract_of affine hull S - {a}"
       
  3852     by (simp add: assms a rel_frontier_retract_of_punctured_affine_hull)
       
  3853   also have "... retract_of {x. closest_point (affine hull S) x \<noteq> a}"
       
  3854     apply (simp add: retract_of_def retraction_def ahS)
       
  3855     apply (rule_tac x="closest_point (affine hull S)" in exI)
       
  3856     apply (auto simp add: False closest_point_self affine_imp_convex closest_point_in_set continuous_on_closest_point)
       
  3857     done
       
  3858   finally have "rel_frontier S retract_of {x. closest_point (affine hull S) x \<noteq> a}" .
       
  3859   moreover have "openin (subtopology euclidean UNIV) {x \<in> UNIV. closest_point (affine hull S) x \<in> - {a}}"
       
  3860     apply (rule continuous_openin_preimage_gen)
       
  3861     apply (auto simp add: False affine_imp_convex continuous_on_closest_point)
       
  3862     done
       
  3863   ultimately show ?thesis
       
  3864     apply (simp add: ENR_def)
       
  3865     apply (rule_tac x = "{x. x \<in> UNIV \<and>
       
  3866                              closest_point (affine hull S) x \<in> (- {a})}" in exI)
       
  3867     apply (simp add: open_openin)
       
  3868     done
       
  3869 qed
       
  3870 
       
  3871 lemma ANR_rel_frontier_convex:
       
  3872                  fixes S :: "'a::euclidean_space set"
       
  3873   assumes "bounded S" "convex S"
       
  3874     shows "ANR(rel_frontier S)"
       
  3875 by (simp add: ENR_imp_ANR ENR_rel_frontier_convex assms)
       
  3876 
       
  3877 (*UNUSED
       
  3878 lemma ENR_Times:
       
  3879   fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
       
  3880   assumes "ENR S" "ENR T" shows "ENR(S \<times> T)"
       
  3881 using assms apply (simp add: ENR_ANR ANR_Times)
       
  3882 thm locally_compact_Times
       
  3883 oops
       
  3884   SIMP_TAC[ENR_ANR; ANR_PCROSS; LOCALLY_COMPACT_PCROSS]);;
       
  3885 *)
       
  3886 
       
  3887 subsection\<open>Borsuk homotopy extension theorem\<close>
       
  3888 
       
  3889 text\<open>It's only this late so we can use the concept of retraction,
       
  3890   saying that the domain sets or range set are ENRs.\<close>
       
  3891 
       
  3892 theorem Borsuk_homotopy_extension_homotopic:
       
  3893   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
  3894   assumes cloTS: "closedin (subtopology euclidean T) S"
       
  3895       and anr: "(ANR S \<and> ANR T) \<or> ANR U"
       
  3896       and contf: "continuous_on T f"
       
  3897       and "f ` T \<subseteq> U"
       
  3898       and "homotopic_with (\<lambda>x. True) S U f g"
       
  3899    obtains g' where "homotopic_with (\<lambda>x. True) T U f g'"
       
  3900                     "continuous_on T g'" "image g' T \<subseteq> U"
       
  3901                     "\<And>x. x \<in> S \<Longrightarrow> g' x = g x"
       
  3902 proof -
       
  3903   have "S \<subseteq> T" using assms closedin_imp_subset by blast
       
  3904   obtain h where conth: "continuous_on ({0..1} \<times> S) h"
       
  3905              and him: "h ` ({0..1} \<times> S) \<subseteq> U"
       
  3906              and [simp]: "\<And>x. h(0, x) = f x" "\<And>x. h(1::real, x) = g x"
       
  3907        using assms by (auto simp: homotopic_with_def)
       
  3908   define h' where "h' \<equiv>  \<lambda>z. if snd z \<in> S then h z else (f o snd) z"
       
  3909   define B where "B \<equiv> {0::real} \<times> T \<union> {0..1} \<times> S"
       
  3910   have clo0T: "closedin (subtopology euclidean ({0..1} \<times> T)) ({0::real} \<times> T)"
       
  3911     by (simp add: closedin_subtopology_refl closedin_Times)
       
  3912   moreover have cloT1S: "closedin (subtopology euclidean ({0..1} \<times> T)) ({0..1} \<times> S)"
       
  3913     by (simp add: closedin_subtopology_refl closedin_Times assms)
       
  3914   ultimately have clo0TB:"closedin (subtopology euclidean ({0..1} \<times> T)) B"
       
  3915     by (auto simp: B_def)
       
  3916   have cloBS: "closedin (subtopology euclidean B) ({0..1} \<times> S)"
       
  3917     by (metis (no_types) Un_subset_iff B_def closedin_subset_trans [OF cloT1S] clo0TB closedin_imp_subset closedin_self)
       
  3918   moreover have cloBT: "closedin (subtopology euclidean B) ({0} \<times> T)"
       
  3919     using \<open>S \<subseteq> T\<close> closedin_subset_trans [OF clo0T]
       
  3920     by (metis B_def Un_upper1 clo0TB closedin_closed inf_le1)
       
  3921   moreover have "continuous_on ({0} \<times> T) (f \<circ> snd)"
       
  3922     apply (rule continuous_intros)+
       
  3923     apply (simp add: contf)
       
  3924     done
       
  3925   ultimately have conth': "continuous_on B h'"
       
  3926     apply (simp add: h'_def B_def Un_commute [of "{0} \<times> T"])
       
  3927     apply (auto intro!: continuous_on_cases_local conth)
       
  3928     done
       
  3929   have "image h' B \<subseteq> U"
       
  3930     using \<open>f ` T \<subseteq> U\<close> him by (auto simp: h'_def B_def)
       
  3931   obtain V k where "B \<subseteq> V" and opeTV: "openin (subtopology euclidean ({0..1} \<times> T)) V"
       
  3932                and contk: "continuous_on V k" and kim: "k ` V \<subseteq> U"
       
  3933                and keq: "\<And>x. x \<in> B \<Longrightarrow> k x = h' x"
       
  3934   using anr
       
  3935   proof
       
  3936     assume ST: "ANR S \<and> ANR T"
       
  3937     have eq: "({0} \<times> T \<inter> {0..1} \<times> S) = {0::real} \<times> S"
       
  3938       using \<open>S \<subseteq> T\<close> by auto
       
  3939     have "ANR B"
       
  3940       apply (simp add: B_def)
       
  3941       apply (rule ANR_closed_Un_local)
       
  3942           apply (metis cloBT B_def)
       
  3943          apply (metis Un_commute cloBS B_def)
       
  3944         apply (simp_all add: ANR_Times convex_imp_ANR ANR_singleton ST eq)
       
  3945       done
       
  3946     note Vk = that
       
  3947     have *: thesis if "openin (subtopology euclidean ({0..1::real} \<times> T)) V"
       
  3948                       "retraction V B r" for V r
       
  3949       using that
       
  3950       apply (clarsimp simp add: retraction_def)
       
  3951       apply (rule Vk [of V "h' o r"], assumption+)
       
  3952         apply (metis continuous_on_compose conth' continuous_on_subset) 
       
  3953       using \<open>h' ` B \<subseteq> U\<close> apply force+
       
  3954       done
       
  3955     show thesis
       
  3956         apply (rule ANR_imp_neighbourhood_retract [OF \<open>ANR B\<close> clo0TB])
       
  3957         apply (auto simp: ANR_Times ANR_singleton ST retract_of_def *)
       
  3958         done
       
  3959   next
       
  3960     assume "ANR U"
       
  3961     with ANR_imp_absolute_neighbourhood_extensor \<open>h' ` B \<subseteq> U\<close> clo0TB conth' that
       
  3962     show ?thesis by blast
       
  3963   qed
       
  3964   define S' where "S' \<equiv> {x. \<exists>u::real. u \<in> {0..1} \<and> (u, x::'a) \<in> {0..1} \<times> T - V}"
       
  3965   have "closedin (subtopology euclidean T) S'"
       
  3966     unfolding S'_def
       
  3967     apply (rule closedin_compact_projection, blast)
       
  3968     using closedin_self opeTV by blast
       
  3969   have S'_def: "S' = {x. \<exists>u::real.  (u, x::'a) \<in> {0..1} \<times> T - V}"
       
  3970     by (auto simp: S'_def)
       
  3971   have cloTS': "closedin (subtopology euclidean T) S'"
       
  3972     using S'_def \<open>closedin (subtopology euclidean T) S'\<close> by blast
       
  3973   have "S \<inter> S' = {}"
       
  3974     using S'_def B_def \<open>B \<subseteq> V\<close> by force
       
  3975   obtain a :: "'a \<Rightarrow> real" where conta: "continuous_on T a"
       
  3976       and "\<And>x. x \<in> T \<Longrightarrow> a x \<in> closed_segment 1 0"
       
  3977       and a1: "\<And>x. x \<in> S \<Longrightarrow> a x = 1"
       
  3978       and a0: "\<And>x. x \<in> S' \<Longrightarrow> a x = 0"
       
  3979     apply (rule Urysohn_local [OF cloTS cloTS' \<open>S \<inter> S' = {}\<close>, of 1 0], blast)
       
  3980     done
       
  3981   then have ain: "\<And>x. x \<in> T \<Longrightarrow> a x \<in> {0..1}"
       
  3982     using closed_segment_eq_real_ivl by auto
       
  3983   have inV: "(u * a t, t) \<in> V" if "t \<in> T" "0 \<le> u" "u \<le> 1" for t u
       
  3984   proof (rule ccontr)
       
  3985     assume "(u * a t, t) \<notin> V"
       
  3986     with ain [OF \<open>t \<in> T\<close>] have "a t = 0"
       
  3987       apply simp
       
  3988       apply (rule a0)
       
  3989       by (metis (no_types, lifting) Diff_iff S'_def SigmaI atLeastAtMost_iff mem_Collect_eq mult_le_one mult_nonneg_nonneg that)
       
  3990     show False
       
  3991       using B_def \<open>(u * a t, t) \<notin> V\<close> \<open>B \<subseteq> V\<close> \<open>a t = 0\<close> that by auto
       
  3992   qed
       
  3993   show ?thesis
       
  3994   proof
       
  3995     show hom: "homotopic_with (\<lambda>x. True) T U f (\<lambda>x. k (a x, x))"
       
  3996     proof (simp add: homotopic_with, intro exI conjI)
       
  3997       show "continuous_on ({0..1} \<times> T) (k \<circ> (\<lambda>z. (fst z *\<^sub>R (a \<circ> snd) z, snd z)))"
       
  3998         apply (intro continuous_on_compose continuous_intros)
       
  3999         apply (rule continuous_on_subset [OF conta], force)
       
  4000         apply (rule continuous_on_subset [OF contk])
       
  4001         apply (force intro: inV)
       
  4002         done
       
  4003       show "(k \<circ> (\<lambda>z. (fst z *\<^sub>R (a \<circ> snd) z, snd z))) ` ({0..1} \<times> T) \<subseteq> U"
       
  4004         using inV kim by auto
       
  4005       show "\<forall>x\<in>T. (k \<circ> (\<lambda>z. (fst z *\<^sub>R (a \<circ> snd) z, snd z))) (0, x) = f x"
       
  4006         by (simp add: B_def h'_def keq)
       
  4007       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)"
       
  4008         by auto
       
  4009     qed
       
  4010   show "continuous_on T (\<lambda>x. k (a x, x))"
       
  4011     using hom homotopic_with_imp_continuous by blast
       
  4012   show "(\<lambda>x. k (a x, x)) ` T \<subseteq> U"
       
  4013   proof clarify
       
  4014     fix t
       
  4015     assume "t \<in> T"
       
  4016     show "k (a t, t) \<in> U"
       
  4017       by (metis \<open>t \<in> T\<close> image_subset_iff inV kim not_one_le_zero linear mult_cancel_right1)
       
  4018   qed
       
  4019   show "\<And>x. x \<in> S \<Longrightarrow> k (a x, x) = g x"
       
  4020     by (simp add: B_def a1 h'_def keq)
       
  4021   qed
       
  4022 qed
       
  4023 
       
  4024 
       
  4025 corollary nullhomotopic_into_ANR_extension:
       
  4026   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
  4027   assumes "closed S"
       
  4028       and contf: "continuous_on S f"
       
  4029       and "ANR T"
       
  4030       and fim: "f ` S \<subseteq> T"
       
  4031       and "S \<noteq> {}"
       
  4032    shows "(\<exists>c. homotopic_with (\<lambda>x. True) S T f (\<lambda>x. c)) \<longleftrightarrow>
       
  4033           (\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> T \<and> (\<forall>x \<in> S. g x = f x))"
       
  4034        (is "?lhs = ?rhs")
       
  4035 proof
       
  4036   assume ?lhs
       
  4037   then obtain c where c: "homotopic_with (\<lambda>x. True) S T (\<lambda>x. c) f"
       
  4038     by (blast intro: homotopic_with_symD elim: )
       
  4039   have "closedin (subtopology euclidean UNIV) S"
       
  4040     using \<open>closed S\<close> closed_closedin by fastforce
       
  4041   then obtain g where "continuous_on UNIV g" "range g \<subseteq> T"
       
  4042                       "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
       
  4043     apply (rule Borsuk_homotopy_extension_homotopic [OF _ _ continuous_on_const _ c, where T=UNIV])
       
  4044     using \<open>ANR T\<close> \<open>S \<noteq> {}\<close> c homotopic_with_imp_subset1 apply fastforce+
       
  4045     done
       
  4046   then show ?rhs by blast
       
  4047 next
       
  4048   assume ?rhs
       
  4049   then obtain g where "continuous_on UNIV g" "range g \<subseteq> T" "\<And>x. x\<in>S \<Longrightarrow> g x = f x"
       
  4050     by blast
       
  4051   then obtain c where "homotopic_with (\<lambda>h. True) UNIV T g (\<lambda>x. c)"
       
  4052     using nullhomotopic_from_contractible [of UNIV g T] contractible_UNIV by blast
       
  4053   then show ?lhs
       
  4054     apply (rule_tac x="c" in exI)
       
  4055     apply (rule homotopic_with_eq [of _ _ _ g "\<lambda>x. c"])
       
  4056     apply (rule homotopic_with_subset_left)
       
  4057     apply (auto simp add: \<open>\<And>x. x \<in> S \<Longrightarrow> g x = f x\<close>)
       
  4058     done
       
  4059 qed
       
  4060 
       
  4061 corollary nullhomotopic_into_rel_frontier_extension:
       
  4062   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
  4063   assumes "closed S"
       
  4064       and contf: "continuous_on S f"
       
  4065       and "convex T" "bounded T"
       
  4066       and fim: "f ` S \<subseteq> rel_frontier T"
       
  4067       and "S \<noteq> {}"
       
  4068    shows "(\<exists>c. homotopic_with (\<lambda>x. True) S (rel_frontier T) f (\<lambda>x. c)) \<longleftrightarrow>
       
  4069           (\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> rel_frontier T \<and> (\<forall>x \<in> S. g x = f x))"
       
  4070 by (simp add: nullhomotopic_into_ANR_extension assms ANR_rel_frontier_convex)
       
  4071 
       
  4072 corollary nullhomotopic_into_sphere_extension:
       
  4073   fixes f :: "'a::euclidean_space \<Rightarrow> 'b :: euclidean_space"
       
  4074   assumes "closed S" and contf: "continuous_on S f"
       
  4075       and "S \<noteq> {}" and fim: "f ` S \<subseteq> sphere a r"
       
  4076     shows "((\<exists>c. homotopic_with (\<lambda>x. True) S (sphere a r) f (\<lambda>x. c)) \<longleftrightarrow>
       
  4077            (\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> sphere a r \<and> (\<forall>x \<in> S. g x = f x)))"
       
  4078            (is "?lhs = ?rhs")
       
  4079 proof (cases "r = 0")
       
  4080   case True with fim show ?thesis
       
  4081     apply (auto simp: )
       
  4082     using fim continuous_on_const apply fastforce
       
  4083     by (metis contf contractible_sing nullhomotopic_into_contractible)
       
  4084 next
       
  4085   case False
       
  4086   then have eq: "sphere a r = rel_frontier (cball a r)" by simp
       
  4087   show ?thesis
       
  4088     using fim unfolding eq
       
  4089     apply (rule nullhomotopic_into_rel_frontier_extension [OF \<open>closed S\<close> contf convex_cball bounded_cball])
       
  4090     apply (rule \<open>S \<noteq> {}\<close>)
       
  4091     done
       
  4092 qed
       
  4093 
       
  4094 proposition Borsuk_map_essential_bounded_component:
       
  4095   fixes a :: "'a :: euclidean_space"
       
  4096   assumes "compact S" and "a \<notin> S"
       
  4097    shows "bounded (connected_component_set (- S) a) \<longleftrightarrow>
       
  4098           ~(\<exists>c. homotopic_with (\<lambda>x. True) S (sphere 0 1)
       
  4099                                (\<lambda>x. inverse(norm(x - a)) *\<^sub>R (x - a)) (\<lambda>x. c))"
       
  4100    (is "?lhs = ?rhs")
       
  4101 proof (cases "S = {}")
       
  4102   case True then show ?thesis
       
  4103     by simp
       
  4104 next
       
  4105   case False
       
  4106   have "closed S" "bounded S"
       
  4107     using \<open>compact S\<close> compact_eq_bounded_closed by auto
       
  4108   have s01: "(\<lambda>x. (x - a) /\<^sub>R norm (x - a)) ` S \<subseteq> sphere 0 1"
       
  4109     using \<open>a \<notin> S\<close>  by clarsimp (metis dist_eq_0_iff dist_norm mult.commute right_inverse)
       
  4110   have aincc: "a \<in> connected_component_set (- S) a"
       
  4111     by (simp add: \<open>a \<notin> S\<close>)
       
  4112   obtain r where "r>0" and r: "S \<subseteq> ball 0 r"
       
  4113     using bounded_subset_ballD \<open>bounded S\<close> by blast
       
  4114   have "~ ?rhs \<longleftrightarrow> ~ ?lhs"
       
  4115   proof
       
  4116     assume notr: "~ ?rhs"
       
  4117     have nog: "\<nexists>g. continuous_on (S \<union> connected_component_set (- S) a) g \<and>
       
  4118                    g ` (S \<union> connected_component_set (- S) a) \<subseteq> sphere 0 1 \<and>
       
  4119                    (\<forall>x\<in>S. g x = (x - a) /\<^sub>R norm (x - a))"
       
  4120          if "bounded (connected_component_set (- S) a)"
       
  4121       apply (rule non_extensible_Borsuk_map [OF \<open>compact S\<close> componentsI _ aincc])
       
  4122       using  \<open>a \<notin> S\<close> that by auto
       
  4123     obtain g where "range g \<subseteq> sphere 0 1" "continuous_on UNIV g"
       
  4124                         "\<And>x. x \<in> S \<Longrightarrow> g x = (x - a) /\<^sub>R norm (x - a)"
       
  4125       using notr
       
  4126       by (auto simp add: nullhomotopic_into_sphere_extension
       
  4127                  [OF \<open>closed S\<close> continuous_on_Borsuk_map [OF \<open>a \<notin> S\<close>] False s01])
       
  4128     with \<open>a \<notin> S\<close> show  "~ ?lhs"
       
  4129       apply (clarsimp simp: Borsuk_map_into_sphere [of a S, symmetric] dest!: nog)
       
  4130       apply (drule_tac x="g" in spec)
       
  4131       using continuous_on_subset by fastforce 
       
  4132   next
       
  4133     assume "~ ?lhs"
       
  4134     then obtain b where b: "b \<in> connected_component_set (- S) a" and "r \<le> norm b"
       
  4135       using bounded_iff linear by blast
       
  4136     then have bnot: "b \<notin> ball 0 r"
       
  4137       by simp
       
  4138     have "homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. (x - a) /\<^sub>R norm (x - a))
       
  4139                                                    (\<lambda>x. (x - b) /\<^sub>R norm (x - b))"
       
  4140       apply (rule Borsuk_maps_homotopic_in_path_component)
       
  4141       using \<open>closed S\<close> b open_Compl open_path_connected_component apply fastforce
       
  4142       done
       
  4143     moreover
       
  4144     obtain c where "homotopic_with (\<lambda>x. True) (ball 0 r) (sphere 0 1)
       
  4145                                    (\<lambda>x. inverse (norm (x - b)) *\<^sub>R (x - b)) (\<lambda>x. c)"
       
  4146     proof (rule nullhomotopic_from_contractible)
       
  4147       show "contractible (ball (0::'a) r)"
       
  4148         by (metis convex_imp_contractible convex_ball)
       
  4149       show "continuous_on (ball 0 r) (\<lambda>x. inverse(norm (x - b)) *\<^sub>R (x - b))"
       
  4150         by (rule continuous_on_Borsuk_map [OF bnot])
       
  4151       show "(\<lambda>x. (x - b) /\<^sub>R norm (x - b)) ` ball 0 r \<subseteq> sphere 0 1"
       
  4152         using bnot Borsuk_map_into_sphere by blast
       
  4153     qed blast
       
  4154     ultimately have "homotopic_with (\<lambda>x. True) S (sphere 0 1)
       
  4155                          (\<lambda>x. (x - a) /\<^sub>R norm (x - a)) (\<lambda>x. c)"
       
  4156       by (meson homotopic_with_subset_left homotopic_with_trans r)
       
  4157     then show "~ ?rhs"
       
  4158       by blast
       
  4159   qed
       
  4160   then show ?thesis by blast
       
  4161 qed
       
  4162 
       
  4163 
       
  4164 end