| author | eberlm <eberlm@in.tum.de> | 
| Sun, 22 Jul 2018 19:29:51 +0200 | |
| changeset 68677 | 99b1cf1e2d48 | 
| parent 68621 | 27432da24236 | 
| child 69286 | e4d5a07fecb6 | 
| permissions | -rw-r--r-- | 
| 53674 | 1 | (* Author: John Harrison | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2 | Author: Robert Himmelmann, TU Muenchen (Translation from HOL light) and LCP | 
| 53674 | 3 | *) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 5 | (* At the moment this is just Brouwer's fixpoint theorem. The proof is from *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 6 | (* Kuhn: "some combinatorial lemmas in topology", IBM J. v4. (1960) p. 518 *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 7 | (* See "http://www.research.ibm.com/journal/rd/045/ibmrd0405K.pdf". *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 8 | (* *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 9 | (* The script below is quite messy, but at least we avoid formalizing any *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 10 | (* topological machinery; we don't even use barycentric subdivision; this is *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 11 | (* the big advantage of Kuhn's proof over the usual Sperner's lemma one. *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 12 | (* *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 13 | (* (c) Copyright, John Harrison 1998-2008 *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 14 | |
| 68617 | 15 | section \<open>Brouwer's Fixed Point Theorem\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 16 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 17 | theory Brouwer_Fixpoint | 
| 63129 | 18 | imports Path_Connected Homeomorphism | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 19 | begin | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 20 | |
| 68617 | 21 | (* FIXME mv topology euclidean space *) | 
| 22 | subsection \<open>Retractions\<close> | |
| 23 | ||
| 24 | 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)" | |
| 25 | ||
| 26 | definition retract_of (infixl "retract'_of" 50) | |
| 27 | where "(T retract_of S) \<longleftrightarrow> (\<exists>r. retraction S T r)" | |
| 28 | ||
| 29 | lemma retraction_idempotent: "retraction S T r \<Longrightarrow> x \<in> S \<Longrightarrow> r (r x) = r x" | |
| 30 | unfolding retraction_def by auto | |
| 31 | ||
| 32 | text \<open>Preservation of fixpoints under (more general notion of) retraction\<close> | |
| 33 | ||
| 34 | lemma invertible_fixpoint_property: | |
| 35 | fixes S :: "'a::euclidean_space set" | |
| 36 | and T :: "'b::euclidean_space set" | |
| 37 | assumes contt: "continuous_on T i" | |
| 38 | and "i ` T \<subseteq> S" | |
| 39 | and contr: "continuous_on S r" | |
| 40 | and "r ` S \<subseteq> T" | |
| 41 | and ri: "\<And>y. y \<in> T \<Longrightarrow> r (i y) = y" | |
| 42 | and FP: "\<And>f. \<lbrakk>continuous_on S f; f ` S \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>x\<in>S. f x = x" | |
| 43 | and contg: "continuous_on T g" | |
| 44 | and "g ` T \<subseteq> T" | |
| 45 | obtains y where "y \<in> T" and "g y = y" | |
| 46 | proof - | |
| 47 | have "\<exists>x\<in>S. (i \<circ> g \<circ> r) x = x" | |
| 48 | proof (rule FP) | |
| 49 | show "continuous_on S (i \<circ> g \<circ> r)" | |
| 50 | by (meson contt contr assms(4) contg assms(8) continuous_on_compose continuous_on_subset) | |
| 51 | show "(i \<circ> g \<circ> r) ` S \<subseteq> S" | |
| 52 | using assms(2,4,8) by force | |
| 53 | qed | |
| 54 | then obtain x where x: "x \<in> S" "(i \<circ> g \<circ> r) x = x" .. | |
| 55 | then have *: "g (r x) \<in> T" | |
| 56 | using assms(4,8) by auto | |
| 57 | have "r ((i \<circ> g \<circ> r) x) = r x" | |
| 58 | using x by auto | |
| 59 | then show ?thesis | |
| 60 | using "*" ri that by auto | |
| 61 | qed | |
| 62 | ||
| 63 | lemma homeomorphic_fixpoint_property: | |
| 64 | fixes S :: "'a::euclidean_space set" | |
| 65 | and T :: "'b::euclidean_space set" | |
| 66 | assumes "S homeomorphic T" | |
| 67 | shows "(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> S \<longrightarrow> (\<exists>x\<in>S. f x = x)) \<longleftrightarrow> | |
| 68 | (\<forall>g. continuous_on T g \<and> g ` T \<subseteq> T \<longrightarrow> (\<exists>y\<in>T. g y = y))" | |
| 69 | (is "?lhs = ?rhs") | |
| 70 | proof - | |
| 71 | obtain r i where r: | |
| 72 | "\<forall>x\<in>S. i (r x) = x" "r ` S = T" "continuous_on S r" | |
| 73 | "\<forall>y\<in>T. r (i y) = y" "i ` T = S" "continuous_on T i" | |
| 74 | using assms unfolding homeomorphic_def homeomorphism_def by blast | |
| 75 | show ?thesis | |
| 76 | proof | |
| 77 | assume ?lhs | |
| 78 | with r show ?rhs | |
| 79 | by (metis invertible_fixpoint_property[of T i S r] order_refl) | |
| 80 | next | |
| 81 | assume ?rhs | |
| 82 | with r show ?lhs | |
| 83 | by (metis invertible_fixpoint_property[of S r T i] order_refl) | |
| 84 | qed | |
| 85 | qed | |
| 86 | ||
| 87 | lemma retract_fixpoint_property: | |
| 88 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 89 | and S :: "'a set" | |
| 90 | assumes "T retract_of S" | |
| 91 | and FP: "\<And>f. \<lbrakk>continuous_on S f; f ` S \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>x\<in>S. f x = x" | |
| 92 | and contg: "continuous_on T g" | |
| 93 | and "g ` T \<subseteq> T" | |
| 94 | obtains y where "y \<in> T" and "g y = y" | |
| 95 | proof - | |
| 96 | obtain h where "retraction S T h" | |
| 97 | using assms(1) unfolding retract_of_def .. | |
| 98 | then show ?thesis | |
| 99 | unfolding retraction_def | |
| 100 | using invertible_fixpoint_property[OF continuous_on_id _ _ _ _ FP] | |
| 101 | by (metis assms(4) contg image_ident that) | |
| 102 | qed | |
| 103 | ||
| 104 | lemma retraction: | |
| 105 | "retraction S T r \<longleftrightarrow> | |
| 106 | T \<subseteq> S \<and> continuous_on S r \<and> r ` S = T \<and> (\<forall>x \<in> T. r x = x)" | |
| 107 | by (force simp: retraction_def) | |
| 108 | ||
| 109 | lemma retract_of_imp_extensible: | |
| 110 | assumes "S retract_of T" and "continuous_on S f" and "f ` S \<subseteq> U" | |
| 111 | obtains g where "continuous_on T g" "g ` T \<subseteq> U" "\<And>x. x \<in> S \<Longrightarrow> g x = f x" | |
| 112 | using assms | |
| 113 | apply (clarsimp simp add: retract_of_def retraction) | |
| 114 | apply (rule_tac g = "f \<circ> r" in that) | |
| 115 | apply (auto simp: continuous_on_compose2) | |
| 116 | done | |
| 117 | ||
| 118 | lemma idempotent_imp_retraction: | |
| 119 | assumes "continuous_on S f" and "f ` S \<subseteq> S" and "\<And>x. x \<in> S \<Longrightarrow> f(f x) = f x" | |
| 120 | shows "retraction S (f ` S) f" | |
| 121 | by (simp add: assms retraction) | |
| 122 | ||
| 123 | lemma retraction_subset: | |
| 124 | assumes "retraction S T r" and "T \<subseteq> s'" and "s' \<subseteq> S" | |
| 125 | shows "retraction s' T r" | |
| 126 | unfolding retraction_def | |
| 127 | by (metis assms continuous_on_subset image_mono retraction) | |
| 128 | ||
| 129 | lemma retract_of_subset: | |
| 130 | assumes "T retract_of S" and "T \<subseteq> s'" and "s' \<subseteq> S" | |
| 131 | shows "T retract_of s'" | |
| 132 | by (meson assms retract_of_def retraction_subset) | |
| 133 | ||
| 134 | lemma retraction_refl [simp]: "retraction S S (\<lambda>x. x)" | |
| 135 | by (simp add: continuous_on_id retraction) | |
| 136 | ||
| 137 | lemma retract_of_refl [iff]: "S retract_of S" | |
| 138 | unfolding retract_of_def retraction_def | |
| 139 | using continuous_on_id by blast | |
| 140 | ||
| 141 | lemma retract_of_imp_subset: | |
| 142 | "S retract_of T \<Longrightarrow> S \<subseteq> T" | |
| 143 | by (simp add: retract_of_def retraction_def) | |
| 144 | ||
| 145 | lemma retract_of_empty [simp]: | |
| 146 |      "({} retract_of S) \<longleftrightarrow> S = {}"  "(S retract_of {}) \<longleftrightarrow> S = {}"
 | |
| 147 | by (auto simp: retract_of_def retraction_def) | |
| 148 | ||
| 149 | lemma retract_of_singleton [iff]: "({x} retract_of S) \<longleftrightarrow> x \<in> S"
 | |
| 150 | unfolding retract_of_def retraction_def by force | |
| 151 | ||
| 152 | lemma retraction_comp: | |
| 153 | "\<lbrakk>retraction S T f; retraction T U g\<rbrakk> | |
| 154 | \<Longrightarrow> retraction S U (g \<circ> f)" | |
| 155 | apply (auto simp: retraction_def intro: continuous_on_compose2) | |
| 156 | by blast | |
| 157 | ||
| 158 | lemma retract_of_trans [trans]: | |
| 159 | assumes "S retract_of T" and "T retract_of U" | |
| 160 | shows "S retract_of U" | |
| 161 | using assms by (auto simp: retract_of_def intro: retraction_comp) | |
| 162 | ||
| 163 | lemma closedin_retract: | |
| 164 | fixes S :: "'a :: real_normed_vector set" | |
| 165 | assumes "S retract_of T" | |
| 166 | shows "closedin (subtopology euclidean T) S" | |
| 167 | proof - | |
| 168 | obtain r where "S \<subseteq> T" "continuous_on T r" "r ` T \<subseteq> S" "\<And>x. x \<in> S \<Longrightarrow> r x = x" | |
| 169 | using assms by (auto simp: retract_of_def retraction_def) | |
| 170 |   then have S: "S = {x \<in> T. (norm(r x - x)) = 0}" by auto
 | |
| 171 | show ?thesis | |
| 172 | apply (subst S) | |
| 173 | apply (rule continuous_closedin_preimage_constant) | |
| 174 | by (simp add: \<open>continuous_on T r\<close> continuous_on_diff continuous_on_id continuous_on_norm) | |
| 175 | qed | |
| 176 | ||
| 177 | lemma closedin_self [simp]: | |
| 178 | fixes S :: "'a :: real_normed_vector set" | |
| 179 | shows "closedin (subtopology euclidean S) S" | |
| 180 | by (simp add: closedin_retract) | |
| 181 | ||
| 182 | lemma retract_of_contractible: | |
| 183 | assumes "contractible T" "S retract_of T" | |
| 184 | shows "contractible S" | |
| 185 | using assms | |
| 186 | apply (clarsimp simp add: retract_of_def contractible_def retraction_def homotopic_with) | |
| 187 | apply (rule_tac x="r a" in exI) | |
| 188 | apply (rule_tac x="r \<circ> h" in exI) | |
| 189 | apply (intro conjI continuous_intros continuous_on_compose) | |
| 190 | apply (erule continuous_on_subset | force)+ | |
| 191 | done | |
| 192 | ||
| 193 | lemma retract_of_compact: | |
| 194 | "\<lbrakk>compact T; S retract_of T\<rbrakk> \<Longrightarrow> compact S" | |
| 195 | by (metis compact_continuous_image retract_of_def retraction) | |
| 196 | ||
| 197 | lemma retract_of_closed: | |
| 198 | fixes S :: "'a :: real_normed_vector set" | |
| 199 | shows "\<lbrakk>closed T; S retract_of T\<rbrakk> \<Longrightarrow> closed S" | |
| 200 | by (metis closedin_retract closedin_closed_eq) | |
| 201 | ||
| 202 | lemma retract_of_connected: | |
| 203 | "\<lbrakk>connected T; S retract_of T\<rbrakk> \<Longrightarrow> connected S" | |
| 204 | by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction) | |
| 205 | ||
| 206 | lemma retract_of_path_connected: | |
| 207 | "\<lbrakk>path_connected T; S retract_of T\<rbrakk> \<Longrightarrow> path_connected S" | |
| 208 | by (metis path_connected_continuous_image retract_of_def retraction) | |
| 209 | ||
| 210 | lemma retract_of_simply_connected: | |
| 211 | "\<lbrakk>simply_connected T; S retract_of T\<rbrakk> \<Longrightarrow> simply_connected S" | |
| 212 | apply (simp add: retract_of_def retraction_def, clarify) | |
| 213 | apply (rule simply_connected_retraction_gen) | |
| 214 | apply (force simp: continuous_on_id elim!: continuous_on_subset)+ | |
| 215 | done | |
| 216 | ||
| 217 | lemma retract_of_homotopically_trivial: | |
| 218 | assumes ts: "T retract_of S" | |
| 219 | and hom: "\<And>f g. \<lbrakk>continuous_on U f; f ` U \<subseteq> S; | |
| 220 | continuous_on U g; g ` U \<subseteq> S\<rbrakk> | |
| 221 | \<Longrightarrow> homotopic_with (\<lambda>x. True) U S f g" | |
| 222 | and "continuous_on U f" "f ` U \<subseteq> T" | |
| 223 | and "continuous_on U g" "g ` U \<subseteq> T" | |
| 224 | shows "homotopic_with (\<lambda>x. True) U T f g" | |
| 225 | proof - | |
| 226 | obtain r where "r ` S \<subseteq> S" "continuous_on S r" "\<forall>x\<in>S. r (r x) = r x" "T = r ` S" | |
| 227 | using ts by (auto simp: retract_of_def retraction) | |
| 228 | then obtain k where "Retracts S r T k" | |
| 229 | unfolding Retracts_def | |
| 230 | by (metis continuous_on_subset dual_order.trans image_iff image_mono) | |
| 231 | then show ?thesis | |
| 232 | apply (rule Retracts.homotopically_trivial_retraction_gen) | |
| 233 | using assms | |
| 234 | apply (force simp: hom)+ | |
| 235 | done | |
| 236 | qed | |
| 237 | ||
| 238 | lemma retract_of_homotopically_trivial_null: | |
| 239 | assumes ts: "T retract_of S" | |
| 240 | and hom: "\<And>f. \<lbrakk>continuous_on U f; f ` U \<subseteq> S\<rbrakk> | |
| 241 | \<Longrightarrow> \<exists>c. homotopic_with (\<lambda>x. True) U S f (\<lambda>x. c)" | |
| 242 | and "continuous_on U f" "f ` U \<subseteq> T" | |
| 243 | obtains c where "homotopic_with (\<lambda>x. True) U T f (\<lambda>x. c)" | |
| 244 | proof - | |
| 245 | obtain r where "r ` S \<subseteq> S" "continuous_on S r" "\<forall>x\<in>S. r (r x) = r x" "T = r ` S" | |
| 246 | using ts by (auto simp: retract_of_def retraction) | |
| 247 | then obtain k where "Retracts S r T k" | |
| 248 | unfolding Retracts_def | |
| 249 | by (metis continuous_on_subset dual_order.trans image_iff image_mono) | |
| 250 | then show ?thesis | |
| 251 | apply (rule Retracts.homotopically_trivial_retraction_null_gen) | |
| 252 | apply (rule TrueI refl assms that | assumption)+ | |
| 253 | done | |
| 254 | qed | |
| 255 | ||
| 256 | lemma retraction_imp_quotient_map: | |
| 257 | "retraction S T r | |
| 258 | \<Longrightarrow> U \<subseteq> T | |
| 259 | \<Longrightarrow> (openin (subtopology euclidean S) (S \<inter> r -` U) \<longleftrightarrow> | |
| 260 | openin (subtopology euclidean T) U)" | |
| 261 | apply (clarsimp simp add: retraction) | |
| 262 | apply (rule continuous_right_inverse_imp_quotient_map [where g=r]) | |
| 263 | apply (auto simp: elim: continuous_on_subset) | |
| 264 | done | |
| 265 | ||
| 266 | lemma retract_of_locally_compact: | |
| 267 |     fixes S :: "'a :: {heine_borel,real_normed_vector} set"
 | |
| 268 | shows "\<lbrakk> locally compact S; T retract_of S\<rbrakk> \<Longrightarrow> locally compact T" | |
| 269 | by (metis locally_compact_closedin closedin_retract) | |
| 270 | ||
| 271 | lemma retract_of_Times: | |
| 272 | "\<lbrakk>S retract_of s'; T retract_of t'\<rbrakk> \<Longrightarrow> (S \<times> T) retract_of (s' \<times> t')" | |
| 273 | apply (simp add: retract_of_def retraction_def Sigma_mono, clarify) | |
| 274 | apply (rename_tac f g) | |
| 275 | apply (rule_tac x="\<lambda>z. ((f \<circ> fst) z, (g \<circ> snd) z)" in exI) | |
| 276 | apply (rule conjI continuous_intros | erule continuous_on_subset | force)+ | |
| 277 | done | |
| 278 | ||
| 279 | lemma homotopic_into_retract: | |
| 280 | "\<lbrakk>f ` S \<subseteq> T; g ` S \<subseteq> T; T retract_of U; homotopic_with (\<lambda>x. True) S U f g\<rbrakk> | |
| 281 | \<Longrightarrow> homotopic_with (\<lambda>x. True) S T f g" | |
| 282 | apply (subst (asm) homotopic_with_def) | |
| 283 | apply (simp add: homotopic_with retract_of_def retraction_def, clarify) | |
| 284 | apply (rule_tac x="r \<circ> h" in exI) | |
| 285 | apply (rule conjI continuous_intros | erule continuous_on_subset | force simp: image_subset_iff)+ | |
| 286 | done | |
| 287 | ||
| 288 | lemma retract_of_locally_connected: | |
| 289 | assumes "locally connected T" "S retract_of T" | |
| 290 | shows "locally connected S" | |
| 291 | using assms | |
| 292 | by (auto simp: retract_of_def retraction intro!: retraction_imp_quotient_map elim!: locally_connected_quotient_image) | |
| 293 | ||
| 294 | lemma retract_of_locally_path_connected: | |
| 295 | assumes "locally path_connected T" "S retract_of T" | |
| 296 | shows "locally path_connected S" | |
| 297 | using assms | |
| 298 | by (auto simp: retract_of_def retraction intro!: retraction_imp_quotient_map elim!: locally_path_connected_quotient_image) | |
| 299 | ||
| 300 | text \<open>A few simple lemmas about deformation retracts\<close> | |
| 301 | ||
| 302 | lemma deformation_retract_imp_homotopy_eqv: | |
| 303 | fixes S :: "'a::euclidean_space set" | |
| 304 | assumes "homotopic_with (\<lambda>x. True) S S id r" and r: "retraction S T r" | |
| 305 | shows "S homotopy_eqv T" | |
| 306 | proof - | |
| 307 | have "homotopic_with (\<lambda>x. True) S S (id \<circ> r) id" | |
| 308 | by (simp add: assms(1) homotopic_with_symD) | |
| 309 | moreover have "homotopic_with (\<lambda>x. True) T T (r \<circ> id) id" | |
| 310 | using r unfolding retraction_def | |
| 311 | by (metis (no_types, lifting) comp_id continuous_on_id' homotopic_with_equal homotopic_with_symD id_def image_id order_refl) | |
| 312 | ultimately | |
| 313 | show ?thesis | |
| 314 | unfolding homotopy_eqv_def | |
| 315 | by (metis continuous_on_id' id_def image_id r retraction_def) | |
| 316 | qed | |
| 317 | ||
| 318 | lemma deformation_retract: | |
| 319 | fixes S :: "'a::euclidean_space set" | |
| 320 | shows "(\<exists>r. homotopic_with (\<lambda>x. True) S S id r \<and> retraction S T r) \<longleftrightarrow> | |
| 321 | T retract_of S \<and> (\<exists>f. homotopic_with (\<lambda>x. True) S S id f \<and> f ` S \<subseteq> T)" | |
| 322 | (is "?lhs = ?rhs") | |
| 323 | proof | |
| 324 | assume ?lhs | |
| 325 | then show ?rhs | |
| 326 | by (auto simp: retract_of_def retraction_def) | |
| 327 | next | |
| 328 | assume ?rhs | |
| 329 | then show ?lhs | |
| 330 | apply (clarsimp simp add: retract_of_def retraction_def) | |
| 331 | apply (rule_tac x=r in exI, simp) | |
| 332 | apply (rule homotopic_with_trans, assumption) | |
| 333 | apply (rule_tac f = "r \<circ> f" and g="r \<circ> id" in homotopic_with_eq) | |
| 334 | apply (rule_tac Y=S in homotopic_compose_continuous_left) | |
| 335 | apply (auto simp: homotopic_with_sym) | |
| 336 | done | |
| 337 | qed | |
| 338 | ||
| 339 | lemma deformation_retract_of_contractible_sing: | |
| 340 | fixes S :: "'a::euclidean_space set" | |
| 341 | assumes "contractible S" "a \<in> S" | |
| 342 |   obtains r where "homotopic_with (\<lambda>x. True) S S id r" "retraction S {a} r"
 | |
| 343 | proof - | |
| 344 |   have "{a} retract_of S"
 | |
| 345 | by (simp add: \<open>a \<in> S\<close>) | |
| 346 | moreover have "homotopic_with (\<lambda>x. True) S S id (\<lambda>x. a)" | |
| 347 | using assms | |
| 348 | by (auto simp: contractible_def continuous_on_const continuous_on_id homotopic_into_contractible image_subset_iff) | |
| 349 |   moreover have "(\<lambda>x. a) ` S \<subseteq> {a}"
 | |
| 350 | by (simp add: image_subsetI) | |
| 351 | ultimately show ?thesis | |
| 352 | using that deformation_retract by metis | |
| 353 | qed | |
| 354 | ||
| 355 | ||
| 356 | lemma continuous_on_compact_surface_projection_aux: | |
| 357 | fixes S :: "'a::t2_space set" | |
| 358 | assumes "compact S" "S \<subseteq> T" "image q T \<subseteq> S" | |
| 359 | and contp: "continuous_on T p" | |
| 360 | and "\<And>x. x \<in> S \<Longrightarrow> q x = x" | |
| 361 | and [simp]: "\<And>x. x \<in> T \<Longrightarrow> q(p x) = q x" | |
| 362 | and "\<And>x. x \<in> T \<Longrightarrow> p(q x) = p x" | |
| 363 | shows "continuous_on T q" | |
| 364 | proof - | |
| 365 | have *: "image p T = image p S" | |
| 366 | using assms by auto (metis imageI subset_iff) | |
| 367 | have contp': "continuous_on S p" | |
| 368 | by (rule continuous_on_subset [OF contp \<open>S \<subseteq> T\<close>]) | |
| 369 | have "continuous_on (p ` T) q" | |
| 370 | by (simp add: "*" assms(1) assms(2) assms(5) continuous_on_inv contp' rev_subsetD) | |
| 371 | then have "continuous_on T (q \<circ> p)" | |
| 372 | by (rule continuous_on_compose [OF contp]) | |
| 373 | then show ?thesis | |
| 374 | by (rule continuous_on_eq [of _ "q \<circ> p"]) (simp add: o_def) | |
| 375 | qed | |
| 376 | ||
| 377 | lemma continuous_on_compact_surface_projection: | |
| 378 | fixes S :: "'a::real_normed_vector set" | |
| 379 | assumes "compact S" | |
| 380 |       and S: "S \<subseteq> V - {0}" and "cone V"
 | |
| 381 |       and iff: "\<And>x k. x \<in> V - {0} \<Longrightarrow> 0 < k \<and> (k *\<^sub>R x) \<in> S \<longleftrightarrow> d x = k"
 | |
| 382 |   shows "continuous_on (V - {0}) (\<lambda>x. d x *\<^sub>R x)"
 | |
| 383 | proof (rule continuous_on_compact_surface_projection_aux [OF \<open>compact S\<close> S]) | |
| 384 |   show "(\<lambda>x. d x *\<^sub>R x) ` (V - {0}) \<subseteq> S"
 | |
| 385 | using iff by auto | |
| 386 |   show "continuous_on (V - {0}) (\<lambda>x. inverse(norm x) *\<^sub>R x)"
 | |
| 387 | by (intro continuous_intros) force | |
| 388 | show "\<And>x. x \<in> S \<Longrightarrow> d x *\<^sub>R x = x" | |
| 389 | by (metis S zero_less_one local.iff scaleR_one subset_eq) | |
| 390 |   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
 | |
| 391 | using iff [of "inverse(norm x) *\<^sub>R x" "norm x * d x", symmetric] iff that \<open>cone V\<close> | |
| 392 | by (simp add: field_simps cone_def zero_less_mult_iff) | |
| 393 |   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
 | |
| 394 | proof - | |
| 395 | have "0 < d x" | |
| 396 | using local.iff that by blast | |
| 397 | then show ?thesis | |
| 398 | by simp | |
| 399 | qed | |
| 400 | qed | |
| 401 | ||
| 402 | subsection \<open>Absolute retracts, absolute neighbourhood retracts (ANR) and Euclidean neighbourhood retracts (ENR)\<close> | |
| 403 | ||
| 404 | text \<open>Absolute retracts (AR), absolute neighbourhood retracts (ANR) and also Euclidean neighbourhood | |
| 405 | retracts (ENR). We define AR and ANR by specializing the standard definitions for a set to embedding | |
| 406 | in spaces of higher dimension. | |
| 407 | ||
| 408 | John Harrison writes: "This turns out to be sufficient (since any set in $\mathbb{R}^n$ can be
 | |
| 409 | embedded as a closed subset of a convex subset of $\mathbb{R}^{n+1}$) to derive the usual
 | |
| 410 | definitions, but we need to split them into two implications because of the lack of type | |
| 411 | quantifiers. Then ENR turns out to be equivalent to ANR plus local compactness."\<close> | |
| 412 | ||
| 413 | definition AR :: "'a::topological_space set => bool" | |
| 414 | where | |
| 415 |    "AR S \<equiv> \<forall>U. \<forall>S'::('a * real) set. S homeomorphic S' \<and> closedin (subtopology euclidean U) S'
 | |
| 416 | \<longrightarrow> S' retract_of U" | |
| 417 | ||
| 418 | definition ANR :: "'a::topological_space set => bool" | |
| 419 | where | |
| 420 |    "ANR S \<equiv> \<forall>U. \<forall>S'::('a * real) set. S homeomorphic S' \<and> closedin (subtopology euclidean U) S'
 | |
| 421 | \<longrightarrow> (\<exists>T. openin (subtopology euclidean U) T \<and> | |
| 422 | S' retract_of T)" | |
| 423 | ||
| 424 | definition ENR :: "'a::topological_space set => bool" | |
| 425 | where "ENR S \<equiv> \<exists>U. open U \<and> S retract_of U" | |
| 426 | ||
| 427 | text \<open>First, show that we do indeed get the "usual" properties of ARs and ANRs.\<close> | |
| 428 | ||
| 429 | lemma AR_imp_absolute_extensor: | |
| 430 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 431 | assumes "AR S" and contf: "continuous_on T f" and "f ` T \<subseteq> S" | |
| 432 | and cloUT: "closedin (subtopology euclidean U) T" | |
| 433 | obtains g where "continuous_on U g" "g ` U \<subseteq> S" "\<And>x. x \<in> T \<Longrightarrow> g x = f x" | |
| 434 | proof - | |
| 435 |   have "aff_dim S < int (DIM('b \<times> real))"
 | |
| 436 | using aff_dim_le_DIM [of S] by simp | |
| 437 |   then obtain C and S' :: "('b * real) set"
 | |
| 438 |           where C: "convex C" "C \<noteq> {}"
 | |
| 439 | and cloCS: "closedin (subtopology euclidean C) S'" | |
| 440 | and hom: "S homeomorphic S'" | |
| 441 | by (metis that homeomorphic_closedin_convex) | |
| 442 | then have "S' retract_of C" | |
| 443 | using \<open>AR S\<close> by (simp add: AR_def) | |
| 444 | then obtain r where "S' \<subseteq> C" and contr: "continuous_on C r" | |
| 445 | and "r ` C \<subseteq> S'" and rid: "\<And>x. x\<in>S' \<Longrightarrow> r x = x" | |
| 446 | by (auto simp: retraction_def retract_of_def) | |
| 447 | obtain g h where "homeomorphism S S' g h" | |
| 448 | using hom by (force simp: homeomorphic_def) | |
| 449 | then have "continuous_on (f ` T) g" | |
| 450 | by (meson \<open>f ` T \<subseteq> S\<close> continuous_on_subset homeomorphism_def) | |
| 451 | then have contgf: "continuous_on T (g \<circ> f)" | |
| 452 | by (metis continuous_on_compose contf) | |
| 453 | have gfTC: "(g \<circ> f) ` T \<subseteq> C" | |
| 454 | proof - | |
| 455 | have "g ` S = S'" | |
| 456 | by (metis (no_types) \<open>homeomorphism S S' g h\<close> homeomorphism_def) | |
| 457 | with \<open>S' \<subseteq> C\<close> \<open>f ` T \<subseteq> S\<close> show ?thesis by force | |
| 458 | qed | |
| 459 | obtain f' where f': "continuous_on U f'" "f' ` U \<subseteq> C" | |
| 460 | "\<And>x. x \<in> T \<Longrightarrow> f' x = (g \<circ> f) x" | |
| 461 | by (metis Dugundji [OF C cloUT contgf gfTC]) | |
| 462 | show ?thesis | |
| 463 | proof (rule_tac g = "h \<circ> r \<circ> f'" in that) | |
| 464 | show "continuous_on U (h \<circ> r \<circ> f')" | |
| 465 | apply (intro continuous_on_compose f') | |
| 466 | using continuous_on_subset contr f' apply blast | |
| 467 | 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) | |
| 468 | show "(h \<circ> r \<circ> f') ` U \<subseteq> S" | |
| 469 | using \<open>homeomorphism S S' g h\<close> \<open>r ` C \<subseteq> S'\<close> \<open>f' ` U \<subseteq> C\<close> | |
| 470 | by (fastforce simp: homeomorphism_def) | |
| 471 | show "\<And>x. x \<in> T \<Longrightarrow> (h \<circ> r \<circ> f') x = f x" | |
| 472 | using \<open>homeomorphism S S' g h\<close> \<open>f ` T \<subseteq> S\<close> f' | |
| 473 | by (auto simp: rid homeomorphism_def) | |
| 474 | qed | |
| 475 | qed | |
| 476 | ||
| 477 | lemma AR_imp_absolute_retract: | |
| 478 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | |
| 479 | assumes "AR S" "S homeomorphic S'" | |
| 480 | and clo: "closedin (subtopology euclidean U) S'" | |
| 481 | shows "S' retract_of U" | |
| 482 | proof - | |
| 483 | obtain g h where hom: "homeomorphism S S' g h" | |
| 484 | using assms by (force simp: homeomorphic_def) | |
| 485 | have h: "continuous_on S' h" " h ` S' \<subseteq> S" | |
| 486 | using hom homeomorphism_def apply blast | |
| 487 | apply (metis hom equalityE homeomorphism_def) | |
| 488 | done | |
| 489 | obtain h' where h': "continuous_on U h'" "h' ` U \<subseteq> S" | |
| 490 | and h'h: "\<And>x. x \<in> S' \<Longrightarrow> h' x = h x" | |
| 491 | by (blast intro: AR_imp_absolute_extensor [OF \<open>AR S\<close> h clo]) | |
| 492 | have [simp]: "S' \<subseteq> U" using clo closedin_limpt by blast | |
| 493 | show ?thesis | |
| 494 | proof (simp add: retraction_def retract_of_def, intro exI conjI) | |
| 495 | show "continuous_on U (g \<circ> h')" | |
| 496 | apply (intro continuous_on_compose h') | |
| 497 | apply (meson hom continuous_on_subset h' homeomorphism_cont1) | |
| 498 | done | |
| 499 | show "(g \<circ> h') ` U \<subseteq> S'" | |
| 500 | using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) | |
| 501 | show "\<forall>x\<in>S'. (g \<circ> h') x = x" | |
| 502 | by clarsimp (metis h'h hom homeomorphism_def) | |
| 503 | qed | |
| 504 | qed | |
| 505 | ||
| 506 | lemma AR_imp_absolute_retract_UNIV: | |
| 507 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | |
| 508 | assumes "AR S" and hom: "S homeomorphic S'" | |
| 509 | and clo: "closed S'" | |
| 510 | shows "S' retract_of UNIV" | |
| 511 | apply (rule AR_imp_absolute_retract [OF \<open>AR S\<close> hom]) | |
| 512 | using clo closed_closedin by auto | |
| 513 | ||
| 514 | lemma absolute_extensor_imp_AR: | |
| 515 | fixes S :: "'a::euclidean_space set" | |
| 516 | assumes "\<And>f :: 'a * real \<Rightarrow> 'a. | |
| 517 | \<And>U T. \<lbrakk>continuous_on T f; f ` T \<subseteq> S; | |
| 518 | closedin (subtopology euclidean U) T\<rbrakk> | |
| 519 | \<Longrightarrow> \<exists>g. continuous_on U g \<and> g ` U \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)" | |
| 520 | shows "AR S" | |
| 521 | proof (clarsimp simp: AR_def) | |
| 522 |   fix U and T :: "('a * real) set"
 | |
| 523 | assume "S homeomorphic T" and clo: "closedin (subtopology euclidean U) T" | |
| 524 | then obtain g h where hom: "homeomorphism S T g h" | |
| 525 | by (force simp: homeomorphic_def) | |
| 526 | have h: "continuous_on T h" " h ` T \<subseteq> S" | |
| 527 | using hom homeomorphism_def apply blast | |
| 528 | apply (metis hom equalityE homeomorphism_def) | |
| 529 | done | |
| 530 | obtain h' where h': "continuous_on U h'" "h' ` U \<subseteq> S" | |
| 531 | and h'h: "\<forall>x\<in>T. h' x = h x" | |
| 532 | using assms [OF h clo] by blast | |
| 533 | have [simp]: "T \<subseteq> U" | |
| 534 | using clo closedin_imp_subset by auto | |
| 535 | show "T retract_of U" | |
| 536 | proof (simp add: retraction_def retract_of_def, intro exI conjI) | |
| 537 | show "continuous_on U (g \<circ> h')" | |
| 538 | apply (intro continuous_on_compose h') | |
| 539 | apply (meson hom continuous_on_subset h' homeomorphism_cont1) | |
| 540 | done | |
| 541 | show "(g \<circ> h') ` U \<subseteq> T" | |
| 542 | using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) | |
| 543 | show "\<forall>x\<in>T. (g \<circ> h') x = x" | |
| 544 | by clarsimp (metis h'h hom homeomorphism_def) | |
| 545 | qed | |
| 546 | qed | |
| 547 | ||
| 548 | lemma AR_eq_absolute_extensor: | |
| 549 | fixes S :: "'a::euclidean_space set" | |
| 550 | shows "AR S \<longleftrightarrow> | |
| 551 | (\<forall>f :: 'a * real \<Rightarrow> 'a. | |
| 552 | \<forall>U T. continuous_on T f \<longrightarrow> f ` T \<subseteq> S \<longrightarrow> | |
| 553 | closedin (subtopology euclidean U) T \<longrightarrow> | |
| 554 | (\<exists>g. continuous_on U g \<and> g ` U \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)))" | |
| 555 | apply (rule iffI) | |
| 556 | apply (metis AR_imp_absolute_extensor) | |
| 557 | apply (simp add: absolute_extensor_imp_AR) | |
| 558 | done | |
| 559 | ||
| 560 | lemma AR_imp_retract: | |
| 561 | fixes S :: "'a::euclidean_space set" | |
| 562 | assumes "AR S \<and> closedin (subtopology euclidean U) S" | |
| 563 | shows "S retract_of U" | |
| 564 | using AR_imp_absolute_retract assms homeomorphic_refl by blast | |
| 565 | ||
| 566 | lemma AR_homeomorphic_AR: | |
| 567 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | |
| 568 | assumes "AR T" "S homeomorphic T" | |
| 569 | shows "AR S" | |
| 570 | unfolding AR_def | |
| 571 | by (metis assms AR_imp_absolute_retract homeomorphic_trans [of _ S] homeomorphic_sym) | |
| 572 | ||
| 573 | lemma homeomorphic_AR_iff_AR: | |
| 574 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | |
| 575 | shows "S homeomorphic T \<Longrightarrow> AR S \<longleftrightarrow> AR T" | |
| 576 | by (metis AR_homeomorphic_AR homeomorphic_sym) | |
| 577 | ||
| 578 | ||
| 579 | lemma ANR_imp_absolute_neighbourhood_extensor: | |
| 580 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 581 | assumes "ANR S" and contf: "continuous_on T f" and "f ` T \<subseteq> S" | |
| 582 | and cloUT: "closedin (subtopology euclidean U) T" | |
| 583 | obtains V g where "T \<subseteq> V" "openin (subtopology euclidean U) V" | |
| 584 | "continuous_on V g" | |
| 585 | "g ` V \<subseteq> S" "\<And>x. x \<in> T \<Longrightarrow> g x = f x" | |
| 586 | proof - | |
| 587 |   have "aff_dim S < int (DIM('b \<times> real))"
 | |
| 588 | using aff_dim_le_DIM [of S] by simp | |
| 589 |   then obtain C and S' :: "('b * real) set"
 | |
| 590 |           where C: "convex C" "C \<noteq> {}"
 | |
| 591 | and cloCS: "closedin (subtopology euclidean C) S'" | |
| 592 | and hom: "S homeomorphic S'" | |
| 593 | by (metis that homeomorphic_closedin_convex) | |
| 594 | then obtain D where opD: "openin (subtopology euclidean C) D" and "S' retract_of D" | |
| 595 | using \<open>ANR S\<close> by (auto simp: ANR_def) | |
| 596 | then obtain r where "S' \<subseteq> D" and contr: "continuous_on D r" | |
| 597 | and "r ` D \<subseteq> S'" and rid: "\<And>x. x \<in> S' \<Longrightarrow> r x = x" | |
| 598 | by (auto simp: retraction_def retract_of_def) | |
| 599 | obtain g h where homgh: "homeomorphism S S' g h" | |
| 600 | using hom by (force simp: homeomorphic_def) | |
| 601 | have "continuous_on (f ` T) g" | |
| 602 | by (meson \<open>f ` T \<subseteq> S\<close> continuous_on_subset homeomorphism_def homgh) | |
| 603 | then have contgf: "continuous_on T (g \<circ> f)" | |
| 604 | by (intro continuous_on_compose contf) | |
| 605 | have gfTC: "(g \<circ> f) ` T \<subseteq> C" | |
| 606 | proof - | |
| 607 | have "g ` S = S'" | |
| 608 | by (metis (no_types) homeomorphism_def homgh) | |
| 609 | then show ?thesis | |
| 610 | by (metis (no_types) assms(3) cloCS closedin_def image_comp image_mono order.trans topspace_euclidean_subtopology) | |
| 611 | qed | |
| 612 | obtain f' where contf': "continuous_on U f'" | |
| 613 | and "f' ` U \<subseteq> C" | |
| 614 | and eq: "\<And>x. x \<in> T \<Longrightarrow> f' x = (g \<circ> f) x" | |
| 615 | by (metis Dugundji [OF C cloUT contgf gfTC]) | |
| 616 | show ?thesis | |
| 617 | proof (rule_tac V = "U \<inter> f' -` D" and g = "h \<circ> r \<circ> f'" in that) | |
| 618 | show "T \<subseteq> U \<inter> f' -` D" | |
| 619 | using cloUT closedin_imp_subset \<open>S' \<subseteq> D\<close> \<open>f ` T \<subseteq> S\<close> eq homeomorphism_image1 homgh | |
| 620 | by fastforce | |
| 621 | show ope: "openin (subtopology euclidean U) (U \<inter> f' -` D)" | |
| 622 | using \<open>f' ` U \<subseteq> C\<close> by (auto simp: opD contf' continuous_openin_preimage) | |
| 623 | have conth: "continuous_on (r ` f' ` (U \<inter> f' -` D)) h" | |
| 624 | apply (rule continuous_on_subset [of S']) | |
| 625 | using homeomorphism_def homgh apply blast | |
| 626 | using \<open>r ` D \<subseteq> S'\<close> by blast | |
| 627 | show "continuous_on (U \<inter> f' -` D) (h \<circ> r \<circ> f')" | |
| 628 | apply (intro continuous_on_compose conth | |
| 629 | continuous_on_subset [OF contr] continuous_on_subset [OF contf'], auto) | |
| 630 | done | |
| 631 | show "(h \<circ> r \<circ> f') ` (U \<inter> f' -` D) \<subseteq> S" | |
| 632 | using \<open>homeomorphism S S' g h\<close> \<open>f' ` U \<subseteq> C\<close> \<open>r ` D \<subseteq> S'\<close> | |
| 633 | by (auto simp: homeomorphism_def) | |
| 634 | show "\<And>x. x \<in> T \<Longrightarrow> (h \<circ> r \<circ> f') x = f x" | |
| 635 | using \<open>homeomorphism S S' g h\<close> \<open>f ` T \<subseteq> S\<close> eq | |
| 636 | by (auto simp: rid homeomorphism_def) | |
| 637 | qed | |
| 638 | qed | |
| 639 | ||
| 640 | ||
| 641 | corollary ANR_imp_absolute_neighbourhood_retract: | |
| 642 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | |
| 643 | assumes "ANR S" "S homeomorphic S'" | |
| 644 | and clo: "closedin (subtopology euclidean U) S'" | |
| 645 | obtains V where "openin (subtopology euclidean U) V" "S' retract_of V" | |
| 646 | proof - | |
| 647 | obtain g h where hom: "homeomorphism S S' g h" | |
| 648 | using assms by (force simp: homeomorphic_def) | |
| 649 | have h: "continuous_on S' h" " h ` S' \<subseteq> S" | |
| 650 | using hom homeomorphism_def apply blast | |
| 651 | apply (metis hom equalityE homeomorphism_def) | |
| 652 | done | |
| 653 | from ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> h clo] | |
| 654 | obtain V h' where "S' \<subseteq> V" and opUV: "openin (subtopology euclidean U) V" | |
| 655 | and h': "continuous_on V h'" "h' ` V \<subseteq> S" | |
| 656 | and h'h:"\<And>x. x \<in> S' \<Longrightarrow> h' x = h x" | |
| 657 | by (blast intro: ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> h clo]) | |
| 658 | have "S' retract_of V" | |
| 659 | proof (simp add: retraction_def retract_of_def, intro exI conjI \<open>S' \<subseteq> V\<close>) | |
| 660 | show "continuous_on V (g \<circ> h')" | |
| 661 | apply (intro continuous_on_compose h') | |
| 662 | apply (meson hom continuous_on_subset h' homeomorphism_cont1) | |
| 663 | done | |
| 664 | show "(g \<circ> h') ` V \<subseteq> S'" | |
| 665 | using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) | |
| 666 | show "\<forall>x\<in>S'. (g \<circ> h') x = x" | |
| 667 | by clarsimp (metis h'h hom homeomorphism_def) | |
| 668 | qed | |
| 669 | then show ?thesis | |
| 670 | by (rule that [OF opUV]) | |
| 671 | qed | |
| 672 | ||
| 673 | corollary ANR_imp_absolute_neighbourhood_retract_UNIV: | |
| 674 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | |
| 675 | assumes "ANR S" and hom: "S homeomorphic S'" and clo: "closed S'" | |
| 676 | obtains V where "open V" "S' retract_of V" | |
| 677 | using ANR_imp_absolute_neighbourhood_retract [OF \<open>ANR S\<close> hom] | |
| 678 | by (metis clo closed_closedin open_openin subtopology_UNIV) | |
| 679 | ||
| 680 | corollary neighbourhood_extension_into_ANR: | |
| 681 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 682 | assumes contf: "continuous_on S f" and fim: "f ` S \<subseteq> T" and "ANR T" "closed S" | |
| 683 | obtains V g where "S \<subseteq> V" "open V" "continuous_on V g" | |
| 684 | "g ` V \<subseteq> T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x" | |
| 685 | using ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> contf fim] | |
| 686 | by (metis \<open>closed S\<close> closed_closedin open_openin subtopology_UNIV) | |
| 687 | ||
| 688 | lemma absolute_neighbourhood_extensor_imp_ANR: | |
| 689 | fixes S :: "'a::euclidean_space set" | |
| 690 | assumes "\<And>f :: 'a * real \<Rightarrow> 'a. | |
| 691 | \<And>U T. \<lbrakk>continuous_on T f; f ` T \<subseteq> S; | |
| 692 | closedin (subtopology euclidean U) T\<rbrakk> | |
| 693 | \<Longrightarrow> \<exists>V g. T \<subseteq> V \<and> openin (subtopology euclidean U) V \<and> | |
| 694 | continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)" | |
| 695 | shows "ANR S" | |
| 696 | proof (clarsimp simp: ANR_def) | |
| 697 |   fix U and T :: "('a * real) set"
 | |
| 698 | assume "S homeomorphic T" and clo: "closedin (subtopology euclidean U) T" | |
| 699 | then obtain g h where hom: "homeomorphism S T g h" | |
| 700 | by (force simp: homeomorphic_def) | |
| 701 | have h: "continuous_on T h" " h ` T \<subseteq> S" | |
| 702 | using hom homeomorphism_def apply blast | |
| 703 | apply (metis hom equalityE homeomorphism_def) | |
| 704 | done | |
| 705 | obtain V h' where "T \<subseteq> V" and opV: "openin (subtopology euclidean U) V" | |
| 706 | and h': "continuous_on V h'" "h' ` V \<subseteq> S" | |
| 707 | and h'h: "\<forall>x\<in>T. h' x = h x" | |
| 708 | using assms [OF h clo] by blast | |
| 709 | have [simp]: "T \<subseteq> U" | |
| 710 | using clo closedin_imp_subset by auto | |
| 711 | have "T retract_of V" | |
| 712 | proof (simp add: retraction_def retract_of_def, intro exI conjI \<open>T \<subseteq> V\<close>) | |
| 713 | show "continuous_on V (g \<circ> h')" | |
| 714 | apply (intro continuous_on_compose h') | |
| 715 | apply (meson hom continuous_on_subset h' homeomorphism_cont1) | |
| 716 | done | |
| 717 | show "(g \<circ> h') ` V \<subseteq> T" | |
| 718 | using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) | |
| 719 | show "\<forall>x\<in>T. (g \<circ> h') x = x" | |
| 720 | by clarsimp (metis h'h hom homeomorphism_def) | |
| 721 | qed | |
| 722 | then show "\<exists>V. openin (subtopology euclidean U) V \<and> T retract_of V" | |
| 723 | using opV by blast | |
| 724 | qed | |
| 725 | ||
| 726 | lemma ANR_eq_absolute_neighbourhood_extensor: | |
| 727 | fixes S :: "'a::euclidean_space set" | |
| 728 | shows "ANR S \<longleftrightarrow> | |
| 729 | (\<forall>f :: 'a * real \<Rightarrow> 'a. | |
| 730 | \<forall>U T. continuous_on T f \<longrightarrow> f ` T \<subseteq> S \<longrightarrow> | |
| 731 | closedin (subtopology euclidean U) T \<longrightarrow> | |
| 732 | (\<exists>V g. T \<subseteq> V \<and> openin (subtopology euclidean U) V \<and> | |
| 733 | continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)))" | |
| 734 | apply (rule iffI) | |
| 735 | apply (metis ANR_imp_absolute_neighbourhood_extensor) | |
| 736 | apply (simp add: absolute_neighbourhood_extensor_imp_ANR) | |
| 737 | done | |
| 738 | ||
| 739 | lemma ANR_imp_neighbourhood_retract: | |
| 740 | fixes S :: "'a::euclidean_space set" | |
| 741 | assumes "ANR S" "closedin (subtopology euclidean U) S" | |
| 742 | obtains V where "openin (subtopology euclidean U) V" "S retract_of V" | |
| 743 | using ANR_imp_absolute_neighbourhood_retract assms homeomorphic_refl by blast | |
| 744 | ||
| 745 | lemma ANR_imp_absolute_closed_neighbourhood_retract: | |
| 746 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | |
| 747 | assumes "ANR S" "S homeomorphic S'" and US': "closedin (subtopology euclidean U) S'" | |
| 748 | obtains V W | |
| 749 | where "openin (subtopology euclidean U) V" | |
| 750 | "closedin (subtopology euclidean U) W" | |
| 751 | "S' \<subseteq> V" "V \<subseteq> W" "S' retract_of W" | |
| 752 | proof - | |
| 753 | obtain Z where "openin (subtopology euclidean U) Z" and S'Z: "S' retract_of Z" | |
| 754 | by (blast intro: assms ANR_imp_absolute_neighbourhood_retract) | |
| 755 | then have UUZ: "closedin (subtopology euclidean U) (U - Z)" | |
| 756 | by auto | |
| 757 |   have "S' \<inter> (U - Z) = {}"
 | |
| 758 | using \<open>S' retract_of Z\<close> closedin_retract closedin_subtopology by fastforce | |
| 759 | then obtain V W | |
| 760 | where "openin (subtopology euclidean U) V" | |
| 761 | and "openin (subtopology euclidean U) W" | |
| 762 |         and "S' \<subseteq> V" "U - Z \<subseteq> W" "V \<inter> W = {}"
 | |
| 763 | using separation_normal_local [OF US' UUZ] by auto | |
| 764 | moreover have "S' retract_of U - W" | |
| 765 | apply (rule retract_of_subset [OF S'Z]) | |
| 766 |     using US' \<open>S' \<subseteq> V\<close> \<open>V \<inter> W = {}\<close> closedin_subset apply fastforce
 | |
| 767 | using Diff_subset_conv \<open>U - Z \<subseteq> W\<close> by blast | |
| 768 | ultimately show ?thesis | |
| 769 | apply (rule_tac V=V and W = "U-W" in that) | |
| 770 | using openin_imp_subset apply force+ | |
| 771 | done | |
| 772 | qed | |
| 773 | ||
| 774 | lemma ANR_imp_closed_neighbourhood_retract: | |
| 775 | fixes S :: "'a::euclidean_space set" | |
| 776 | assumes "ANR S" "closedin (subtopology euclidean U) S" | |
| 777 | obtains V W where "openin (subtopology euclidean U) V" | |
| 778 | "closedin (subtopology euclidean U) W" | |
| 779 | "S \<subseteq> V" "V \<subseteq> W" "S retract_of W" | |
| 780 | by (meson ANR_imp_absolute_closed_neighbourhood_retract assms homeomorphic_refl) | |
| 781 | ||
| 782 | lemma ANR_homeomorphic_ANR: | |
| 783 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | |
| 784 | assumes "ANR T" "S homeomorphic T" | |
| 785 | shows "ANR S" | |
| 786 | unfolding ANR_def | |
| 787 | by (metis assms ANR_imp_absolute_neighbourhood_retract homeomorphic_trans [of _ S] homeomorphic_sym) | |
| 788 | ||
| 789 | lemma homeomorphic_ANR_iff_ANR: | |
| 790 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | |
| 791 | shows "S homeomorphic T \<Longrightarrow> ANR S \<longleftrightarrow> ANR T" | |
| 792 | by (metis ANR_homeomorphic_ANR homeomorphic_sym) | |
| 793 | ||
| 794 | subsubsection \<open>Analogous properties of ENRs\<close> | |
| 795 | ||
| 796 | lemma ENR_imp_absolute_neighbourhood_retract: | |
| 797 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | |
| 798 | assumes "ENR S" and hom: "S homeomorphic S'" | |
| 799 | and "S' \<subseteq> U" | |
| 800 | obtains V where "openin (subtopology euclidean U) V" "S' retract_of V" | |
| 801 | proof - | |
| 802 | obtain X where "open X" "S retract_of X" | |
| 803 | using \<open>ENR S\<close> by (auto simp: ENR_def) | |
| 804 | then obtain r where "retraction X S r" | |
| 805 | by (auto simp: retract_of_def) | |
| 806 | have "locally compact S'" | |
| 807 | using retract_of_locally_compact open_imp_locally_compact | |
| 808 | homeomorphic_local_compactness \<open>S retract_of X\<close> \<open>open X\<close> hom by blast | |
| 809 | then obtain W where UW: "openin (subtopology euclidean U) W" | |
| 810 | and WS': "closedin (subtopology euclidean W) S'" | |
| 811 | apply (rule locally_compact_closedin_open) | |
| 812 | apply (rename_tac W) | |
| 813 | apply (rule_tac W = "U \<inter> W" in that, blast) | |
| 814 | by (simp add: \<open>S' \<subseteq> U\<close> closedin_limpt) | |
| 815 | obtain f g where hom: "homeomorphism S S' f g" | |
| 816 | using assms by (force simp: homeomorphic_def) | |
| 817 | have contg: "continuous_on S' g" | |
| 818 | using hom homeomorphism_def by blast | |
| 819 | moreover have "g ` S' \<subseteq> S" by (metis hom equalityE homeomorphism_def) | |
| 820 | ultimately obtain h where conth: "continuous_on W h" and hg: "\<And>x. x \<in> S' \<Longrightarrow> h x = g x" | |
| 821 | using Tietze_unbounded [of S' g W] WS' by blast | |
| 822 | have "W \<subseteq> U" using UW openin_open by auto | |
| 823 | have "S' \<subseteq> W" using WS' closedin_closed by auto | |
| 824 | have him: "\<And>x. x \<in> S' \<Longrightarrow> h x \<in> X" | |
| 825 | 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) | |
| 826 | have "S' retract_of (W \<inter> h -` X)" | |
| 827 | proof (simp add: retraction_def retract_of_def, intro exI conjI) | |
| 828 | show "S' \<subseteq> W" "S' \<subseteq> h -` X" | |
| 829 | using him WS' closedin_imp_subset by blast+ | |
| 830 | show "continuous_on (W \<inter> h -` X) (f \<circ> r \<circ> h)" | |
| 831 | proof (intro continuous_on_compose) | |
| 832 | show "continuous_on (W \<inter> h -` X) h" | |
| 833 | by (meson conth continuous_on_subset inf_le1) | |
| 834 | show "continuous_on (h ` (W \<inter> h -` X)) r" | |
| 835 | proof - | |
| 836 | have "h ` (W \<inter> h -` X) \<subseteq> X" | |
| 837 | by blast | |
| 838 | then show "continuous_on (h ` (W \<inter> h -` X)) r" | |
| 839 | by (meson \<open>retraction X S r\<close> continuous_on_subset retraction) | |
| 840 | qed | |
| 841 | show "continuous_on (r ` h ` (W \<inter> h -` X)) f" | |
| 842 | apply (rule continuous_on_subset [of S]) | |
| 843 | using hom homeomorphism_def apply blast | |
| 844 | apply clarify | |
| 845 | apply (meson \<open>retraction X S r\<close> subsetD imageI retraction_def) | |
| 846 | done | |
| 847 | qed | |
| 848 | show "(f \<circ> r \<circ> h) ` (W \<inter> h -` X) \<subseteq> S'" | |
| 849 | using \<open>retraction X S r\<close> hom | |
| 850 | by (auto simp: retraction_def homeomorphism_def) | |
| 851 | show "\<forall>x\<in>S'. (f \<circ> r \<circ> h) x = x" | |
| 852 | using \<open>retraction X S r\<close> hom by (auto simp: retraction_def homeomorphism_def hg) | |
| 853 | qed | |
| 854 | then show ?thesis | |
| 855 | apply (rule_tac V = "W \<inter> h -` X" in that) | |
| 856 | apply (rule openin_trans [OF _ UW]) | |
| 857 | using \<open>continuous_on W h\<close> \<open>open X\<close> continuous_openin_preimage_eq apply blast+ | |
| 858 | done | |
| 859 | qed | |
| 860 | ||
| 861 | corollary ENR_imp_absolute_neighbourhood_retract_UNIV: | |
| 862 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | |
| 863 | assumes "ENR S" "S homeomorphic S'" | |
| 864 | obtains T' where "open T'" "S' retract_of T'" | |
| 865 | by (metis ENR_imp_absolute_neighbourhood_retract UNIV_I assms(1) assms(2) open_openin subsetI subtopology_UNIV) | |
| 866 | ||
| 867 | lemma ENR_homeomorphic_ENR: | |
| 868 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | |
| 869 | assumes "ENR T" "S homeomorphic T" | |
| 870 | shows "ENR S" | |
| 871 | unfolding ENR_def | |
| 872 | by (meson ENR_imp_absolute_neighbourhood_retract_UNIV assms homeomorphic_sym) | |
| 873 | ||
| 874 | lemma homeomorphic_ENR_iff_ENR: | |
| 875 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | |
| 876 | assumes "S homeomorphic T" | |
| 877 | shows "ENR S \<longleftrightarrow> ENR T" | |
| 878 | by (meson ENR_homeomorphic_ENR assms homeomorphic_sym) | |
| 879 | ||
| 880 | lemma ENR_translation: | |
| 881 | fixes S :: "'a::euclidean_space set" | |
| 882 | shows "ENR(image (\<lambda>x. a + x) S) \<longleftrightarrow> ENR S" | |
| 883 | by (meson homeomorphic_sym homeomorphic_translation homeomorphic_ENR_iff_ENR) | |
| 884 | ||
| 885 | lemma ENR_linear_image_eq: | |
| 886 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 887 | assumes "linear f" "inj f" | |
| 888 | shows "ENR (image f S) \<longleftrightarrow> ENR S" | |
| 889 | apply (rule homeomorphic_ENR_iff_ENR) | |
| 890 | using assms homeomorphic_sym linear_homeomorphic_image by auto | |
| 891 | ||
| 892 | text \<open>Some relations among the concepts. We also relate AR to being a retract of UNIV, which is | |
| 893 | often a more convenient proxy in the closed case.\<close> | |
| 894 | ||
| 895 | lemma AR_imp_ANR: "AR S \<Longrightarrow> ANR S" | |
| 896 | using ANR_def AR_def by fastforce | |
| 897 | ||
| 898 | lemma ENR_imp_ANR: | |
| 899 | fixes S :: "'a::euclidean_space set" | |
| 900 | shows "ENR S \<Longrightarrow> ANR S" | |
| 901 | apply (simp add: ANR_def) | |
| 902 | by (metis ENR_imp_absolute_neighbourhood_retract closedin_imp_subset) | |
| 903 | ||
| 904 | lemma ENR_ANR: | |
| 905 | fixes S :: "'a::euclidean_space set" | |
| 906 | shows "ENR S \<longleftrightarrow> ANR S \<and> locally compact S" | |
| 907 | proof | |
| 908 | assume "ENR S" | |
| 909 | then have "locally compact S" | |
| 910 | using ENR_def open_imp_locally_compact retract_of_locally_compact by auto | |
| 911 | then show "ANR S \<and> locally compact S" | |
| 912 | using ENR_imp_ANR \<open>ENR S\<close> by blast | |
| 913 | next | |
| 914 | assume "ANR S \<and> locally compact S" | |
| 915 | then have "ANR S" "locally compact S" by auto | |
| 916 |   then obtain T :: "('a * real) set" where "closed T" "S homeomorphic T"
 | |
| 917 | using locally_compact_homeomorphic_closed | |
| 918 | by (metis DIM_prod DIM_real Suc_eq_plus1 lessI) | |
| 919 | then show "ENR S" | |
| 920 | using \<open>ANR S\<close> | |
| 921 | apply (simp add: ANR_def) | |
| 922 | apply (drule_tac x=UNIV in spec) | |
| 923 | apply (drule_tac x=T in spec, clarsimp) | |
| 924 | apply (meson ENR_def ENR_homeomorphic_ENR open_openin) | |
| 925 | done | |
| 926 | qed | |
| 927 | ||
| 928 | ||
| 929 | lemma AR_ANR: | |
| 930 | fixes S :: "'a::euclidean_space set" | |
| 931 |   shows "AR S \<longleftrightarrow> ANR S \<and> contractible S \<and> S \<noteq> {}"
 | |
| 932 | (is "?lhs = ?rhs") | |
| 933 | proof | |
| 934 | assume ?lhs | |
| 935 |   obtain C and S' :: "('a * real) set"
 | |
| 936 |     where "convex C" "C \<noteq> {}" "closedin (subtopology euclidean C) S'" "S homeomorphic S'"
 | |
| 937 | apply (rule homeomorphic_closedin_convex [of S, where 'n = "'a * real"]) | |
| 938 | using aff_dim_le_DIM [of S] by auto | |
| 939 | with \<open>AR S\<close> have "contractible S" | |
| 940 | apply (simp add: AR_def) | |
| 941 | apply (drule_tac x=C in spec) | |
| 942 | apply (drule_tac x="S'" in spec, simp) | |
| 943 | using convex_imp_contractible homeomorphic_contractible_eq retract_of_contractible by fastforce | |
| 944 | with \<open>AR S\<close> show ?rhs | |
| 945 | apply (auto simp: AR_imp_ANR) | |
| 946 | apply (force simp: AR_def) | |
| 947 | done | |
| 948 | next | |
| 949 | assume ?rhs | |
| 950 | then obtain a and h:: "real \<times> 'a \<Rightarrow> 'a" | |
| 951 |       where conth: "continuous_on ({0..1} \<times> S) h"
 | |
| 952 |         and hS: "h ` ({0..1} \<times> S) \<subseteq> S"
 | |
| 953 | and [simp]: "\<And>x. h(0, x) = x" | |
| 954 | and [simp]: "\<And>x. h(1, x) = a" | |
| 955 |         and "ANR S" "S \<noteq> {}"
 | |
| 956 | by (auto simp: contractible_def homotopic_with_def) | |
| 957 | then have "a \<in> S" | |
| 958 | by (metis all_not_in_conv atLeastAtMost_iff image_subset_iff mem_Sigma_iff order_refl zero_le_one) | |
| 959 | have "\<exists>g. continuous_on W g \<and> g ` W \<subseteq> S \<and> (\<forall>x\<in>T. g x = f x)" | |
| 960 | if f: "continuous_on T f" "f ` T \<subseteq> S" | |
| 961 | and WT: "closedin (subtopology euclidean W) T" | |
| 962 | for W T and f :: "'a \<times> real \<Rightarrow> 'a" | |
| 963 | proof - | |
| 964 | obtain U g | |
| 965 | where "T \<subseteq> U" and WU: "openin (subtopology euclidean W) U" | |
| 966 | and contg: "continuous_on U g" | |
| 967 | and "g ` U \<subseteq> S" and gf: "\<And>x. x \<in> T \<Longrightarrow> g x = f x" | |
| 968 | using iffD1 [OF ANR_eq_absolute_neighbourhood_extensor \<open>ANR S\<close>, rule_format, OF f WT] | |
| 969 | by auto | |
| 970 | have WWU: "closedin (subtopology euclidean W) (W - U)" | |
| 971 | using WU closedin_diff by fastforce | |
| 972 |     moreover have "(W - U) \<inter> T = {}"
 | |
| 973 | using \<open>T \<subseteq> U\<close> by auto | |
| 974 | ultimately obtain V V' | |
| 975 | where WV': "openin (subtopology euclidean W) V'" | |
| 976 | and WV: "openin (subtopology euclidean W) V" | |
| 977 |         and "W - U \<subseteq> V'" "T \<subseteq> V" "V' \<inter> V = {}"
 | |
| 978 | using separation_normal_local [of W "W-U" T] WT by blast | |
| 979 |     then have WVT: "T \<inter> (W - V) = {}"
 | |
| 980 | by auto | |
| 981 | have WWV: "closedin (subtopology euclidean W) (W - V)" | |
| 982 | using WV closedin_diff by fastforce | |
| 983 | obtain j :: " 'a \<times> real \<Rightarrow> real" | |
| 984 | where contj: "continuous_on W j" | |
| 985 |         and j:  "\<And>x. x \<in> W \<Longrightarrow> j x \<in> {0..1}"
 | |
| 986 | and j0: "\<And>x. x \<in> W - V \<Longrightarrow> j x = 1" | |
| 987 | and j1: "\<And>x. x \<in> T \<Longrightarrow> j x = 0" | |
| 988 | by (rule Urysohn_local [OF WT WWV WVT, of 0 "1::real"]) (auto simp: in_segment) | |
| 989 | have Weq: "W = (W - V) \<union> (W - V')" | |
| 990 |       using \<open>V' \<inter> V = {}\<close> by force
 | |
| 991 | show ?thesis | |
| 992 | proof (intro conjI exI) | |
| 993 | have *: "continuous_on (W - V') (\<lambda>x. h (j x, g x))" | |
| 994 | apply (rule continuous_on_compose2 [OF conth continuous_on_Pair]) | |
| 995 | apply (rule continuous_on_subset [OF contj Diff_subset]) | |
| 996 | apply (rule continuous_on_subset [OF contg]) | |
| 997 | apply (metis Diff_subset_conv Un_commute \<open>W - U \<subseteq> V'\<close>) | |
| 998 | using j \<open>g ` U \<subseteq> S\<close> \<open>W - U \<subseteq> V'\<close> apply fastforce | |
| 999 | done | |
| 1000 | show "continuous_on W (\<lambda>x. if x \<in> W - V then a else h (j x, g x))" | |
| 1001 | apply (subst Weq) | |
| 1002 | apply (rule continuous_on_cases_local) | |
| 1003 | apply (simp_all add: Weq [symmetric] WWV continuous_on_const *) | |
| 1004 | using WV' closedin_diff apply fastforce | |
| 1005 | apply (auto simp: j0 j1) | |
| 1006 | done | |
| 1007 | next | |
| 1008 | have "h (j (x, y), g (x, y)) \<in> S" if "(x, y) \<in> W" "(x, y) \<in> V" for x y | |
| 1009 | proof - | |
| 1010 |         have "j(x, y) \<in> {0..1}"
 | |
| 1011 | using j that by blast | |
| 1012 | moreover have "g(x, y) \<in> S" | |
| 1013 |           using \<open>V' \<inter> V = {}\<close> \<open>W - U \<subseteq> V'\<close> \<open>g ` U \<subseteq> S\<close> that by fastforce
 | |
| 1014 | ultimately show ?thesis | |
| 1015 | using hS by blast | |
| 1016 | qed | |
| 1017 | with \<open>a \<in> S\<close> \<open>g ` U \<subseteq> S\<close> | |
| 1018 | show "(\<lambda>x. if x \<in> W - V then a else h (j x, g x)) ` W \<subseteq> S" | |
| 1019 | by auto | |
| 1020 | next | |
| 1021 | show "\<forall>x\<in>T. (if x \<in> W - V then a else h (j x, g x)) = f x" | |
| 1022 | using \<open>T \<subseteq> V\<close> by (auto simp: j0 j1 gf) | |
| 1023 | qed | |
| 1024 | qed | |
| 1025 | then show ?lhs | |
| 1026 | by (simp add: AR_eq_absolute_extensor) | |
| 1027 | qed | |
| 1028 | ||
| 1029 | ||
| 1030 | lemma ANR_retract_of_ANR: | |
| 1031 | fixes S :: "'a::euclidean_space set" | |
| 1032 | assumes "ANR T" "S retract_of T" | |
| 1033 | shows "ANR S" | |
| 1034 | using assms | |
| 1035 | apply (simp add: ANR_eq_absolute_neighbourhood_extensor retract_of_def retraction_def) | |
| 1036 | apply (clarsimp elim!: all_forward) | |
| 1037 | apply (erule impCE, metis subset_trans) | |
| 1038 | apply (clarsimp elim!: ex_forward) | |
| 1039 | apply (rule_tac x="r \<circ> g" in exI) | |
| 1040 | by (metis comp_apply continuous_on_compose continuous_on_subset subsetD imageI image_comp image_mono subset_trans) | |
| 1041 | ||
| 1042 | lemma AR_retract_of_AR: | |
| 1043 | fixes S :: "'a::euclidean_space set" | |
| 1044 | shows "\<lbrakk>AR T; S retract_of T\<rbrakk> \<Longrightarrow> AR S" | |
| 1045 | using ANR_retract_of_ANR AR_ANR retract_of_contractible by fastforce | |
| 1046 | ||
| 1047 | lemma ENR_retract_of_ENR: | |
| 1048 | "\<lbrakk>ENR T; S retract_of T\<rbrakk> \<Longrightarrow> ENR S" | |
| 1049 | by (meson ENR_def retract_of_trans) | |
| 1050 | ||
| 1051 | lemma retract_of_UNIV: | |
| 1052 | fixes S :: "'a::euclidean_space set" | |
| 1053 | shows "S retract_of UNIV \<longleftrightarrow> AR S \<and> closed S" | |
| 1054 | 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) | |
| 1055 | ||
| 1056 | lemma compact_AR: | |
| 1057 | fixes S :: "'a::euclidean_space set" | |
| 1058 | shows "compact S \<and> AR S \<longleftrightarrow> compact S \<and> S retract_of UNIV" | |
| 1059 | using compact_imp_closed retract_of_UNIV by blast | |
| 1060 | ||
| 1061 | text \<open>More properties of ARs, ANRs and ENRs\<close> | |
| 1062 | ||
| 1063 | lemma not_AR_empty [simp]: "~ AR({})"
 | |
| 1064 | by (auto simp: AR_def) | |
| 1065 | ||
| 1066 | lemma ENR_empty [simp]: "ENR {}"
 | |
| 1067 | by (simp add: ENR_def) | |
| 1068 | ||
| 1069 | lemma ANR_empty [simp]: "ANR ({} :: 'a::euclidean_space set)"
 | |
| 1070 | by (simp add: ENR_imp_ANR) | |
| 1071 | ||
| 1072 | lemma convex_imp_AR: | |
| 1073 | fixes S :: "'a::euclidean_space set" | |
| 1074 |   shows "\<lbrakk>convex S; S \<noteq> {}\<rbrakk> \<Longrightarrow> AR S"
 | |
| 1075 | apply (rule absolute_extensor_imp_AR) | |
| 1076 | apply (rule Dugundji, assumption+) | |
| 1077 | by blast | |
| 1078 | ||
| 1079 | lemma convex_imp_ANR: | |
| 1080 | fixes S :: "'a::euclidean_space set" | |
| 1081 | shows "convex S \<Longrightarrow> ANR S" | |
| 1082 | using ANR_empty AR_imp_ANR convex_imp_AR by blast | |
| 1083 | ||
| 1084 | lemma ENR_convex_closed: | |
| 1085 | fixes S :: "'a::euclidean_space set" | |
| 1086 | shows "\<lbrakk>closed S; convex S\<rbrakk> \<Longrightarrow> ENR S" | |
| 1087 | using ENR_def ENR_empty convex_imp_AR retract_of_UNIV by blast | |
| 1088 | ||
| 1089 | lemma AR_UNIV [simp]: "AR (UNIV :: 'a::euclidean_space set)" | |
| 1090 | using retract_of_UNIV by auto | |
| 1091 | ||
| 1092 | lemma ANR_UNIV [simp]: "ANR (UNIV :: 'a::euclidean_space set)" | |
| 1093 | by (simp add: AR_imp_ANR) | |
| 1094 | ||
| 1095 | lemma ENR_UNIV [simp]:"ENR UNIV" | |
| 1096 | using ENR_def by blast | |
| 1097 | ||
| 1098 | lemma AR_singleton: | |
| 1099 | fixes a :: "'a::euclidean_space" | |
| 1100 |     shows "AR {a}"
 | |
| 1101 | using retract_of_UNIV by blast | |
| 1102 | ||
| 1103 | lemma ANR_singleton: | |
| 1104 | fixes a :: "'a::euclidean_space" | |
| 1105 |     shows "ANR {a}"
 | |
| 1106 | by (simp add: AR_imp_ANR AR_singleton) | |
| 1107 | ||
| 1108 | lemma ENR_singleton: "ENR {a}"
 | |
| 1109 | using ENR_def by blast | |
| 1110 | ||
| 1111 | text \<open>ARs closed under union\<close> | |
| 1112 | ||
| 1113 | lemma AR_closed_Un_local_aux: | |
| 1114 | fixes U :: "'a::euclidean_space set" | |
| 1115 | assumes "closedin (subtopology euclidean U) S" | |
| 1116 | "closedin (subtopology euclidean U) T" | |
| 1117 | "AR S" "AR T" "AR(S \<inter> T)" | |
| 1118 | shows "(S \<union> T) retract_of U" | |
| 1119 | proof - | |
| 1120 |   have "S \<inter> T \<noteq> {}"
 | |
| 1121 | using assms AR_def by fastforce | |
| 1122 | have "S \<subseteq> U" "T \<subseteq> U" | |
| 1123 | using assms by (auto simp: closedin_imp_subset) | |
| 1124 |   define S' where "S' \<equiv> {x \<in> U. setdist {x} S \<le> setdist {x} T}"
 | |
| 1125 |   define T' where "T' \<equiv> {x \<in> U. setdist {x} T \<le> setdist {x} S}"
 | |
| 1126 |   define W  where "W \<equiv> {x \<in> U. setdist {x} S = setdist {x} T}"
 | |
| 1127 | have US': "closedin (subtopology euclidean U) S'" | |
| 1128 |     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} S - setdist {x} T" "{..0}"]
 | |
| 1129 | by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) | |
| 1130 | have UT': "closedin (subtopology euclidean U) T'" | |
| 1131 |     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} T - setdist {x} S" "{..0}"]
 | |
| 1132 | by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) | |
| 1133 | have "S \<subseteq> S'" | |
| 1134 | using S'_def \<open>S \<subseteq> U\<close> setdist_sing_in_set by fastforce | |
| 1135 | have "T \<subseteq> T'" | |
| 1136 | using T'_def \<open>T \<subseteq> U\<close> setdist_sing_in_set by fastforce | |
| 1137 | have "S \<inter> T \<subseteq> W" "W \<subseteq> U" | |
| 1138 | using \<open>S \<subseteq> U\<close> by (auto simp: W_def setdist_sing_in_set) | |
| 1139 | have "(S \<inter> T) retract_of W" | |
| 1140 | apply (rule AR_imp_absolute_retract [OF \<open>AR(S \<inter> T)\<close>]) | |
| 1141 | apply (simp add: homeomorphic_refl) | |
| 1142 | apply (rule closedin_subset_trans [of U]) | |
| 1143 | apply (simp_all add: assms closedin_Int \<open>S \<inter> T \<subseteq> W\<close> \<open>W \<subseteq> U\<close>) | |
| 1144 | done | |
| 1145 | then obtain r0 | |
| 1146 | where "S \<inter> T \<subseteq> W" and contr0: "continuous_on W r0" | |
| 1147 | and "r0 ` W \<subseteq> S \<inter> T" | |
| 1148 | and r0 [simp]: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> r0 x = x" | |
| 1149 | by (auto simp: retract_of_def retraction_def) | |
| 1150 | have ST: "x \<in> W \<Longrightarrow> x \<in> S \<longleftrightarrow> x \<in> T" for x | |
| 1151 |     using setdist_eq_0_closedin \<open>S \<inter> T \<noteq> {}\<close> assms
 | |
| 1152 | by (force simp: W_def setdist_sing_in_set) | |
| 1153 | have "S' \<inter> T' = W" | |
| 1154 | by (auto simp: S'_def T'_def W_def) | |
| 1155 | then have cloUW: "closedin (subtopology euclidean U) W" | |
| 1156 | using closedin_Int US' UT' by blast | |
| 1157 | define r where "r \<equiv> \<lambda>x. if x \<in> W then r0 x else x" | |
| 1158 | have "r ` (W \<union> S) \<subseteq> S" "r ` (W \<union> T) \<subseteq> T" | |
| 1159 | using \<open>r0 ` W \<subseteq> S \<inter> T\<close> r_def by auto | |
| 1160 | have contr: "continuous_on (W \<union> (S \<union> T)) r" | |
| 1161 | unfolding r_def | |
| 1162 | proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) | |
| 1163 | show "closedin (subtopology euclidean (W \<union> (S \<union> T))) W" | |
| 1164 | 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 | |
| 1165 | show "closedin (subtopology euclidean (W \<union> (S \<union> T))) (S \<union> T)" | |
| 1166 | 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) | |
| 1167 | show "\<And>x. x \<in> W \<and> x \<notin> W \<or> x \<in> S \<union> T \<and> x \<in> W \<Longrightarrow> r0 x = x" | |
| 1168 | by (auto simp: ST) | |
| 1169 | qed | |
| 1170 | have cloUWS: "closedin (subtopology euclidean U) (W \<union> S)" | |
| 1171 | by (simp add: cloUW assms closedin_Un) | |
| 1172 | obtain g where contg: "continuous_on U g" | |
| 1173 | and "g ` U \<subseteq> S" and geqr: "\<And>x. x \<in> W \<union> S \<Longrightarrow> g x = r x" | |
| 1174 | apply (rule AR_imp_absolute_extensor [OF \<open>AR S\<close> _ _ cloUWS]) | |
| 1175 | apply (rule continuous_on_subset [OF contr]) | |
| 1176 | using \<open>r ` (W \<union> S) \<subseteq> S\<close> apply auto | |
| 1177 | done | |
| 1178 | have cloUWT: "closedin (subtopology euclidean U) (W \<union> T)" | |
| 1179 | by (simp add: cloUW assms closedin_Un) | |
| 1180 | obtain h where conth: "continuous_on U h" | |
| 1181 | and "h ` U \<subseteq> T" and heqr: "\<And>x. x \<in> W \<union> T \<Longrightarrow> h x = r x" | |
| 1182 | apply (rule AR_imp_absolute_extensor [OF \<open>AR T\<close> _ _ cloUWT]) | |
| 1183 | apply (rule continuous_on_subset [OF contr]) | |
| 1184 | using \<open>r ` (W \<union> T) \<subseteq> T\<close> apply auto | |
| 1185 | done | |
| 1186 | have "U = S' \<union> T'" | |
| 1187 | by (force simp: S'_def T'_def) | |
| 1188 | then have cont: "continuous_on U (\<lambda>x. if x \<in> S' then g x else h x)" | |
| 1189 | apply (rule ssubst) | |
| 1190 | apply (rule continuous_on_cases_local) | |
| 1191 | using US' UT' \<open>S' \<inter> T' = W\<close> \<open>U = S' \<union> T'\<close> | |
| 1192 | contg conth continuous_on_subset geqr heqr apply auto | |
| 1193 | done | |
| 1194 | have UST: "(\<lambda>x. if x \<in> S' then g x else h x) ` U \<subseteq> S \<union> T" | |
| 1195 | using \<open>g ` U \<subseteq> S\<close> \<open>h ` U \<subseteq> T\<close> by auto | |
| 1196 | show ?thesis | |
| 1197 | apply (simp add: retract_of_def retraction_def \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close>) | |
| 1198 | apply (rule_tac x="\<lambda>x. if x \<in> S' then g x else h x" in exI) | |
| 1199 | apply (intro conjI cont UST) | |
| 1200 | 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) | |
| 1201 | qed | |
| 1202 | ||
| 1203 | ||
| 1204 | lemma AR_closed_Un_local: | |
| 1205 | fixes S :: "'a::euclidean_space set" | |
| 1206 | assumes STS: "closedin (subtopology euclidean (S \<union> T)) S" | |
| 1207 | and STT: "closedin (subtopology euclidean (S \<union> T)) T" | |
| 1208 | and "AR S" "AR T" "AR(S \<inter> T)" | |
| 1209 | shows "AR(S \<union> T)" | |
| 1210 | proof - | |
| 1211 | have "C retract_of U" | |
| 1212 | if hom: "S \<union> T homeomorphic C" and UC: "closedin (subtopology euclidean U) C" | |
| 1213 |        for U and C :: "('a * real) set"
 | |
| 1214 | proof - | |
| 1215 | obtain f g where hom: "homeomorphism (S \<union> T) C f g" | |
| 1216 | using hom by (force simp: homeomorphic_def) | |
| 1217 | have US: "closedin (subtopology euclidean U) (C \<inter> g -` S)" | |
| 1218 | apply (rule closedin_trans [OF _ UC]) | |
| 1219 | apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) | |
| 1220 | using hom homeomorphism_def apply blast | |
| 1221 | apply (metis hom homeomorphism_def set_eq_subset) | |
| 1222 | done | |
| 1223 | have UT: "closedin (subtopology euclidean U) (C \<inter> g -` T)" | |
| 1224 | apply (rule closedin_trans [OF _ UC]) | |
| 1225 | apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) | |
| 1226 | using hom homeomorphism_def apply blast | |
| 1227 | apply (metis hom homeomorphism_def set_eq_subset) | |
| 1228 | done | |
| 1229 | have ARS: "AR (C \<inter> g -` S)" | |
| 1230 | apply (rule AR_homeomorphic_AR [OF \<open>AR S\<close>]) | |
| 1231 | apply (simp add: homeomorphic_def) | |
| 1232 | apply (rule_tac x=g in exI) | |
| 1233 | apply (rule_tac x=f in exI) | |
| 1234 | using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) | |
| 1235 | apply (rule_tac x="f x" in image_eqI, auto) | |
| 1236 | done | |
| 1237 | have ART: "AR (C \<inter> g -` T)" | |
| 1238 | apply (rule AR_homeomorphic_AR [OF \<open>AR T\<close>]) | |
| 1239 | apply (simp add: homeomorphic_def) | |
| 1240 | apply (rule_tac x=g in exI) | |
| 1241 | apply (rule_tac x=f in exI) | |
| 1242 | using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) | |
| 1243 | apply (rule_tac x="f x" in image_eqI, auto) | |
| 1244 | done | |
| 1245 | have ARI: "AR ((C \<inter> g -` S) \<inter> (C \<inter> g -` T))" | |
| 1246 | apply (rule AR_homeomorphic_AR [OF \<open>AR (S \<inter> T)\<close>]) | |
| 1247 | apply (simp add: homeomorphic_def) | |
| 1248 | apply (rule_tac x=g in exI) | |
| 1249 | apply (rule_tac x=f in exI) | |
| 1250 | using hom | |
| 1251 | apply (auto simp: homeomorphism_def elim!: continuous_on_subset) | |
| 1252 | apply (rule_tac x="f x" in image_eqI, auto) | |
| 1253 | done | |
| 1254 | have "C = (C \<inter> g -` S) \<union> (C \<inter> g -` T)" | |
| 1255 | using hom by (auto simp: homeomorphism_def) | |
| 1256 | then show ?thesis | |
| 1257 | by (metis AR_closed_Un_local_aux [OF US UT ARS ART ARI]) | |
| 1258 | qed | |
| 1259 | then show ?thesis | |
| 1260 | by (force simp: AR_def) | |
| 1261 | qed | |
| 1262 | ||
| 1263 | corollary AR_closed_Un: | |
| 1264 | fixes S :: "'a::euclidean_space set" | |
| 1265 | shows "\<lbrakk>closed S; closed T; AR S; AR T; AR (S \<inter> T)\<rbrakk> \<Longrightarrow> AR (S \<union> T)" | |
| 1266 | by (metis AR_closed_Un_local_aux closed_closedin retract_of_UNIV subtopology_UNIV) | |
| 1267 | ||
| 1268 | text \<open>ANRs closed under union\<close> | |
| 1269 | ||
| 1270 | lemma ANR_closed_Un_local_aux: | |
| 1271 | fixes U :: "'a::euclidean_space set" | |
| 1272 | assumes US: "closedin (subtopology euclidean U) S" | |
| 1273 | and UT: "closedin (subtopology euclidean U) T" | |
| 1274 | and "ANR S" "ANR T" "ANR(S \<inter> T)" | |
| 1275 | obtains V where "openin (subtopology euclidean U) V" "(S \<union> T) retract_of V" | |
| 1276 | proof (cases "S = {} \<or> T = {}")
 | |
| 1277 | case True with assms that show ?thesis | |
| 1278 | by (metis ANR_imp_neighbourhood_retract Un_commute inf_bot_right sup_inf_absorb) | |
| 1279 | next | |
| 1280 | case False | |
| 1281 |   then have [simp]: "S \<noteq> {}" "T \<noteq> {}" by auto
 | |
| 1282 | have "S \<subseteq> U" "T \<subseteq> U" | |
| 1283 | using assms by (auto simp: closedin_imp_subset) | |
| 1284 |   define S' where "S' \<equiv> {x \<in> U. setdist {x} S \<le> setdist {x} T}"
 | |
| 1285 |   define T' where "T' \<equiv> {x \<in> U. setdist {x} T \<le> setdist {x} S}"
 | |
| 1286 |   define W  where "W \<equiv> {x \<in> U. setdist {x} S = setdist {x} T}"
 | |
| 1287 | have cloUS': "closedin (subtopology euclidean U) S'" | |
| 1288 |     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} S - setdist {x} T" "{..0}"]
 | |
| 1289 | by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) | |
| 1290 | have cloUT': "closedin (subtopology euclidean U) T'" | |
| 1291 |     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} T - setdist {x} S" "{..0}"]
 | |
| 1292 | by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) | |
| 1293 | have "S \<subseteq> S'" | |
| 1294 | using S'_def \<open>S \<subseteq> U\<close> setdist_sing_in_set by fastforce | |
| 1295 | have "T \<subseteq> T'" | |
| 1296 | using T'_def \<open>T \<subseteq> U\<close> setdist_sing_in_set by fastforce | |
| 1297 | have "S' \<union> T' = U" | |
| 1298 | by (auto simp: S'_def T'_def) | |
| 1299 | have "W \<subseteq> S'" | |
| 1300 | by (simp add: Collect_mono S'_def W_def) | |
| 1301 | have "W \<subseteq> T'" | |
| 1302 | by (simp add: Collect_mono T'_def W_def) | |
| 1303 | have ST_W: "S \<inter> T \<subseteq> W" and "W \<subseteq> U" | |
| 1304 | using \<open>S \<subseteq> U\<close> by (force simp: W_def setdist_sing_in_set)+ | |
| 1305 | have "S' \<inter> T' = W" | |
| 1306 | by (auto simp: S'_def T'_def W_def) | |
| 1307 | then have cloUW: "closedin (subtopology euclidean U) W" | |
| 1308 | using closedin_Int cloUS' cloUT' by blast | |
| 1309 | obtain W' W0 where "openin (subtopology euclidean W) W'" | |
| 1310 | and cloWW0: "closedin (subtopology euclidean W) W0" | |
| 1311 | and "S \<inter> T \<subseteq> W'" "W' \<subseteq> W0" | |
| 1312 | and ret: "(S \<inter> T) retract_of W0" | |
| 1313 | apply (rule ANR_imp_closed_neighbourhood_retract [OF \<open>ANR(S \<inter> T)\<close>]) | |
| 1314 | apply (rule closedin_subset_trans [of U, OF _ ST_W \<open>W \<subseteq> U\<close>]) | |
| 1315 | apply (blast intro: assms)+ | |
| 1316 | done | |
| 1317 | then obtain U0 where opeUU0: "openin (subtopology euclidean U) U0" | |
| 1318 | and U0: "S \<inter> T \<subseteq> U0" "U0 \<inter> W \<subseteq> W0" | |
| 1319 | unfolding openin_open using \<open>W \<subseteq> U\<close> by blast | |
| 1320 | have "W0 \<subseteq> U" | |
| 1321 | using \<open>W \<subseteq> U\<close> cloWW0 closedin_subset by fastforce | |
| 1322 | obtain r0 | |
| 1323 | where "S \<inter> T \<subseteq> W0" and contr0: "continuous_on W0 r0" and "r0 ` W0 \<subseteq> S \<inter> T" | |
| 1324 | and r0 [simp]: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> r0 x = x" | |
| 1325 | using ret by (force simp: retract_of_def retraction_def) | |
| 1326 | have ST: "x \<in> W \<Longrightarrow> x \<in> S \<longleftrightarrow> x \<in> T" for x | |
| 1327 | using assms by (auto simp: W_def setdist_sing_in_set dest!: setdist_eq_0_closedin) | |
| 1328 | define r where "r \<equiv> \<lambda>x. if x \<in> W0 then r0 x else x" | |
| 1329 | have "r ` (W0 \<union> S) \<subseteq> S" "r ` (W0 \<union> T) \<subseteq> T" | |
| 1330 | using \<open>r0 ` W0 \<subseteq> S \<inter> T\<close> r_def by auto | |
| 1331 | have contr: "continuous_on (W0 \<union> (S \<union> T)) r" | |
| 1332 | unfolding r_def | |
| 1333 | proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) | |
| 1334 | show "closedin (subtopology euclidean (W0 \<union> (S \<union> T))) W0" | |
| 1335 | apply (rule closedin_subset_trans [of U]) | |
| 1336 | using cloWW0 cloUW closedin_trans \<open>W0 \<subseteq> U\<close> \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> apply blast+ | |
| 1337 | done | |
| 1338 | show "closedin (subtopology euclidean (W0 \<union> (S \<union> T))) (S \<union> T)" | |
| 1339 | 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) | |
| 1340 | show "\<And>x. x \<in> W0 \<and> x \<notin> W0 \<or> x \<in> S \<union> T \<and> x \<in> W0 \<Longrightarrow> r0 x = x" | |
| 1341 | using ST cloWW0 closedin_subset by fastforce | |
| 1342 | qed | |
| 1343 | have cloS'WS: "closedin (subtopology euclidean S') (W0 \<union> S)" | |
| 1344 | by (meson closedin_subset_trans US cloUS' \<open>S \<subseteq> S'\<close> \<open>W \<subseteq> S'\<close> cloUW cloWW0 | |
| 1345 | closedin_Un closedin_imp_subset closedin_trans) | |
| 1346 | obtain W1 g where "W0 \<union> S \<subseteq> W1" and contg: "continuous_on W1 g" | |
| 1347 | and opeSW1: "openin (subtopology euclidean S') W1" | |
| 1348 | and "g ` W1 \<subseteq> S" and geqr: "\<And>x. x \<in> W0 \<union> S \<Longrightarrow> g x = r x" | |
| 1349 | apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> _ \<open>r ` (W0 \<union> S) \<subseteq> S\<close> cloS'WS]) | |
| 1350 | apply (rule continuous_on_subset [OF contr], blast+) | |
| 1351 | done | |
| 1352 | have cloT'WT: "closedin (subtopology euclidean T') (W0 \<union> T)" | |
| 1353 | by (meson closedin_subset_trans UT cloUT' \<open>T \<subseteq> T'\<close> \<open>W \<subseteq> T'\<close> cloUW cloWW0 | |
| 1354 | closedin_Un closedin_imp_subset closedin_trans) | |
| 1355 | obtain W2 h where "W0 \<union> T \<subseteq> W2" and conth: "continuous_on W2 h" | |
| 1356 | and opeSW2: "openin (subtopology euclidean T') W2" | |
| 1357 | and "h ` W2 \<subseteq> T" and heqr: "\<And>x. x \<in> W0 \<union> T \<Longrightarrow> h x = r x" | |
| 1358 | apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> _ \<open>r ` (W0 \<union> T) \<subseteq> T\<close> cloT'WT]) | |
| 1359 | apply (rule continuous_on_subset [OF contr], blast+) | |
| 1360 | done | |
| 1361 | have "S' \<inter> T' = W" | |
| 1362 | by (force simp: S'_def T'_def W_def) | |
| 1363 | obtain O1 O2 where "open O1" "W1 = S' \<inter> O1" "open O2" "W2 = T' \<inter> O2" | |
| 1364 | using opeSW1 opeSW2 by (force simp: openin_open) | |
| 1365 | show ?thesis | |
| 1366 | proof | |
| 1367 | have eq: "W1 - (W - U0) \<union> (W2 - (W - U0)) = | |
| 1368 | ((U - T') \<inter> O1 \<union> (U - S') \<inter> O2 \<union> U \<inter> O1 \<inter> O2) - (W - U0)" | |
| 1369 | using \<open>U0 \<inter> W \<subseteq> W0\<close> \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W0 \<union> T \<subseteq> W2\<close> | |
| 1370 | 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>) | |
| 1371 | show "openin (subtopology euclidean U) (W1 - (W - U0) \<union> (W2 - (W - U0)))" | |
| 1372 | apply (subst eq) | |
| 1373 | apply (intro openin_Un openin_Int_open openin_diff closedin_diff cloUW opeUU0 cloUS' cloUT' \<open>open O1\<close> \<open>open O2\<close>, simp_all) | |
| 1374 | done | |
| 1375 | have cloW1: "closedin (subtopology euclidean (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W1 - (W - U0))" | |
| 1376 | using cloUS' apply (simp add: closedin_closed) | |
| 1377 | apply (erule ex_forward) | |
| 1378 | using U0 \<open>W0 \<union> S \<subseteq> W1\<close> | |
| 1379 | apply (auto simp: \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<union> T' = U\<close> [symmetric]\<open>S' \<inter> T' = W\<close> [symmetric]) | |
| 1380 | done | |
| 1381 | have cloW2: "closedin (subtopology euclidean (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W2 - (W - U0))" | |
| 1382 | using cloUT' apply (simp add: closedin_closed) | |
| 1383 | apply (erule ex_forward) | |
| 1384 | using U0 \<open>W0 \<union> T \<subseteq> W2\<close> | |
| 1385 | apply (auto simp: \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<union> T' = U\<close> [symmetric]\<open>S' \<inter> T' = W\<close> [symmetric]) | |
| 1386 | done | |
| 1387 | have *: "\<forall>x\<in>S \<union> T. (if x \<in> S' then g x else h x) = x" | |
| 1388 | using ST \<open>S' \<inter> T' = W\<close> cloT'WT closedin_subset geqr heqr | |
| 1389 | apply (auto simp: r_def, fastforce) | |
| 1390 | 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 | |
| 1391 | have "\<exists>r. continuous_on (W1 - (W - U0) \<union> (W2 - (W - U0))) r \<and> | |
| 1392 | r ` (W1 - (W - U0) \<union> (W2 - (W - U0))) \<subseteq> S \<union> T \<and> | |
| 1393 | (\<forall>x\<in>S \<union> T. r x = x)" | |
| 1394 | apply (rule_tac x = "\<lambda>x. if x \<in> S' then g x else h x" in exI) | |
| 1395 | apply (intro conjI *) | |
| 1396 | apply (rule continuous_on_cases_local | |
| 1397 | [OF cloW1 cloW2 continuous_on_subset [OF contg] continuous_on_subset [OF conth]]) | |
| 1398 | using \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<inter> T' = W\<close> | |
| 1399 | \<open>g ` W1 \<subseteq> S\<close> \<open>h ` W2 \<subseteq> T\<close> apply auto | |
| 1400 | using \<open>U0 \<inter> W \<subseteq> W0\<close> \<open>W0 \<union> S \<subseteq> W1\<close> apply (fastforce simp add: geqr heqr)+ | |
| 1401 | done | |
| 1402 | then show "S \<union> T retract_of W1 - (W - U0) \<union> (W2 - (W - U0))" | |
| 1403 | using \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W0 \<union> T \<subseteq> W2\<close> ST opeUU0 U0 | |
| 1404 | by (auto simp: retract_of_def retraction_def) | |
| 1405 | qed | |
| 1406 | qed | |
| 1407 | ||
| 1408 | ||
| 1409 | lemma ANR_closed_Un_local: | |
| 1410 | fixes S :: "'a::euclidean_space set" | |
| 1411 | assumes STS: "closedin (subtopology euclidean (S \<union> T)) S" | |
| 1412 | and STT: "closedin (subtopology euclidean (S \<union> T)) T" | |
| 1413 | and "ANR S" "ANR T" "ANR(S \<inter> T)" | |
| 1414 | shows "ANR(S \<union> T)" | |
| 1415 | proof - | |
| 1416 | have "\<exists>T. openin (subtopology euclidean U) T \<and> C retract_of T" | |
| 1417 | if hom: "S \<union> T homeomorphic C" and UC: "closedin (subtopology euclidean U) C" | |
| 1418 |        for U and C :: "('a * real) set"
 | |
| 1419 | proof - | |
| 1420 | obtain f g where hom: "homeomorphism (S \<union> T) C f g" | |
| 1421 | using hom by (force simp: homeomorphic_def) | |
| 1422 | have US: "closedin (subtopology euclidean U) (C \<inter> g -` S)" | |
| 1423 | apply (rule closedin_trans [OF _ UC]) | |
| 1424 | apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) | |
| 1425 | using hom [unfolded homeomorphism_def] apply blast | |
| 1426 | apply (metis hom homeomorphism_def set_eq_subset) | |
| 1427 | done | |
| 1428 | have UT: "closedin (subtopology euclidean U) (C \<inter> g -` T)" | |
| 1429 | apply (rule closedin_trans [OF _ UC]) | |
| 1430 | apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) | |
| 1431 | using hom [unfolded homeomorphism_def] apply blast | |
| 1432 | apply (metis hom homeomorphism_def set_eq_subset) | |
| 1433 | done | |
| 1434 | have ANRS: "ANR (C \<inter> g -` S)" | |
| 1435 | apply (rule ANR_homeomorphic_ANR [OF \<open>ANR S\<close>]) | |
| 1436 | apply (simp add: homeomorphic_def) | |
| 1437 | apply (rule_tac x=g in exI) | |
| 1438 | apply (rule_tac x=f in exI) | |
| 1439 | using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) | |
| 1440 | apply (rule_tac x="f x" in image_eqI, auto) | |
| 1441 | done | |
| 1442 | have ANRT: "ANR (C \<inter> g -` T)" | |
| 1443 | apply (rule ANR_homeomorphic_ANR [OF \<open>ANR T\<close>]) | |
| 1444 | apply (simp add: homeomorphic_def) | |
| 1445 | apply (rule_tac x=g in exI) | |
| 1446 | apply (rule_tac x=f in exI) | |
| 1447 | using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) | |
| 1448 | apply (rule_tac x="f x" in image_eqI, auto) | |
| 1449 | done | |
| 1450 | have ANRI: "ANR ((C \<inter> g -` S) \<inter> (C \<inter> g -` T))" | |
| 1451 | apply (rule ANR_homeomorphic_ANR [OF \<open>ANR (S \<inter> T)\<close>]) | |
| 1452 | apply (simp add: homeomorphic_def) | |
| 1453 | apply (rule_tac x=g in exI) | |
| 1454 | apply (rule_tac x=f in exI) | |
| 1455 | using hom | |
| 1456 | apply (auto simp: homeomorphism_def elim!: continuous_on_subset) | |
| 1457 | apply (rule_tac x="f x" in image_eqI, auto) | |
| 1458 | done | |
| 1459 | have "C = (C \<inter> g -` S) \<union> (C \<inter> g -` T)" | |
| 1460 | using hom by (auto simp: homeomorphism_def) | |
| 1461 | then show ?thesis | |
| 1462 | by (metis ANR_closed_Un_local_aux [OF US UT ANRS ANRT ANRI]) | |
| 1463 | qed | |
| 1464 | then show ?thesis | |
| 1465 | by (auto simp: ANR_def) | |
| 1466 | qed | |
| 1467 | ||
| 1468 | corollary ANR_closed_Un: | |
| 1469 | fixes S :: "'a::euclidean_space set" | |
| 1470 | shows "\<lbrakk>closed S; closed T; ANR S; ANR T; ANR (S \<inter> T)\<rbrakk> \<Longrightarrow> ANR (S \<union> T)" | |
| 1471 | by (simp add: ANR_closed_Un_local closedin_def diff_eq open_Compl openin_open_Int) | |
| 1472 | ||
| 1473 | lemma ANR_openin: | |
| 1474 | fixes S :: "'a::euclidean_space set" | |
| 1475 | assumes "ANR T" and opeTS: "openin (subtopology euclidean T) S" | |
| 1476 | shows "ANR S" | |
| 1477 | proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) | |
| 1478 | fix f :: "'a \<times> real \<Rightarrow> 'a" and U C | |
| 1479 | assume contf: "continuous_on C f" and fim: "f ` C \<subseteq> S" | |
| 1480 | and cloUC: "closedin (subtopology euclidean U) C" | |
| 1481 | have "f ` C \<subseteq> T" | |
| 1482 | using fim opeTS openin_imp_subset by blast | |
| 1483 | obtain W g where "C \<subseteq> W" | |
| 1484 | and UW: "openin (subtopology euclidean U) W" | |
| 1485 | and contg: "continuous_on W g" | |
| 1486 | and gim: "g ` W \<subseteq> T" | |
| 1487 | and geq: "\<And>x. x \<in> C \<Longrightarrow> g x = f x" | |
| 1488 | apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> contf \<open>f ` C \<subseteq> T\<close> cloUC]) | |
| 1489 | using fim by auto | |
| 1490 | 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)" | |
| 1491 | proof (intro exI conjI) | |
| 1492 | show "C \<subseteq> W \<inter> g -` S" | |
| 1493 | using \<open>C \<subseteq> W\<close> fim geq by blast | |
| 1494 | show "openin (subtopology euclidean U) (W \<inter> g -` S)" | |
| 1495 | by (metis (mono_tags, lifting) UW contg continuous_openin_preimage gim opeTS openin_trans) | |
| 1496 | show "continuous_on (W \<inter> g -` S) g" | |
| 1497 | by (blast intro: continuous_on_subset [OF contg]) | |
| 1498 | show "g ` (W \<inter> g -` S) \<subseteq> S" | |
| 1499 | using gim by blast | |
| 1500 | show "\<forall>x\<in>C. g x = f x" | |
| 1501 | using geq by blast | |
| 1502 | qed | |
| 1503 | qed | |
| 1504 | ||
| 1505 | lemma ENR_openin: | |
| 1506 | fixes S :: "'a::euclidean_space set" | |
| 1507 | assumes "ENR T" and opeTS: "openin (subtopology euclidean T) S" | |
| 1508 | shows "ENR S" | |
| 1509 | using assms apply (simp add: ENR_ANR) | |
| 1510 | using ANR_openin locally_open_subset by blast | |
| 1511 | ||
| 1512 | lemma ANR_neighborhood_retract: | |
| 1513 | fixes S :: "'a::euclidean_space set" | |
| 1514 | assumes "ANR U" "S retract_of T" "openin (subtopology euclidean U) T" | |
| 1515 | shows "ANR S" | |
| 1516 | using ANR_openin ANR_retract_of_ANR assms by blast | |
| 1517 | ||
| 1518 | lemma ENR_neighborhood_retract: | |
| 1519 | fixes S :: "'a::euclidean_space set" | |
| 1520 | assumes "ENR U" "S retract_of T" "openin (subtopology euclidean U) T" | |
| 1521 | shows "ENR S" | |
| 1522 | using ENR_openin ENR_retract_of_ENR assms by blast | |
| 1523 | ||
| 1524 | lemma ANR_rel_interior: | |
| 1525 | fixes S :: "'a::euclidean_space set" | |
| 1526 | shows "ANR S \<Longrightarrow> ANR(rel_interior S)" | |
| 1527 | by (blast intro: ANR_openin openin_set_rel_interior) | |
| 1528 | ||
| 1529 | lemma ANR_delete: | |
| 1530 | fixes S :: "'a::euclidean_space set" | |
| 1531 |   shows "ANR S \<Longrightarrow> ANR(S - {a})"
 | |
| 1532 | by (blast intro: ANR_openin openin_delete openin_subtopology_self) | |
| 1533 | ||
| 1534 | lemma ENR_rel_interior: | |
| 1535 | fixes S :: "'a::euclidean_space set" | |
| 1536 | shows "ENR S \<Longrightarrow> ENR(rel_interior S)" | |
| 1537 | by (blast intro: ENR_openin openin_set_rel_interior) | |
| 1538 | ||
| 1539 | lemma ENR_delete: | |
| 1540 | fixes S :: "'a::euclidean_space set" | |
| 1541 |   shows "ENR S \<Longrightarrow> ENR(S - {a})"
 | |
| 1542 | by (blast intro: ENR_openin openin_delete openin_subtopology_self) | |
| 1543 | ||
| 1544 | lemma open_imp_ENR: "open S \<Longrightarrow> ENR S" | |
| 1545 | using ENR_def by blast | |
| 1546 | ||
| 1547 | lemma open_imp_ANR: | |
| 1548 | fixes S :: "'a::euclidean_space set" | |
| 1549 | shows "open S \<Longrightarrow> ANR S" | |
| 1550 | by (simp add: ENR_imp_ANR open_imp_ENR) | |
| 1551 | ||
| 1552 | lemma ANR_ball [iff]: | |
| 1553 | fixes a :: "'a::euclidean_space" | |
| 1554 | shows "ANR(ball a r)" | |
| 1555 | by (simp add: convex_imp_ANR) | |
| 1556 | ||
| 1557 | lemma ENR_ball [iff]: "ENR(ball a r)" | |
| 1558 | by (simp add: open_imp_ENR) | |
| 1559 | ||
| 1560 | lemma AR_ball [simp]: | |
| 1561 | fixes a :: "'a::euclidean_space" | |
| 1562 | shows "AR(ball a r) \<longleftrightarrow> 0 < r" | |
| 1563 | by (auto simp: AR_ANR convex_imp_contractible) | |
| 1564 | ||
| 1565 | lemma ANR_cball [iff]: | |
| 1566 | fixes a :: "'a::euclidean_space" | |
| 1567 | shows "ANR(cball a r)" | |
| 1568 | by (simp add: convex_imp_ANR) | |
| 1569 | ||
| 1570 | lemma ENR_cball: | |
| 1571 | fixes a :: "'a::euclidean_space" | |
| 1572 | shows "ENR(cball a r)" | |
| 1573 | using ENR_convex_closed by blast | |
| 1574 | ||
| 1575 | lemma AR_cball [simp]: | |
| 1576 | fixes a :: "'a::euclidean_space" | |
| 1577 | shows "AR(cball a r) \<longleftrightarrow> 0 \<le> r" | |
| 1578 | by (auto simp: AR_ANR convex_imp_contractible) | |
| 1579 | ||
| 1580 | lemma ANR_box [iff]: | |
| 1581 | fixes a :: "'a::euclidean_space" | |
| 1582 | shows "ANR(cbox a b)" "ANR(box a b)" | |
| 1583 | by (auto simp: convex_imp_ANR open_imp_ANR) | |
| 1584 | ||
| 1585 | lemma ENR_box [iff]: | |
| 1586 | fixes a :: "'a::euclidean_space" | |
| 1587 | shows "ENR(cbox a b)" "ENR(box a b)" | |
| 1588 | apply (simp add: ENR_convex_closed closed_cbox) | |
| 1589 | by (simp add: open_box open_imp_ENR) | |
| 1590 | ||
| 1591 | lemma AR_box [simp]: | |
| 1592 |     "AR(cbox a b) \<longleftrightarrow> cbox a b \<noteq> {}" "AR(box a b) \<longleftrightarrow> box a b \<noteq> {}"
 | |
| 1593 | by (auto simp: AR_ANR convex_imp_contractible) | |
| 1594 | ||
| 1595 | lemma ANR_interior: | |
| 1596 | fixes S :: "'a::euclidean_space set" | |
| 1597 | shows "ANR(interior S)" | |
| 1598 | by (simp add: open_imp_ANR) | |
| 1599 | ||
| 1600 | lemma ENR_interior: | |
| 1601 | fixes S :: "'a::euclidean_space set" | |
| 1602 | shows "ENR(interior S)" | |
| 1603 | by (simp add: open_imp_ENR) | |
| 1604 | ||
| 1605 | lemma AR_imp_contractible: | |
| 1606 | fixes S :: "'a::euclidean_space set" | |
| 1607 | shows "AR S \<Longrightarrow> contractible S" | |
| 1608 | by (simp add: AR_ANR) | |
| 1609 | ||
| 1610 | lemma ENR_imp_locally_compact: | |
| 1611 | fixes S :: "'a::euclidean_space set" | |
| 1612 | shows "ENR S \<Longrightarrow> locally compact S" | |
| 1613 | by (simp add: ENR_ANR) | |
| 1614 | ||
| 1615 | lemma ANR_imp_locally_path_connected: | |
| 1616 | fixes S :: "'a::euclidean_space set" | |
| 1617 | assumes "ANR S" | |
| 1618 | shows "locally path_connected S" | |
| 1619 | proof - | |
| 1620 |   obtain U and T :: "('a \<times> real) set"
 | |
| 1621 |      where "convex U" "U \<noteq> {}"
 | |
| 1622 | and UT: "closedin (subtopology euclidean U) T" | |
| 1623 | and "S homeomorphic T" | |
| 1624 | apply (rule homeomorphic_closedin_convex [of S]) | |
| 1625 | using aff_dim_le_DIM [of S] apply auto | |
| 1626 | done | |
| 1627 | then have "locally path_connected T" | |
| 1628 | by (meson ANR_imp_absolute_neighbourhood_retract | |
| 1629 | assms convex_imp_locally_path_connected locally_open_subset retract_of_locally_path_connected) | |
| 1630 | then have S: "locally path_connected S" | |
| 1631 |       if "openin (subtopology euclidean U) V" "T retract_of V" "U \<noteq> {}" for V
 | |
| 1632 | using \<open>S homeomorphic T\<close> homeomorphic_locally homeomorphic_path_connectedness by blast | |
| 1633 | show ?thesis | |
| 1634 | using assms | |
| 1635 | apply (clarsimp simp: ANR_def) | |
| 1636 | apply (drule_tac x=U in spec) | |
| 1637 | apply (drule_tac x=T in spec) | |
| 1638 |     using \<open>S homeomorphic T\<close> \<open>U \<noteq> {}\<close> UT  apply (blast intro: S)
 | |
| 1639 | done | |
| 1640 | qed | |
| 1641 | ||
| 1642 | lemma ANR_imp_locally_connected: | |
| 1643 | fixes S :: "'a::euclidean_space set" | |
| 1644 | assumes "ANR S" | |
| 1645 | shows "locally connected S" | |
| 1646 | using locally_path_connected_imp_locally_connected ANR_imp_locally_path_connected assms by auto | |
| 1647 | ||
| 1648 | lemma AR_imp_locally_path_connected: | |
| 1649 | fixes S :: "'a::euclidean_space set" | |
| 1650 | assumes "AR S" | |
| 1651 | shows "locally path_connected S" | |
| 1652 | by (simp add: ANR_imp_locally_path_connected AR_imp_ANR assms) | |
| 1653 | ||
| 1654 | lemma AR_imp_locally_connected: | |
| 1655 | fixes S :: "'a::euclidean_space set" | |
| 1656 | assumes "AR S" | |
| 1657 | shows "locally connected S" | |
| 1658 | using ANR_imp_locally_connected AR_ANR assms by blast | |
| 1659 | ||
| 1660 | lemma ENR_imp_locally_path_connected: | |
| 1661 | fixes S :: "'a::euclidean_space set" | |
| 1662 | assumes "ENR S" | |
| 1663 | shows "locally path_connected S" | |
| 1664 | by (simp add: ANR_imp_locally_path_connected ENR_imp_ANR assms) | |
| 1665 | ||
| 1666 | lemma ENR_imp_locally_connected: | |
| 1667 | fixes S :: "'a::euclidean_space set" | |
| 1668 | assumes "ENR S" | |
| 1669 | shows "locally connected S" | |
| 1670 | using ANR_imp_locally_connected ENR_ANR assms by blast | |
| 1671 | ||
| 1672 | lemma ANR_Times: | |
| 1673 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | |
| 1674 | assumes "ANR S" "ANR T" shows "ANR(S \<times> T)" | |
| 1675 | proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) | |
| 1676 |   fix f :: " ('a \<times> 'b) \<times> real \<Rightarrow> 'a \<times> 'b" and U C
 | |
| 1677 | assume "continuous_on C f" and fim: "f ` C \<subseteq> S \<times> T" | |
| 1678 | and cloUC: "closedin (subtopology euclidean U) C" | |
| 1679 | have contf1: "continuous_on C (fst \<circ> f)" | |
| 1680 | by (simp add: \<open>continuous_on C f\<close> continuous_on_fst) | |
| 1681 | obtain W1 g where "C \<subseteq> W1" | |
| 1682 | and UW1: "openin (subtopology euclidean U) W1" | |
| 1683 | and contg: "continuous_on W1 g" | |
| 1684 | and gim: "g ` W1 \<subseteq> S" | |
| 1685 | and geq: "\<And>x. x \<in> C \<Longrightarrow> g x = (fst \<circ> f) x" | |
| 1686 | apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> contf1 _ cloUC]) | |
| 1687 | using fim apply auto | |
| 1688 | done | |
| 1689 | have contf2: "continuous_on C (snd \<circ> f)" | |
| 1690 | by (simp add: \<open>continuous_on C f\<close> continuous_on_snd) | |
| 1691 | obtain W2 h where "C \<subseteq> W2" | |
| 1692 | and UW2: "openin (subtopology euclidean U) W2" | |
| 1693 | and conth: "continuous_on W2 h" | |
| 1694 | and him: "h ` W2 \<subseteq> T" | |
| 1695 | and heq: "\<And>x. x \<in> C \<Longrightarrow> h x = (snd \<circ> f) x" | |
| 1696 | apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> contf2 _ cloUC]) | |
| 1697 | using fim apply auto | |
| 1698 | done | |
| 1699 | show "\<exists>V g. C \<subseteq> V \<and> | |
| 1700 | openin (subtopology euclidean U) V \<and> | |
| 1701 | continuous_on V g \<and> g ` V \<subseteq> S \<times> T \<and> (\<forall>x\<in>C. g x = f x)" | |
| 1702 | proof (intro exI conjI) | |
| 1703 | show "C \<subseteq> W1 \<inter> W2" | |
| 1704 | by (simp add: \<open>C \<subseteq> W1\<close> \<open>C \<subseteq> W2\<close>) | |
| 1705 | show "openin (subtopology euclidean U) (W1 \<inter> W2)" | |
| 1706 | by (simp add: UW1 UW2 openin_Int) | |
| 1707 | show "continuous_on (W1 \<inter> W2) (\<lambda>x. (g x, h x))" | |
| 1708 | by (metis (no_types) contg conth continuous_on_Pair continuous_on_subset inf_commute inf_le1) | |
| 1709 | show "(\<lambda>x. (g x, h x)) ` (W1 \<inter> W2) \<subseteq> S \<times> T" | |
| 1710 | using gim him by blast | |
| 1711 | show "(\<forall>x\<in>C. (g x, h x) = f x)" | |
| 1712 | using geq heq by auto | |
| 1713 | qed | |
| 1714 | qed | |
| 1715 | ||
| 1716 | lemma AR_Times: | |
| 1717 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | |
| 1718 | assumes "AR S" "AR T" shows "AR(S \<times> T)" | |
| 1719 | using assms by (simp add: AR_ANR ANR_Times contractible_Times) | |
| 1720 | ||
| 1721 | subsection \<open>Kuhn Simplices\<close> | |
| 1722 | ||
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1723 | lemma bij_betw_singleton_eq: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1724 | assumes f: "bij_betw f A B" and g: "bij_betw g A B" and a: "a \<in> A" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1725 | assumes eq: "(\<And>x. x \<in> A \<Longrightarrow> x \<noteq> a \<Longrightarrow> f x = g x)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1726 | shows "f a = g a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1727 | proof - | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1728 |   have "f ` (A - {a}) = g ` (A - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1729 | by (intro image_cong) (simp_all add: eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1730 |   then have "B - {f a} = B - {g a}"
 | 
| 60303 | 1731 | using f g a by (auto simp: bij_betw_def inj_on_image_set_diff set_eq_iff Diff_subset) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1732 | moreover have "f a \<in> B" "g a \<in> B" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1733 | using f g a by (auto simp: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1734 | ultimately show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1735 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1736 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1737 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1738 | lemma swap_image: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1739 |   "Fun.swap i j f ` A = (if i \<in> A then (if j \<in> A then f ` A else f ` ((A - {i}) \<union> {j}))
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1740 |                                   else (if j \<in> A then f ` ((A - {j}) \<union> {i}) else f ` A))"
 | 
| 68022 | 1741 | by (auto simp: swap_def image_def) metis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1742 | |
| 63365 | 1743 | lemmas swap_apply1 = swap_apply(1) | 
| 1744 | lemmas swap_apply2 = swap_apply(2) | |
| 1745 | lemmas Zero_notin_Suc = zero_notin_Suc_image | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1746 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1747 | lemma pointwise_minimal_pointwise_maximal: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1748 | fixes s :: "(nat \<Rightarrow> nat) set" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1749 | assumes "finite s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1750 |     and "s \<noteq> {}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1751 | and "\<forall>x\<in>s. \<forall>y\<in>s. x \<le> y \<or> y \<le> x" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1752 | shows "\<exists>a\<in>s. \<forall>x\<in>s. a \<le> x" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1753 | and "\<exists>a\<in>s. \<forall>x\<in>s. x \<le> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1754 | using assms | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1755 | proof (induct s rule: finite_ne_induct) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1756 | case (insert b s) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1757 | assume *: "\<forall>x\<in>insert b s. \<forall>y\<in>insert b s. x \<le> y \<or> y \<le> x" | 
| 63540 | 1758 | 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" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1759 | using insert by auto | 
| 63540 | 1760 | 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" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1761 | using *[rule_format, of b u] *[rule_format, of b l] by (metis insert_iff order.trans)+ | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1762 | qed auto | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1763 | |
| 68617 | 1764 | (* FIXME mv *) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1765 | lemma brouwer_compactness_lemma: | 
| 56226 | 1766 | fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" | 
| 53674 | 1767 | assumes "compact s" | 
| 1768 | and "continuous_on s f" | |
| 53688 | 1769 | and "\<not> (\<exists>x\<in>s. f x = 0)" | 
| 53674 | 1770 | obtains d where "0 < d" and "\<forall>x\<in>s. d \<le> norm (f x)" | 
| 53185 | 1771 | proof (cases "s = {}")
 | 
| 53674 | 1772 | case True | 
| 53688 | 1773 | show thesis | 
| 1774 | by (rule that [of 1]) (auto simp: True) | |
| 53674 | 1775 | next | 
| 49374 | 1776 | case False | 
| 1777 | have "continuous_on s (norm \<circ> f)" | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56273diff
changeset | 1778 | by (rule continuous_intros continuous_on_norm assms(2))+ | 
| 53674 | 1779 | with False obtain x where x: "x \<in> s" "\<forall>y\<in>s. (norm \<circ> f) x \<le> (norm \<circ> f) y" | 
| 1780 | using continuous_attains_inf[OF assms(1), of "norm \<circ> f"] | |
| 1781 | unfolding o_def | |
| 1782 | by auto | |
| 1783 | have "(norm \<circ> f) x > 0" | |
| 1784 | using assms(3) and x(1) | |
| 1785 | by auto | |
| 1786 | then show ?thesis | |
| 1787 | by (rule that) (insert x(2), auto simp: o_def) | |
| 49555 | 1788 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1789 | |
| 49555 | 1790 | lemma kuhn_labelling_lemma: | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1791 | fixes P Q :: "'a::euclidean_space \<Rightarrow> bool" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1792 | assumes "\<forall>x. P x \<longrightarrow> P (f x)" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1793 | and "\<forall>x. P x \<longrightarrow> (\<forall>i\<in>Basis. Q i \<longrightarrow> 0 \<le> x\<bullet>i \<and> x\<bullet>i \<le> 1)" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1794 | shows "\<exists>l. (\<forall>x.\<forall>i\<in>Basis. l x i \<le> (1::nat)) \<and> | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1795 | (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (x\<bullet>i = 0) \<longrightarrow> (l x i = 0)) \<and> | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1796 | (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (x\<bullet>i = 1) \<longrightarrow> (l x i = 1)) \<and> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1797 | (\<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> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1798 | (\<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)" | 
| 49374 | 1799 | proof - | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1800 |   { fix x i
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1801 | let ?R = "\<lambda>y. (P x \<and> Q i \<and> x \<bullet> i = 0 \<longrightarrow> y = (0::nat)) \<and> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1802 | (P x \<and> Q i \<and> x \<bullet> i = 1 \<longrightarrow> y = 1) \<and> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1803 | (P x \<and> Q i \<and> y = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i) \<and> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1804 | (P x \<and> Q i \<and> y = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1805 |     { 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 }
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1806 | then have "i \<in> Basis \<Longrightarrow> ?R 0 \<or> ?R 1" by auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1807 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1808 | unfolding all_conj_distrib[symmetric] Ball_def (* FIXME: shouldn't this work by metis? *) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1809 | by (subst choice_iff[symmetric])+ blast | 
| 49374 | 1810 | qed | 
| 1811 | ||
| 53185 | 1812 | |
| 68617 | 1813 | subsubsection \<open>The key "counting" observation, somewhat abstracted\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1814 | |
| 53252 | 1815 | lemma kuhn_counting_lemma: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1816 | fixes bnd compo compo' face S F | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1817 |   defines "nF s == card {f\<in>F. face f s \<and> compo' f}"
 | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67399diff
changeset | 1818 | assumes [simp, intro]: "finite F" \<comment> \<open>faces\<close> and [simp, intro]: "finite S" \<comment> \<open>simplices\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1819 |     and "\<And>f. f \<in> F \<Longrightarrow> bnd f \<Longrightarrow> card {s\<in>S. face f s} = 1"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1820 |     and "\<And>f. f \<in> F \<Longrightarrow> \<not> bnd f \<Longrightarrow> card {s\<in>S. face f s} = 2"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1821 | and "\<And>s. s \<in> S \<Longrightarrow> compo s \<Longrightarrow> nF s = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1822 | and "\<And>s. s \<in> S \<Longrightarrow> \<not> compo s \<Longrightarrow> nF s = 0 \<or> nF s = 2" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1823 |     and "odd (card {f\<in>F. compo' f \<and> bnd f})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1824 |   shows "odd (card {s\<in>S. compo s})"
 | 
| 53185 | 1825 | proof - | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1826 | 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)" | 
| 64267 | 1827 | by (subst sum.union_disjoint[symmetric]) (auto intro!: sum.cong) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1828 |   also have "\<dots> = (\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> bnd f}. face f s}) +
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1829 |                   (\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> \<not> bnd f}. face f s})"
 | 
| 64267 | 1830 | unfolding sum.distrib[symmetric] | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1831 | by (subst card_Un_disjoint[symmetric]) | 
| 64267 | 1832 | (auto simp: nF_def intro!: sum.cong arg_cong[where f=card]) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1833 |   also have "\<dots> = 1 * card {f\<in>F. compo' f \<and> bnd f} + 2 * card {f\<in>F. compo' f \<and> \<not> bnd f}"
 | 
| 67399 | 1834 | using assms(4,5) by (fastforce intro!: arg_cong2[where f="(+)"] sum_multicount) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1835 |   finally have "odd ((\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) + card {s\<in>S. compo s})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1836 | using assms(6,8) by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1837 | moreover have "(\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) = | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1838 | (\<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)" | 
| 64267 | 1839 | using assms(7) by (subst sum.union_disjoint[symmetric]) (fastforce intro!: sum.cong)+ | 
| 53688 | 1840 | ultimately show ?thesis | 
| 1841 | by auto | |
| 53186 | 1842 | qed | 
| 1843 | ||
| 68617 | 1844 | subsubsection \<open>The odd/even result for faces of complete vertices, generalized\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1845 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1846 | lemma kuhn_complete_lemma: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1847 | assumes [simp]: "finite simplices" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1848 |     and face: "\<And>f s. face f s \<longleftrightarrow> (\<exists>a\<in>s. f = s - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1849 | and card_s[simp]: "\<And>s. s \<in> simplices \<Longrightarrow> card s = n + 2" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1850 |     and rl_bd: "\<And>s. s \<in> simplices \<Longrightarrow> rl ` s \<subseteq> {..Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1851 |     and bnd: "\<And>f s. s \<in> simplices \<Longrightarrow> face f s \<Longrightarrow> bnd f \<Longrightarrow> card {s\<in>simplices. face f s} = 1"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1852 |     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"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1853 |     and odd_card: "odd (card {f. (\<exists>s\<in>simplices. face f s) \<and> rl ` f = {..n} \<and> bnd f})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1854 |   shows "odd (card {s\<in>simplices. (rl ` s = {..Suc n})})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1855 | proof (rule kuhn_counting_lemma) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1856 | have finite_s[simp]: "\<And>s. s \<in> simplices \<Longrightarrow> finite s" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 1857 | by (metis add_is_0 zero_neq_numeral card_infinite assms(3)) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1858 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1859 |   let ?F = "{f. \<exists>s\<in>simplices. face f s}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1860 |   have F_eq: "?F = (\<Union>s\<in>simplices. \<Union>a\<in>s. {s - {a}})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1861 | by (auto simp: face) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1862 | show "finite ?F" | 
| 60420 | 1863 | using \<open>finite simplices\<close> unfolding F_eq by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1864 | |
| 60421 | 1865 |   show "card {s \<in> simplices. face f s} = 1" if "f \<in> ?F" "bnd f" for f
 | 
| 60449 | 1866 | using bnd that by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1867 | |
| 60421 | 1868 |   show "card {s \<in> simplices. face f s} = 2" if "f \<in> ?F" "\<not> bnd f" for f
 | 
| 60449 | 1869 | using nbnd that by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1870 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1871 |   show "odd (card {f \<in> {f. \<exists>s\<in>simplices. face f s}. rl ` f = {..n} \<and> bnd f})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1872 | using odd_card by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1873 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1874 | fix s assume s[simp]: "s \<in> simplices" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1875 |   let ?S = "{f \<in> {f. \<exists>s\<in>simplices. face f s}. face f s \<and> rl ` f = {..n}}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1876 |   have "?S = (\<lambda>a. s - {a}) ` {a\<in>s. rl ` (s - {a}) = {..n}}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1877 | using s by (fastforce simp: face) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1878 |   then have card_S: "card ?S = card {a\<in>s. rl ` (s - {a}) = {..n}}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1879 | by (auto intro!: card_image inj_onI) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1880 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1881 |   { assume rl: "rl ` s = {..Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1882 | then have inj_rl: "inj_on rl s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1883 | by (intro eq_card_imp_inj_on) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1884 | moreover obtain a where "rl a = Suc n" "a \<in> s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1885 | by (metis atMost_iff image_iff le_Suc_eq rl) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1886 |     ultimately have n: "{..n} = rl ` (s - {a})"
 | 
| 68022 | 1887 | by (auto simp: inj_on_image_set_diff Diff_subset rl) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1888 |     have "{a\<in>s. rl ` (s - {a}) = {..n}} = {a}"
 | 
| 68022 | 1889 | using inj_rl \<open>a \<in> s\<close> by (auto simp: n inj_on_image_eq_iff[OF inj_rl] Diff_subset) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1890 | then show "card ?S = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1891 | unfolding card_S by simp } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1892 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1893 |   { assume rl: "rl ` s \<noteq> {..Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1894 | show "card ?S = 0 \<or> card ?S = 2" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1895 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1896 |       assume *: "{..n} \<subseteq> rl ` s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1897 |       with rl rl_bd[OF s] have rl_s: "rl ` s = {..n}"
 | 
| 68022 | 1898 | by (auto simp: atMost_Suc subset_insert_iff split: if_split_asm) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1899 | then have "\<not> inj_on rl s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1900 | by (intro pigeonhole) simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1901 | then obtain a b where ab: "a \<in> s" "b \<in> s" "rl a = rl b" "a \<noteq> b" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1902 | by (auto simp: inj_on_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1903 |       then have eq: "rl ` (s - {a}) = rl ` s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1904 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1905 |       with ab have inj: "inj_on rl (s - {a})"
 | 
| 68022 | 1906 | by (intro eq_card_imp_inj_on) (auto simp: rl_s card_Diff_singleton_if) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1907 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1908 |       { fix x assume "x \<in> s" "x \<notin> {a, b}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1909 |         then have "rl ` s - {rl x} = rl ` ((s - {a}) - {x})"
 | 
| 60303 | 1910 | by (auto simp: eq Diff_subset inj_on_image_set_diff[OF inj]) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1911 |         also have "\<dots> = rl ` (s - {x})"
 | 
| 60420 | 1912 |           using ab \<open>x \<notin> {a, b}\<close> by auto
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1913 | also assume "\<dots> = rl ` s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1914 | finally have False | 
| 60420 | 1915 | using \<open>x\<in>s\<close> by auto } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1916 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1917 |       { fix x assume "x \<in> {a, b}" with ab have "x \<in> s \<and> rl ` (s - {x}) = rl ` s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1918 | by (simp add: set_eq_iff image_iff Bex_def) metis } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1919 |       ultimately have "{a\<in>s. rl ` (s - {a}) = {..n}} = {a, b}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1920 | unfolding rl_s[symmetric] by fastforce | 
| 60420 | 1921 | with \<open>a \<noteq> b\<close> show "card ?S = 0 \<or> card ?S = 2" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1922 | unfolding card_S by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1923 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1924 |       assume "\<not> {..n} \<subseteq> rl ` s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1925 |       then have "\<And>x. rl ` (s - {x}) \<noteq> {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1926 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1927 | then show "card ?S = 0 \<or> card ?S = 2" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1928 | unfolding card_S by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1929 | qed } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1930 | qed fact | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1931 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1932 | locale kuhn_simplex = | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1933 | fixes p n and base upd and s :: "(nat \<Rightarrow> nat) set" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1934 |   assumes base: "base \<in> {..< n} \<rightarrow> {..< p}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1935 | assumes base_out: "\<And>i. n \<le> i \<Longrightarrow> base i = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1936 |   assumes upd: "bij_betw upd {..< n} {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1937 |   assumes s_pre: "s = (\<lambda>i j. if j \<in> upd`{..< i} then Suc (base j) else base j) ` {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1938 | begin | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1939 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1940 | definition "enum i j = (if j \<in> upd`{..< i} then Suc (base j) else base j)"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1941 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1942 | lemma s_eq: "s = enum ` {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1943 | unfolding s_pre enum_def[abs_def] .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1944 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1945 | lemma upd_space: "i < n \<Longrightarrow> upd i < n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1946 | using upd by (auto dest!: bij_betwE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1947 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1948 | lemma s_space: "s \<subseteq> {..< n} \<rightarrow> {.. p}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1949 | proof - | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1950 |   { fix i assume "i \<le> n" then have "enum i \<in> {..< n} \<rightarrow> {.. p}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1951 | proof (induct i) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1952 | case 0 then show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1953 | using base by (auto simp: Pi_iff less_imp_le enum_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1954 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1955 | case (Suc i) with base show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1956 | by (auto simp: Pi_iff Suc_le_eq less_imp_le enum_def intro: upd_space) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1957 | qed } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1958 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1959 | by (auto simp: s_eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1960 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1961 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1962 | lemma inj_upd: "inj_on upd {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1963 | using upd by (simp add: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1964 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1965 | lemma inj_enum: "inj_on enum {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1966 | proof - | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1967 |   { fix x y :: nat assume "x \<noteq> y" "x \<le> n" "y \<le> n"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1968 |     with upd have "upd ` {..< x} \<noteq> upd ` {..< y}"
 | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 1969 |       by (subst inj_on_image_eq_iff[where C="{..< n}"]) (auto simp: bij_betw_def)
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1970 | then have "enum x \<noteq> enum y" | 
| 68022 | 1971 | by (auto simp: enum_def fun_eq_iff) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1972 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1973 | by (auto simp: inj_on_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1974 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1975 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1976 | lemma enum_0: "enum 0 = base" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1977 | by (simp add: enum_def[abs_def]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1978 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1979 | lemma base_in_s: "base \<in> s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1980 | unfolding s_eq by (subst enum_0[symmetric]) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1981 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1982 | lemma enum_in: "i \<le> n \<Longrightarrow> enum i \<in> s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1983 | unfolding s_eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1984 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1985 | lemma one_step: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1986 | assumes a: "a \<in> s" "j < n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1987 | assumes *: "\<And>a'. a' \<in> s \<Longrightarrow> a' \<noteq> a \<Longrightarrow> a' j = p'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1988 | shows "a j \<noteq> p'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1989 | proof | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1990 | assume "a j = p'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1991 | with * a have "\<And>a'. a' \<in> s \<Longrightarrow> a' j = p'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1992 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1993 | then have "\<And>i. i \<le> n \<Longrightarrow> enum i j = p'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1994 | unfolding s_eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1995 |   from this[of 0] this[of n] have "j \<notin> upd ` {..< n}"
 | 
| 62390 | 1996 | by (auto simp: enum_def fun_eq_iff split: if_split_asm) | 
| 60420 | 1997 | with upd \<open>j < n\<close> show False | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1998 | by (auto simp: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1999 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2000 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2001 | lemma upd_inj: "i < n \<Longrightarrow> j < n \<Longrightarrow> upd i = upd j \<longleftrightarrow> i = j" | 
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61284diff
changeset | 2002 | using upd by (auto simp: bij_betw_def inj_on_eq_iff) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2003 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2004 | lemma upd_surj: "upd ` {..< n} = {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2005 | using upd by (auto simp: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2006 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2007 | lemma in_upd_image: "A \<subseteq> {..< n} \<Longrightarrow> i < n \<Longrightarrow> upd i \<in> upd ` A \<longleftrightarrow> i \<in> A"
 | 
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61284diff
changeset | 2008 |   using inj_on_image_mem_iff[of upd "{..< n}"] upd
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2009 | by (auto simp: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2010 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2011 | lemma enum_inj: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i = enum j \<longleftrightarrow> i = j" | 
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61284diff
changeset | 2012 | using inj_enum by (auto simp: inj_on_eq_iff) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2013 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2014 | lemma in_enum_image: "A \<subseteq> {.. n} \<Longrightarrow> i \<le> n \<Longrightarrow> enum i \<in> enum ` A \<longleftrightarrow> i \<in> A"
 | 
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61284diff
changeset | 2015 | using inj_on_image_mem_iff[OF inj_enum] by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2016 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2017 | lemma enum_mono: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i \<le> enum j \<longleftrightarrow> i \<le> j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2018 | by (auto simp: enum_def le_fun_def in_upd_image Ball_def[symmetric]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2019 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2020 | lemma enum_strict_mono: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i < enum j \<longleftrightarrow> i < j" | 
| 68022 | 2021 | using enum_mono[of i j] enum_inj[of i j] by (auto simp: le_less) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2022 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2023 | lemma chain: "a \<in> s \<Longrightarrow> b \<in> s \<Longrightarrow> a \<le> b \<or> b \<le> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2024 | by (auto simp: s_eq enum_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2025 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2026 | lemma less: "a \<in> s \<Longrightarrow> b \<in> s \<Longrightarrow> a i < b i \<Longrightarrow> a < b" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2027 | using chain[of a b] by (auto simp: less_fun_def le_fun_def not_le[symmetric]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2028 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2029 | lemma enum_0_bot: "a \<in> s \<Longrightarrow> a = enum 0 \<longleftrightarrow> (\<forall>a'\<in>s. a \<le> a')" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2030 | unfolding s_eq by (auto simp: enum_mono Ball_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2031 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2032 | lemma enum_n_top: "a \<in> s \<Longrightarrow> a = enum n \<longleftrightarrow> (\<forall>a'\<in>s. a' \<le> a)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2033 | unfolding s_eq by (auto simp: enum_mono Ball_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2034 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2035 | lemma enum_Suc: "i < n \<Longrightarrow> enum (Suc i) = (enum i)(upd i := Suc (enum i (upd i)))" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2036 | by (auto simp: fun_eq_iff enum_def upd_inj) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2037 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2038 | lemma enum_eq_p: "i \<le> n \<Longrightarrow> n \<le> j \<Longrightarrow> enum i j = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2039 | by (induct i) (auto simp: enum_Suc enum_0 base_out upd_space not_less[symmetric]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2040 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2041 | lemma out_eq_p: "a \<in> s \<Longrightarrow> n \<le> j \<Longrightarrow> a j = p" | 
| 68022 | 2042 | unfolding s_eq by (auto simp: enum_eq_p) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2043 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2044 | lemma s_le_p: "a \<in> s \<Longrightarrow> a j \<le> p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2045 | using out_eq_p[of a j] s_space by (cases "j < n") auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2046 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2047 | lemma le_Suc_base: "a \<in> s \<Longrightarrow> a j \<le> Suc (base j)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2048 | unfolding s_eq by (auto simp: enum_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2049 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2050 | lemma base_le: "a \<in> s \<Longrightarrow> base j \<le> a j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2051 | unfolding s_eq by (auto simp: enum_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2052 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2053 | lemma enum_le_p: "i \<le> n \<Longrightarrow> j < n \<Longrightarrow> enum i j \<le> p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2054 | using enum_in[of i] s_space by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2055 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2056 | lemma enum_less: "a \<in> s \<Longrightarrow> i < n \<Longrightarrow> enum i < a \<longleftrightarrow> enum (Suc i) \<le> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2057 | unfolding s_eq by (auto simp: enum_strict_mono enum_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2058 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2059 | lemma ksimplex_0: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2060 |   "n = 0 \<Longrightarrow> s = {(\<lambda>x. p)}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2061 | using s_eq enum_def base_out by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2062 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2063 | lemma replace_0: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2064 |   assumes "j < n" "a \<in> s" and p: "\<forall>x\<in>s - {a}. x j = 0" and "x \<in> s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2065 | shows "x \<le> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2066 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2067 | assume "x \<noteq> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2068 | have "a j \<noteq> 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2069 | using assms by (intro one_step[where a=a]) auto | 
| 60420 | 2070 | 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> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2071 | show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2072 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2073 | qed simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2074 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2075 | lemma replace_1: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2076 |   assumes "j < n" "a \<in> s" and p: "\<forall>x\<in>s - {a}. x j = p" and "x \<in> s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2077 | shows "a \<le> x" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2078 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2079 | assume "x \<noteq> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2080 | have "a j \<noteq> p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2081 | using assms by (intro one_step[where a=a]) auto | 
| 60420 | 2082 | with enum_le_p[of _ j] \<open>j < n\<close> \<open>a\<in>s\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2083 | have "a j < p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2084 | by (auto simp: less_le s_eq) | 
| 60420 | 2085 | 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> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2086 | show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2087 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2088 | qed simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2089 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2090 | end | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2091 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2092 | locale kuhn_simplex_pair = s: kuhn_simplex p n b_s u_s s + t: kuhn_simplex p n b_t u_t t | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2093 | for p n b_s u_s s b_t u_t t | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2094 | begin | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2095 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2096 | lemma enum_eq: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2097 | assumes l: "i \<le> l" "l \<le> j" and "j + d \<le> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2098 |   assumes eq: "s.enum ` {i .. j} = t.enum ` {i + d .. j + d}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2099 | shows "s.enum l = t.enum (l + d)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2100 | using l proof (induct l rule: dec_induct) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2101 | case base | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2102 |   then have s: "s.enum i \<in> t.enum ` {i + d .. j + d}" and t: "t.enum (i + d) \<in> s.enum ` {i .. j}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2103 | using eq by auto | 
| 60420 | 2104 | from t \<open>i \<le> j\<close> \<open>j + d \<le> n\<close> have "s.enum i \<le> t.enum (i + d)" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2105 | by (auto simp: s.enum_mono) | 
| 60420 | 2106 | moreover from s \<open>i \<le> j\<close> \<open>j + d \<le> n\<close> have "t.enum (i + d) \<le> s.enum i" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2107 | by (auto simp: t.enum_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2108 | ultimately show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2109 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2110 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2111 | case (step l) | 
| 60420 | 2112 | moreover from step.prems \<open>j + d \<le> n\<close> have | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2113 | "s.enum l < s.enum (Suc l)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2114 | "t.enum (l + d) < t.enum (Suc l + d)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2115 | by (simp_all add: s.enum_strict_mono t.enum_strict_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2116 | moreover have | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2117 |       "s.enum (Suc l) \<in> t.enum ` {i + d .. j + d}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2118 |       "t.enum (Suc l + d) \<in> s.enum ` {i .. j}"
 | 
| 60420 | 2119 | using step \<open>j + d \<le> n\<close> eq by (auto simp: s.enum_inj t.enum_inj) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2120 | ultimately have "s.enum (Suc l) = t.enum (Suc (l + d))" | 
| 60420 | 2121 | using \<open>j + d \<le> n\<close> | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 2122 | by (intro antisym s.enum_less[THEN iffD1] t.enum_less[THEN iffD1]) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2123 | (auto intro!: s.enum_in t.enum_in) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2124 | then show ?case by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2125 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2126 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2127 | lemma ksimplex_eq_bot: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2128 | assumes a: "a \<in> s" "\<And>a'. a' \<in> s \<Longrightarrow> a \<le> a'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2129 | assumes b: "b \<in> t" "\<And>b'. b' \<in> t \<Longrightarrow> b \<le> b'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2130 |   assumes eq: "s - {a} = t - {b}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2131 | shows "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2132 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2133 | assume "n = 0" with s.ksimplex_0 t.ksimplex_0 show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2134 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2135 | assume "n \<noteq> 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2136 | have "s.enum 0 = (s.enum (Suc 0)) (u_s 0 := s.enum (Suc 0) (u_s 0) - 1)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2137 | "t.enum 0 = (t.enum (Suc 0)) (u_t 0 := t.enum (Suc 0) (u_t 0) - 1)" | 
| 60420 | 2138 | using \<open>n \<noteq> 0\<close> by (simp_all add: s.enum_Suc t.enum_Suc) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2139 | moreover have e0: "a = s.enum 0" "b = t.enum 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2140 | using a b by (simp_all add: s.enum_0_bot t.enum_0_bot) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2141 | moreover | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 2142 |   { fix j assume "0 < j" "j \<le> n"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2143 |     moreover have "s - {a} = s.enum ` {Suc 0 .. n}" "t - {b} = t.enum ` {Suc 0 .. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2144 | unfolding s.s_eq t.s_eq e0 by (auto simp: s.enum_inj t.enum_inj) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2145 | ultimately have "s.enum j = t.enum j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2146 | using enum_eq[of "1" j n 0] eq by auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2147 | note enum_eq = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2148 | then have "s.enum (Suc 0) = t.enum (Suc 0)" | 
| 60420 | 2149 | using \<open>n \<noteq> 0\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2150 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2151 |   { fix j assume "Suc j < n"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2152 | with enum_eq[of "Suc j"] enum_eq[of "Suc (Suc j)"] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2153 | have "u_s (Suc j) = u_t (Suc j)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2154 | using s.enum_Suc[of "Suc j"] t.enum_Suc[of "Suc j"] | 
| 62390 | 2155 | by (auto simp: fun_eq_iff split: if_split_asm) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2156 | then have "\<And>j. 0 < j \<Longrightarrow> j < n \<Longrightarrow> u_s j = u_t j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2157 | by (auto simp: gr0_conv_Suc) | 
| 60420 | 2158 | with \<open>n \<noteq> 0\<close> have "u_t 0 = u_s 0" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2159 | by (intro bij_betw_singleton_eq[OF t.upd s.upd, of 0]) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2160 | ultimately have "a = b" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2161 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2162 | with assms show "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2163 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2164 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2165 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2166 | lemma ksimplex_eq_top: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2167 | assumes a: "a \<in> s" "\<And>a'. a' \<in> s \<Longrightarrow> a' \<le> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2168 | assumes b: "b \<in> t" "\<And>b'. b' \<in> t \<Longrightarrow> b' \<le> b" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2169 |   assumes eq: "s - {a} = t - {b}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2170 | shows "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2171 | proof (cases n) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2172 | assume "n = 0" with s.ksimplex_0 t.ksimplex_0 show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2173 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2174 | case (Suc n') | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2175 | have "s.enum n = (s.enum n') (u_s n' := Suc (s.enum n' (u_s n')))" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2176 | "t.enum n = (t.enum n') (u_t n' := Suc (t.enum n' (u_t n')))" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2177 | using Suc by (simp_all add: s.enum_Suc t.enum_Suc) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2178 | moreover have en: "a = s.enum n" "b = t.enum n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2179 | using a b by (simp_all add: s.enum_n_top t.enum_n_top) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2180 | moreover | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 2181 |   { fix j assume "j < n"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2182 |     moreover have "s - {a} = s.enum ` {0 .. n'}" "t - {b} = t.enum ` {0 .. n'}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2183 | unfolding s.s_eq t.s_eq en by (auto simp: s.enum_inj t.enum_inj Suc) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2184 | ultimately have "s.enum j = t.enum j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2185 | using enum_eq[of "0" j n' 0] eq Suc by auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2186 | note enum_eq = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2187 | then have "s.enum n' = t.enum n'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2188 | using Suc by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2189 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2190 |   { fix j assume "j < n'"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2191 | with enum_eq[of j] enum_eq[of "Suc j"] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2192 | have "u_s j = u_t j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2193 | using s.enum_Suc[of j] t.enum_Suc[of j] | 
| 62390 | 2194 | by (auto simp: Suc fun_eq_iff split: if_split_asm) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2195 | then have "\<And>j. j < n' \<Longrightarrow> u_s j = u_t j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2196 | by (auto simp: gr0_conv_Suc) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2197 | then have "u_t n' = u_s n'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2198 | by (intro bij_betw_singleton_eq[OF t.upd s.upd, of n']) (auto simp: Suc) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2199 | ultimately have "a = b" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2200 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2201 | with assms show "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2202 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2203 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2204 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2205 | end | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2206 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2207 | inductive ksimplex for p n :: nat where | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2208 | ksimplex: "kuhn_simplex p n base upd s \<Longrightarrow> ksimplex p n s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2209 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2210 | lemma finite_ksimplexes: "finite {s. ksimplex p n s}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2211 | proof (rule finite_subset) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2212 |   { fix a s assume "ksimplex p n s" "a \<in> s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2213 | then obtain b u where "kuhn_simplex p n b u s" by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2214 | then interpret kuhn_simplex p n b u s . | 
| 60420 | 2215 | from s_space \<open>a \<in> s\<close> out_eq_p[OF \<open>a \<in> s\<close>] | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2216 |     have "a \<in> (\<lambda>f x. if n \<le> x then p else f x) ` ({..< n} \<rightarrow>\<^sub>E {.. p})"
 | 
| 62390 | 2217 | by (auto simp: image_iff subset_eq Pi_iff split: if_split_asm | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2218 |                intro!: bexI[of _ "restrict a {..< n}"]) }
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2219 |   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}))"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2220 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2221 | qed (simp add: finite_PiE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2222 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2223 | lemma ksimplex_card: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2224 | assumes "ksimplex p n s" shows "card s = Suc n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2225 | using assms proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2226 | case (ksimplex u b) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2227 | then interpret kuhn_simplex p n u b s . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2228 | show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2229 | by (simp add: card_image s_eq inj_enum) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2230 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2231 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2232 | lemma simplex_top_face: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2233 | assumes "0 < p" "\<forall>x\<in>s'. x n = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2234 |   shows "ksimplex p n s' \<longleftrightarrow> (\<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> s' = s - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2235 | using assms | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2236 | proof safe | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2237 |   fix s a assume "ksimplex p (Suc n) s" and a: "a \<in> s" and na: "\<forall>x\<in>s - {a}. x n = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2238 |   then show "ksimplex p n (s - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2239 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2240 | case (ksimplex base upd) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2241 | then interpret kuhn_simplex p "Suc n" base upd "s" . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2242 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2243 | have "a n < p" | 
| 60420 | 2244 | using one_step[of a n p] na \<open>a\<in>s\<close> s_space by (auto simp: less_le) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2245 | then have "a = enum 0" | 
| 60420 | 2246 | using \<open>a \<in> s\<close> na by (subst enum_0_bot) (auto simp: le_less intro!: less[of a _ n]) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2247 |     then have s_eq: "s - {a} = enum ` Suc ` {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2248 | using s_eq by (simp add: atMost_Suc_eq_insert_0 insert_ident Zero_notin_Suc in_enum_image subset_eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2249 |     then have "enum 1 \<in> s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2250 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2251 | then have "upd 0 = n" | 
| 60420 | 2252 | using \<open>a n < p\<close> \<open>a = enum 0\<close> na[rule_format, of "enum 1"] | 
| 62390 | 2253 | by (auto simp: fun_eq_iff enum_Suc split: if_split_asm) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2254 |     then have "bij_betw upd (Suc ` {..< n}) {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2255 | using upd | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2256 | by (subst notIn_Un_bij_betw3[where b=0]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2257 | (auto simp: lessThan_Suc[symmetric] lessThan_Suc_eq_insert_0) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2258 |     then have "bij_betw (upd\<circ>Suc) {..<n} {..<n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2259 | by (rule bij_betw_trans[rotated]) (auto simp: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2260 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2261 | have "a n = p - 1" | 
| 60420 | 2262 |       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>)
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2263 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2264 | show ?thesis | 
| 61169 | 2265 | proof (rule ksimplex.intros, standard) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2266 |       show "bij_betw (upd\<circ>Suc) {..< n} {..< n}" by fact
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2267 |       show "base(n := p) \<in> {..<n} \<rightarrow> {..<p}" "\<And>i. n\<le>i \<Longrightarrow> (base(n := p)) i = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2268 | using base base_out by (auto simp: Pi_iff) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2269 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2270 |       have "\<And>i. Suc ` {..< i} = {..< Suc i} - {0}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2271 | by (auto simp: image_iff Ball_def) arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2272 |       then have upd_Suc: "\<And>i. i \<le> n \<Longrightarrow> (upd\<circ>Suc) ` {..< i} = upd ` {..< Suc i} - {n}"
 | 
| 60420 | 2273 | using \<open>upd 0 = n\<close> upd_inj | 
| 68022 | 2274 | by (auto simp: image_comp[symmetric] inj_on_image_set_diff[OF inj_upd]) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2275 |       have n_in_upd: "\<And>i. n \<in> upd ` {..< Suc i}"
 | 
| 60420 | 2276 | using \<open>upd 0 = n\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2277 | |
| 63040 | 2278 | define f' where "f' i j = | 
| 2279 |         (if j \<in> (upd\<circ>Suc)`{..< i} then Suc ((base(n := p)) j) else (base(n := p)) j)" for i j
 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2280 |       { fix x i assume i[arith]: "i \<le> n" then have "enum (Suc i) x = f' i x"
 | 
| 60420 | 2281 | 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> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2282 | by (simp add: upd_Suc enum_0 n_in_upd) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2283 |       then show "s - {a} = f' ` {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2284 | unfolding s_eq image_comp by (intro image_cong) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2285 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2286 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2287 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2288 | assume "ksimplex p n s'" and *: "\<forall>x\<in>s'. x n = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2289 |   then show "\<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> s' = s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2290 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2291 | case (ksimplex base upd) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2292 | then interpret kuhn_simplex p n base upd s' . | 
| 63040 | 2293 | define b where "b = base (n := p - 1)" | 
| 2294 | define u where "u i = (case i of 0 \<Rightarrow> n | Suc i \<Rightarrow> upd i)" for i | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2295 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2296 |     have "ksimplex p (Suc n) (s' \<union> {b})"
 | 
| 61169 | 2297 | proof (rule ksimplex.intros, standard) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2298 |       show "b \<in> {..<Suc n} \<rightarrow> {..<p}"
 | 
| 60420 | 2299 | using base \<open>0 < p\<close> unfolding lessThan_Suc b_def by (auto simp: PiE_iff) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2300 | show "\<And>i. Suc n \<le> i \<Longrightarrow> b i = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2301 | using base_out by (auto simp: b_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2302 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2303 |       have "bij_betw u (Suc ` {..< n} \<union> {0}) ({..<n} \<union> {u 0})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2304 | using upd | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2305 | by (intro notIn_Un_bij_betw) (auto simp: u_def bij_betw_def image_comp comp_def inj_on_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2306 |       then show "bij_betw u {..<Suc n} {..<Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2307 | by (simp add: u_def lessThan_Suc[symmetric] lessThan_Suc_eq_insert_0) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2308 | |
| 63040 | 2309 |       define f' where "f' i j = (if j \<in> u`{..< i} then Suc (b j) else b j)" for i j
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2310 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2311 |       have u_eq: "\<And>i. i \<le> n \<Longrightarrow> u ` {..< Suc i} = upd ` {..< i} \<union> { n }"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2312 | by (auto simp: u_def image_iff upd_inj Ball_def split: nat.split) arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2313 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2314 |       { fix x have "x \<le> n \<Longrightarrow> n \<notin> upd ` {..<x}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2315 | using upd_space by (simp add: image_iff neq_iff) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2316 | note n_not_upd = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2317 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2318 |       have *: "f' ` {.. Suc n} = f' ` (Suc ` {.. n} \<union> {0})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2319 | unfolding atMost_Suc_eq_insert_0 by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2320 |       also have "\<dots> = (f' \<circ> Suc) ` {.. n} \<union> {b}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2321 | by (auto simp: f'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2322 |       also have "(f' \<circ> Suc) ` {.. n} = s'"
 | 
| 60420 | 2323 | using \<open>0 < p\<close> base_out[of n] | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2324 | unfolding s_eq enum_def[abs_def] f'_def[abs_def] upd_space | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2325 | by (intro image_cong) (simp_all add: u_eq b_def fun_eq_iff n_not_upd) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2326 |       finally show "s' \<union> {b} = f' ` {.. Suc n}" ..
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2327 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2328 | moreover have "b \<notin> s'" | 
| 60420 | 2329 | using * \<open>0 < p\<close> by (auto simp: b_def) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2330 | ultimately show ?thesis by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2331 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2332 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2333 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2334 | lemma ksimplex_replace_0: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2335 | assumes s: "ksimplex p n s" and a: "a \<in> s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2336 |   assumes j: "j < n" and p: "\<forall>x\<in>s - {a}. x j = 0"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2337 |   shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 1"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2338 | using s | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2339 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2340 | case (ksimplex b_s u_s) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2341 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 2342 |   { fix t b assume "ksimplex p n t"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2343 | then obtain b_t u_t where "kuhn_simplex p n b_t u_t t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2344 | by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2345 | interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2346 | by intro_locales fact+ | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2347 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2348 |     assume b: "b \<in> t" "t - {b} = s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2349 | with a j p s.replace_0[of _ a] t.replace_0[of _ b] have "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2350 | by (intro ksimplex_eq_top[of a b]) auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2351 |   then have "{s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = {s}"
 | 
| 60420 | 2352 | using s \<open>a \<in> s\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2353 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2354 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2355 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2356 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2357 | lemma ksimplex_replace_1: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2358 | assumes s: "ksimplex p n s" and a: "a \<in> s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2359 |   assumes j: "j < n" and p: "\<forall>x\<in>s - {a}. x j = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2360 |   shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 1"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2361 | using s | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2362 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2363 | case (ksimplex b_s u_s) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2364 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 2365 |   { fix t b assume "ksimplex p n t"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2366 | then obtain b_t u_t where "kuhn_simplex p n b_t u_t t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2367 | by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2368 | interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2369 | by intro_locales fact+ | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2370 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2371 |     assume b: "b \<in> t" "t - {b} = s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2372 | with a j p s.replace_1[of _ a] t.replace_1[of _ b] have "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2373 | by (intro ksimplex_eq_bot[of a b]) auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2374 |   then have "{s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = {s}"
 | 
| 60420 | 2375 | using s \<open>a \<in> s\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2376 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2377 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2378 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2379 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2380 | 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))" | 
| 68022 | 2381 | by (auto simp: card_Suc_eq eval_nat_numeral) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2382 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2383 | lemma ksimplex_replace_2: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2384 | assumes s: "ksimplex p n s" and "a \<in> s" and "n \<noteq> 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2385 |     and lb: "\<forall>j<n. \<exists>x\<in>s - {a}. x j \<noteq> 0"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2386 |     and ub: "\<forall>j<n. \<exists>x\<in>s - {a}. x j \<noteq> p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2387 |   shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 2"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2388 | using s | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2389 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2390 | case (ksimplex base upd) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2391 | then interpret kuhn_simplex p n base upd s . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2392 | |
| 60420 | 2393 | from \<open>a \<in> s\<close> obtain i where "i \<le> n" "a = enum i" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2394 | unfolding s_eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2395 | |
| 60420 | 2396 | from \<open>i \<le> n\<close> have "i = 0 \<or> i = n \<or> (0 < i \<and> i < n)" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2397 | by linarith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2398 |   then have "\<exists>!s'. s' \<noteq> s \<and> ksimplex p n s' \<and> (\<exists>b\<in>s'. s - {a} = s'- {b})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2399 | proof (elim disjE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2400 | assume "i = 0" | 
| 63040 | 2401 | define rot where [abs_def]: "rot i = (if i + 1 = n then 0 else i + 1)" for i | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2402 | let ?upd = "upd \<circ> rot" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2403 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2404 |     have rot: "bij_betw rot {..< n} {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2405 | by (auto simp: bij_betw_def inj_on_def image_iff Ball_def rot_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2406 | arith+ | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2407 |     from rot upd have "bij_betw ?upd {..<n} {..<n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2408 | by (rule bij_betw_trans) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2409 | |
| 63040 | 2410 | define f' where [abs_def]: "f' i j = | 
| 2411 |       (if j \<in> ?upd`{..< i} then Suc (enum (Suc 0) j) else enum (Suc 0) j)" for i j
 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2412 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2413 |     interpret b: kuhn_simplex p n "enum (Suc 0)" "upd \<circ> rot" "f' ` {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2414 | proof | 
| 60420 | 2415 | from \<open>a = enum i\<close> ub \<open>n \<noteq> 0\<close> \<open>i = 0\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2416 | obtain i' where "i' \<le> n" "enum i' \<noteq> enum 0" "enum i' (upd 0) \<noteq> p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2417 | unfolding s_eq by (auto intro: upd_space simp: enum_inj) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2418 | then have "enum 1 \<le> enum i'" "enum i' (upd 0) < p" | 
| 68022 | 2419 | using enum_le_p[of i' "upd 0"] by (auto simp: enum_inj enum_mono upd_space) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2420 | then have "enum 1 (upd 0) < p" | 
| 68022 | 2421 | by (auto simp: le_fun_def intro: le_less_trans) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2422 |       then show "enum (Suc 0) \<in> {..<n} \<rightarrow> {..<p}"
 | 
| 68022 | 2423 | using base \<open>n \<noteq> 0\<close> by (auto simp: enum_0 enum_Suc PiE_iff extensional_def upd_space) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2424 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2425 |       { fix i assume "n \<le> i" then show "enum (Suc 0) i = p"
 | 
| 60420 | 2426 | using \<open>n \<noteq> 0\<close> by (auto simp: enum_eq_p) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2427 |       show "bij_betw ?upd {..<n} {..<n}" by fact
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2428 | qed (simp add: f'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2429 |     have ks_f': "ksimplex p n (f' ` {.. n})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2430 | by rule unfold_locales | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2431 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2432 | have b_enum: "b.enum = f'" unfolding f'_def b.enum_def[abs_def] .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2433 |     with b.inj_enum have inj_f': "inj_on f' {.. n}" by simp
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2434 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2435 |     have [simp]: "\<And>j. j < n \<Longrightarrow> rot ` {..< j} = {0 <..< Suc j}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2436 | by (auto simp: rot_def image_iff Ball_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2437 | arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2438 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2439 |     { fix j assume j: "j < n"
 | 
| 60420 | 2440 | from j \<open>n \<noteq> 0\<close> have "f' j = enum (Suc j)" | 
| 68022 | 2441 | by (auto simp: f'_def enum_def upd_inj in_upd_image image_comp[symmetric] fun_eq_iff) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2442 | note f'_eq_enum = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2443 |     then have "enum ` Suc ` {..< n} = f' ` {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2444 | by (force simp: enum_inj) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2445 |     also have "Suc ` {..< n} = {.. n} - {0}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2446 | by (auto simp: image_iff Ball_def) arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2447 |     also have "{..< n} = {.. n} - {n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2448 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2449 |     finally have eq: "s - {a} = f' ` {.. n} - {f' n}"
 | 
| 60420 | 2450 | unfolding s_eq \<open>a = enum i\<close> \<open>i = 0\<close> | 
| 60303 | 2451 | by (simp add: Diff_subset inj_on_image_set_diff[OF inj_enum] inj_on_image_set_diff[OF inj_f']) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2452 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2453 | have "enum 0 < f' 0" | 
| 60420 | 2454 | using \<open>n \<noteq> 0\<close> by (simp add: enum_strict_mono f'_eq_enum) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2455 | also have "\<dots> < f' n" | 
| 60420 | 2456 | using \<open>n \<noteq> 0\<close> b.enum_strict_mono[of 0 n] unfolding b_enum by simp | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2457 | finally have "a \<noteq> f' n" | 
| 60420 | 2458 | using \<open>a = enum i\<close> \<open>i = 0\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2459 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2460 |     { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2461 | obtain b u where "kuhn_simplex p n b u t" | 
| 60420 | 2462 | using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2463 | then interpret t: kuhn_simplex p n b u t . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2464 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2465 |       { fix x assume "x \<in> s" "x \<noteq> a"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2466 | then have "x (upd 0) = enum (Suc 0) (upd 0)" | 
| 60420 | 2467 | by (auto simp: \<open>a = enum i\<close> \<open>i = 0\<close> s_eq enum_def enum_inj) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2468 |       then have eq_upd0: "\<forall>x\<in>t-{c}. x (upd 0) = enum (Suc 0) (upd 0)"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2469 | unfolding eq_sma[symmetric] by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2470 | then have "c (upd 0) \<noteq> enum (Suc 0) (upd 0)" | 
| 60420 | 2471 | using \<open>n \<noteq> 0\<close> by (intro t.one_step[OF \<open>c\<in>t\<close> ]) (auto simp: upd_space) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2472 | then have "c (upd 0) < enum (Suc 0) (upd 0) \<or> c (upd 0) > enum (Suc 0) (upd 0)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2473 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2474 |       then have "t = s \<or> t = f' ` {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2475 | proof (elim disjE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2476 | assume *: "c (upd 0) < enum (Suc 0) (upd 0)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2477 | interpret st: kuhn_simplex_pair p n base upd s b u t .. | 
| 60420 | 2478 |         { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "c \<le> x"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2479 | by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2480 | note top = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2481 | have "s = t" | 
| 60420 | 2482 | using \<open>a = enum i\<close> \<open>i = 0\<close> \<open>c \<in> t\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2483 | by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq_sma]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2484 | (auto simp: s_eq enum_mono t.s_eq t.enum_mono top) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2485 | then show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2486 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2487 | assume *: "c (upd 0) > enum (Suc 0) (upd 0)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2488 |         interpret st: kuhn_simplex_pair p n "enum (Suc 0)" "upd \<circ> rot" "f' ` {.. n}" b u t ..
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2489 |         have eq: "f' ` {..n} - {f' n} = t - {c}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2490 | using eq_sma eq by simp | 
| 60420 | 2491 |         { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "x \<le> c"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2492 | by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2493 | note top = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2494 |         have "f' ` {..n} = t"
 | 
| 60420 | 2495 | using \<open>a = enum i\<close> \<open>i = 0\<close> \<open>c \<in> t\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2496 | by (intro st.ksimplex_eq_top[OF _ _ _ _ eq]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2497 | (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono b_enum[symmetric] top) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2498 | then show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2499 | qed } | 
| 60420 | 2500 | with ks_f' eq \<open>a \<noteq> f' n\<close> \<open>n \<noteq> 0\<close> show ?thesis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2501 |       apply (intro ex1I[of _ "f' ` {.. n}"])
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2502 | apply auto [] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2503 | apply metis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2504 | done | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2505 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2506 | assume "i = n" | 
| 60420 | 2507 | from \<open>n \<noteq> 0\<close> obtain n' where n': "n = Suc n'" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2508 | by (cases n) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2509 | |
| 63040 | 2510 | define rot where "rot i = (case i of 0 \<Rightarrow> n' | Suc i \<Rightarrow> i)" for i | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2511 | let ?upd = "upd \<circ> rot" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2512 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2513 |     have rot: "bij_betw rot {..< n} {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2514 | by (auto simp: bij_betw_def inj_on_def image_iff Bex_def rot_def n' split: nat.splits) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2515 | arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2516 |     from rot upd have "bij_betw ?upd {..<n} {..<n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2517 | by (rule bij_betw_trans) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2518 | |
| 63040 | 2519 | define b where "b = base (upd n' := base (upd n') - 1)" | 
| 2520 |     define f' where [abs_def]: "f' i j = (if j \<in> ?upd`{..< i} then Suc (b j) else b j)" for i j
 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2521 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2522 |     interpret b: kuhn_simplex p n b "upd \<circ> rot" "f' ` {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2523 | proof | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2524 |       { fix i assume "n \<le> i" then show "b i = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2525 | using base_out[of i] upd_space[of n'] by (auto simp: b_def n') } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2526 |       show "b \<in> {..<n} \<rightarrow> {..<p}"
 | 
| 60420 | 2527 | using base \<open>n \<noteq> 0\<close> upd_space[of n'] | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2528 | by (auto simp: b_def PiE_def Pi_iff Ball_def upd_space extensional_def n') | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2529 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2530 |       show "bij_betw ?upd {..<n} {..<n}" by fact
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2531 | qed (simp add: f'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2532 | have f': "b.enum = f'" unfolding f'_def b.enum_def[abs_def] .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2533 |     have ks_f': "ksimplex p n (b.enum ` {.. n})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2534 | unfolding f' by rule unfold_locales | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2535 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 2536 | have "0 < n" | 
| 60420 | 2537 | using \<open>n \<noteq> 0\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2538 | |
| 60420 | 2539 |     { from \<open>a = enum i\<close> \<open>n \<noteq> 0\<close> \<open>i = n\<close> lb upd_space[of n']
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2540 | obtain i' where "i' \<le> n" "enum i' \<noteq> enum n" "0 < enum i' (upd n')" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2541 | unfolding s_eq by (auto simp: enum_inj n') | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2542 | moreover have "enum i' (upd n') = base (upd n')" | 
| 60420 | 2543 | unfolding enum_def using \<open>i' \<le> n\<close> \<open>enum i' \<noteq> enum n\<close> by (auto simp: n' upd_inj enum_inj) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2544 | ultimately have "0 < base (upd n')" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2545 | by auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2546 | then have benum1: "b.enum (Suc 0) = base" | 
| 60420 | 2547 | unfolding b.enum_Suc[OF \<open>0<n\<close>] b.enum_0 by (auto simp: b_def rot_def) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2548 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2549 |     have [simp]: "\<And>j. Suc j < n \<Longrightarrow> rot ` {..< Suc j} = {n'} \<union> {..< j}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2550 | by (auto simp: rot_def image_iff Ball_def split: nat.splits) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2551 | have rot_simps: "\<And>j. rot (Suc j) = j" "rot 0 = n'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2552 | by (simp_all add: rot_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2553 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2554 |     { fix j assume j: "Suc j \<le> n" then have "b.enum (Suc j) = enum j"
 | 
| 68022 | 2555 | by (induct j) (auto simp: benum1 enum_0 b.enum_Suc enum_Suc rot_simps) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2556 | note b_enum_eq_enum = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2557 |     then have "enum ` {..< n} = b.enum ` Suc ` {..< n}"
 | 
| 68022 | 2558 | by (auto simp: image_comp intro!: image_cong) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2559 |     also have "Suc ` {..< n} = {.. n} - {0}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2560 | by (auto simp: image_iff Ball_def) arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2561 |     also have "{..< n} = {.. n} - {n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2562 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2563 |     finally have eq: "s - {a} = b.enum ` {.. n} - {b.enum 0}"
 | 
| 60420 | 2564 | unfolding s_eq \<open>a = enum i\<close> \<open>i = n\<close> | 
| 60303 | 2565 |       using inj_on_image_set_diff[OF inj_enum Diff_subset, of "{n}"]
 | 
| 2566 |             inj_on_image_set_diff[OF b.inj_enum Diff_subset, of "{0}"]
 | |
| 68022 | 2567 | by (simp add: comp_def) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2568 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2569 | have "b.enum 0 \<le> b.enum n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2570 | by (simp add: b.enum_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2571 | also have "b.enum n < enum n" | 
| 60420 | 2572 | using \<open>n \<noteq> 0\<close> by (simp add: enum_strict_mono b_enum_eq_enum n') | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2573 | finally have "a \<noteq> b.enum 0" | 
| 60420 | 2574 | using \<open>a = enum i\<close> \<open>i = n\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2575 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2576 |     { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2577 | obtain b' u where "kuhn_simplex p n b' u t" | 
| 60420 | 2578 | using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2579 | then interpret t: kuhn_simplex p n b' u t . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2580 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2581 |       { fix x assume "x \<in> s" "x \<noteq> a"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2582 | then have "x (upd n') = enum n' (upd n')" | 
| 60420 | 2583 | by (auto simp: \<open>a = enum i\<close> n' \<open>i = n\<close> s_eq enum_def enum_inj in_upd_image) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2584 |       then have eq_upd0: "\<forall>x\<in>t-{c}. x (upd n') = enum n' (upd n')"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2585 | unfolding eq_sma[symmetric] by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2586 | then have "c (upd n') \<noteq> enum n' (upd n')" | 
| 60420 | 2587 | using \<open>n \<noteq> 0\<close> by (intro t.one_step[OF \<open>c\<in>t\<close> ]) (auto simp: n' upd_space[unfolded n']) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2588 | then have "c (upd n') < enum n' (upd n') \<or> c (upd n') > enum n' (upd n')" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2589 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2590 |       then have "t = s \<or> t = b.enum ` {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2591 | proof (elim disjE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2592 | assume *: "c (upd n') > enum n' (upd n')" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2593 | interpret st: kuhn_simplex_pair p n base upd s b' u t .. | 
| 60420 | 2594 |         { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "x \<le> c"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2595 | by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2596 | note top = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2597 | have "s = t" | 
| 60420 | 2598 | using \<open>a = enum i\<close> \<open>i = n\<close> \<open>c \<in> t\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2599 | by (intro st.ksimplex_eq_top[OF _ _ _ _ eq_sma]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2600 | (auto simp: s_eq enum_mono t.s_eq t.enum_mono top) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2601 | then show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2602 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2603 | assume *: "c (upd n') < enum n' (upd n')" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2604 |         interpret st: kuhn_simplex_pair p n b "upd \<circ> rot" "f' ` {.. n}" b' u t ..
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2605 |         have eq: "f' ` {..n} - {b.enum 0} = t - {c}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2606 | using eq_sma eq f' by simp | 
| 60420 | 2607 |         { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "c \<le> x"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2608 | by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2609 | note bot = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2610 |         have "f' ` {..n} = t"
 | 
| 60420 | 2611 | using \<open>a = enum i\<close> \<open>i = n\<close> \<open>c \<in> t\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2612 | by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2613 | (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono bot) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2614 | with f' show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2615 | qed } | 
| 60420 | 2616 | with ks_f' eq \<open>a \<noteq> b.enum 0\<close> \<open>n \<noteq> 0\<close> show ?thesis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2617 |       apply (intro ex1I[of _ "b.enum ` {.. n}"])
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2618 | apply auto [] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2619 | apply metis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2620 | done | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2621 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2622 | assume i: "0 < i" "i < n" | 
| 63040 | 2623 | define i' where "i' = i - 1" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2624 | with i have "Suc i' < n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2625 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2626 | with i have Suc_i': "Suc i' = i" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2627 | by (simp add: i'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2628 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2629 | let ?upd = "Fun.swap i' i upd" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2630 |     from i upd have "bij_betw ?upd {..< n} {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2631 | by (subst bij_betw_swap_iff) (auto simp: i'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2632 | |
| 63040 | 2633 |     define f' where [abs_def]: "f' i j = (if j \<in> ?upd`{..< i} then Suc (base j) else base j)"
 | 
| 2634 | for i j | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2635 |     interpret b: kuhn_simplex p n base ?upd "f' ` {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2636 | proof | 
| 67682 
00c436488398
tuned proofs -- prefer explicit names for facts from 'interpret';
 wenzelm parents: 
67673diff
changeset | 2637 |       show "base \<in> {..<n} \<rightarrow> {..<p}" by (rule base)
 | 
| 
00c436488398
tuned proofs -- prefer explicit names for facts from 'interpret';
 wenzelm parents: 
67673diff
changeset | 2638 |       { fix i assume "n \<le> i" then show "base i = p" by (rule base_out) }
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2639 |       show "bij_betw ?upd {..<n} {..<n}" by fact
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2640 | qed (simp add: f'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2641 | have f': "b.enum = f'" unfolding f'_def b.enum_def[abs_def] .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2642 |     have ks_f': "ksimplex p n (b.enum ` {.. n})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2643 | unfolding f' by rule unfold_locales | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2644 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2645 |     have "{i} \<subseteq> {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2646 | using i by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2647 |     { fix j assume "j \<le> n"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2648 | moreover have "j < i \<or> i = j \<or> i < j" by arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2649 | moreover note i | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2650 | ultimately have "enum j = b.enum j \<longleftrightarrow> j \<noteq> i" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2651 | unfolding enum_def[abs_def] b.enum_def[abs_def] | 
| 68022 | 2652 | by (auto simp: fun_eq_iff swap_image i'_def | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2653 | in_upd_image inj_on_image_set_diff[OF inj_upd]) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2654 | note enum_eq_benum = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2655 |     then have "enum ` ({.. n} - {i}) = b.enum ` ({.. n} - {i})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2656 | by (intro image_cong) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2657 |     then have eq: "s - {a} = b.enum ` {.. n} - {b.enum i}"
 | 
| 60420 | 2658 | unfolding s_eq \<open>a = enum i\<close> | 
| 2659 |       using inj_on_image_set_diff[OF inj_enum Diff_subset \<open>{i} \<subseteq> {..n}\<close>]
 | |
| 2660 |             inj_on_image_set_diff[OF b.inj_enum Diff_subset \<open>{i} \<subseteq> {..n}\<close>]
 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2661 | by (simp add: comp_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2662 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2663 | have "a \<noteq> b.enum i" | 
| 60420 | 2664 | using \<open>a = enum i\<close> enum_eq_benum i by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2665 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2666 |     { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2667 | obtain b' u where "kuhn_simplex p n b' u t" | 
| 60420 | 2668 | using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2669 | then interpret t: kuhn_simplex p n b' u t . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2670 |       have "enum i' \<in> s - {a}" "enum (i + 1) \<in> s - {a}"
 | 
| 60420 | 2671 | using \<open>a = enum i\<close> i enum_in by (auto simp: enum_inj i'_def) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2672 | then obtain l k where | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2673 | l: "t.enum l = enum i'" "l \<le> n" "t.enum l \<noteq> c" and | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2674 | k: "t.enum k = enum (i + 1)" "k \<le> n" "t.enum k \<noteq> c" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2675 | unfolding eq_sma by (auto simp: t.s_eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2676 | with i have "t.enum l < t.enum k" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2677 | by (simp add: enum_strict_mono i'_def) | 
| 60420 | 2678 | with \<open>l \<le> n\<close> \<open>k \<le> n\<close> have "l < k" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2679 | by (simp add: t.enum_strict_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2680 |       { assume "Suc l = k"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2681 | have "enum (Suc (Suc i')) = t.enum (Suc l)" | 
| 60420 | 2682 | using i by (simp add: k \<open>Suc l = k\<close> i'_def) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2683 | then have False | 
| 60420 | 2684 | using \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close> | 
| 62390 | 2685 | by (auto simp: t.enum_Suc enum_Suc l upd_inj fun_eq_iff split: if_split_asm) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2686 | (metis Suc_lessD n_not_Suc_n upd_inj) } | 
| 60420 | 2687 | with \<open>l < k\<close> have "Suc l < k" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2688 | by arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2689 | have c_eq: "c = t.enum (Suc l)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2690 | proof (rule ccontr) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2691 | assume "c \<noteq> t.enum (Suc l)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2692 |         then have "t.enum (Suc l) \<in> s - {a}"
 | 
| 60420 | 2693 | using \<open>l < k\<close> \<open>k \<le> n\<close> by (simp add: t.s_eq eq_sma) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2694 | then obtain j where "t.enum (Suc l) = enum j" "j \<le> n" "enum j \<noteq> enum i" | 
| 60420 | 2695 | unfolding s_eq \<open>a = enum i\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2696 | with i have "t.enum (Suc l) \<le> t.enum l \<or> t.enum k \<le> t.enum (Suc l)" | 
| 68022 | 2697 | by (auto simp: i'_def enum_mono enum_inj l k) | 
| 60420 | 2698 | with \<open>Suc l < k\<close> \<open>k \<le> n\<close> show False | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2699 | by (simp add: t.enum_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2700 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2701 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2702 |       { have "t.enum (Suc (Suc l)) \<in> s - {a}"
 | 
| 60420 | 2703 | 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) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2704 | then obtain j where eq: "t.enum (Suc (Suc l)) = enum j" and "j \<le> n" "j \<noteq> i" | 
| 60420 | 2705 | by (auto simp: s_eq \<open>a = enum i\<close>) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2706 | moreover have "enum i' < t.enum (Suc (Suc l))" | 
| 60420 | 2707 | unfolding l(1)[symmetric] using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (auto simp: t.enum_strict_mono) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2708 | ultimately have "i' < j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2709 | using i by (simp add: enum_strict_mono i'_def) | 
| 60420 | 2710 | with \<open>j \<noteq> i\<close> \<open>j \<le> n\<close> have "t.enum k \<le> t.enum (Suc (Suc l))" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2711 | unfolding i'_def by (simp add: enum_mono k eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2712 | then have "k \<le> Suc (Suc l)" | 
| 60420 | 2713 | using \<open>k \<le> n\<close> \<open>Suc l < k\<close> by (simp add: t.enum_mono) } | 
| 2714 | with \<open>Suc l < k\<close> have "Suc (Suc l) = k" by simp | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2715 | then have "enum (Suc (Suc i')) = t.enum (Suc (Suc l))" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2716 | using i by (simp add: k i'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2717 | also have "\<dots> = (enum i') (u l := Suc (enum i' (u l)), u (Suc l) := Suc (enum i' (u (Suc l))))" | 
| 60420 | 2718 | using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (simp add: t.enum_Suc l t.upd_inj) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 2719 | finally have "(u l = upd i' \<and> u (Suc l) = upd (Suc i')) \<or> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2720 | (u l = upd (Suc i') \<and> u (Suc l) = upd i')" | 
| 62390 | 2721 | using \<open>Suc i' < n\<close> by (auto simp: enum_Suc fun_eq_iff split: if_split_asm) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2722 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2723 |       then have "t = s \<or> t = b.enum ` {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2724 | proof (elim disjE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2725 | assume u: "u l = upd i'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2726 | have "c = t.enum (Suc l)" unfolding c_eq .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2727 | also have "t.enum (Suc l) = enum (Suc i')" | 
| 60420 | 2728 | 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) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2729 | also have "\<dots> = a" | 
| 60420 | 2730 | using \<open>a = enum i\<close> i by (simp add: i'_def) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2731 | finally show ?thesis | 
| 60420 | 2732 | using eq_sma \<open>a \<in> s\<close> \<open>c \<in> t\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2733 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2734 | assume u: "u l = upd (Suc i')" | 
| 63040 | 2735 |         define B where "B = b.enum ` {..n}"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2736 | have "b.enum i' = enum i'" | 
| 68022 | 2737 | using enum_eq_benum[of i'] i by (auto simp: i'_def gr0_conv_Suc) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2738 | have "c = t.enum (Suc l)" unfolding c_eq .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2739 | also have "t.enum (Suc l) = b.enum (Suc i')" | 
| 60420 | 2740 | using u \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close> | 
| 2741 | by (simp_all add: enum_Suc t.enum_Suc l b.enum_Suc \<open>b.enum i' = enum i'\<close> swap_apply1) | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2742 | (simp add: Suc_i') | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2743 | also have "\<dots> = b.enum i" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2744 | using i by (simp add: i'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2745 | finally have "c = b.enum i" . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2746 |         then have "t - {c} = B - {c}" "c \<in> B"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2747 | unfolding eq_sma[symmetric] eq B_def using i by auto | 
| 60420 | 2748 | with \<open>c \<in> t\<close> have "t = B" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2749 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2750 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2751 | by (simp add: B_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2752 | qed } | 
| 60420 | 2753 | with ks_f' eq \<open>a \<noteq> b.enum i\<close> \<open>n \<noteq> 0\<close> \<open>i \<le> n\<close> show ?thesis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2754 |       apply (intro ex1I[of _ "b.enum ` {.. n}"])
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2755 | apply auto [] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2756 | apply metis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2757 | done | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2758 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2759 | then show ?thesis | 
| 60420 | 2760 | using s \<open>a \<in> s\<close> by (simp add: card_2_exists Ex1_def) metis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2761 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2762 | |
| 60420 | 2763 | text \<open>Hence another step towards concreteness.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2764 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2765 | lemma kuhn_simplex_lemma: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2766 |   assumes "\<forall>s. ksimplex p (Suc n) s \<longrightarrow> rl ` s \<subseteq> {.. Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2767 |     and "odd (card {f. \<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> (f = s - {a}) \<and>
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2768 |       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))})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2769 |   shows "odd (card {s. ksimplex p (Suc n) s \<and> rl ` s = {..Suc n}})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2770 | proof (rule kuhn_complete_lemma[OF finite_ksimplexes refl, unfolded mem_Collect_eq, | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2771 |     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)"],
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2772 | safe del: notI) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2773 | |
| 53186 | 2774 | have *: "\<And>x y. x = y \<Longrightarrow> odd (card x) \<Longrightarrow> odd (card y)" | 
| 2775 | by auto | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2776 |   show "odd (card {f. (\<exists>s\<in>{s. ksimplex p (Suc n) s}. \<exists>a\<in>s. f = s - {a}) \<and>
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2777 |     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))})"
 | 
| 53186 | 2778 | apply (rule *[OF _ assms(2)]) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2779 | apply (auto simp: atLeast0AtMost) | 
| 53186 | 2780 | done | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2781 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2782 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2783 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2784 | fix s assume s: "ksimplex p (Suc n) s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2785 | then show "card s = n + 2" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2786 | by (simp add: ksimplex_card) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2787 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2788 | fix a assume a: "a \<in> s" then show "rl a \<le> Suc n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2789 | using assms(1) s by (auto simp: subset_eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2790 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2791 |   let ?S = "{t. ksimplex p (Suc n) t \<and> (\<exists>b\<in>t. s - {a} = t - {b})}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2792 |   { fix j assume j: "j \<le> n" "\<forall>x\<in>s - {a}. x j = 0"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2793 | with s a show "card ?S = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2794 | using ksimplex_replace_0[of p "n + 1" s a j] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2795 | by (subst eq_commute) simp } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2796 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2797 |   { fix j assume j: "j \<le> n" "\<forall>x\<in>s - {a}. x j = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2798 | with s a show "card ?S = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2799 | using ksimplex_replace_1[of p "n + 1" s a j] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2800 | by (subst eq_commute) simp } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2801 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2802 |   { assume "card ?S \<noteq> 2" "\<not> (\<exists>j\<in>{..n}. \<forall>x\<in>s - {a}. x j = p)"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2803 |     with s a show "\<exists>j\<in>{..n}. \<forall>x\<in>s - {a}. x j = 0"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2804 | using ksimplex_replace_2[of p "n + 1" s a] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2805 | by (subst (asm) eq_commute) auto } | 
| 53186 | 2806 | qed | 
| 2807 | ||
| 68617 | 2808 | subsubsection \<open>Reduced labelling\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2809 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2810 | definition reduced :: "nat \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> nat" where "reduced n x = (LEAST k. k = n \<or> x k \<noteq> 0)" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2811 | |
| 53186 | 2812 | lemma reduced_labelling: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2813 | shows "reduced n x \<le> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2814 | and "\<forall>i<reduced n x. x i = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2815 | and "reduced n x = n \<or> x (reduced n x) \<noteq> 0" | 
| 53186 | 2816 | proof - | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2817 | show "reduced n x \<le> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2818 | unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2819 | show "\<forall>i<reduced n x. x i = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2820 | unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) fastforce+ | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2821 | show "reduced n x = n \<or> x (reduced n x) \<noteq> 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2822 | unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) fastforce+ | 
| 53186 | 2823 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2824 | |
| 53186 | 2825 | lemma reduced_labelling_unique: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2826 | "r \<le> n \<Longrightarrow> \<forall>i<r. x i = 0 \<Longrightarrow> r = n \<or> x r \<noteq> 0 \<Longrightarrow> reduced n x = r" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2827 | unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) (metis le_less not_le)+ | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2828 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2829 | lemma reduced_labelling_zero: "j < n \<Longrightarrow> x j = 0 \<Longrightarrow> reduced n x \<noteq> j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2830 | using reduced_labelling[of n x] by auto | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2831 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2832 | lemma reduce_labelling_zero[simp]: "reduced 0 x = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2833 | by (rule reduced_labelling_unique) auto | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2834 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2835 | lemma reduced_labelling_nonzero: "j < n \<Longrightarrow> x j \<noteq> 0 \<Longrightarrow> reduced n x \<le> j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2836 | using reduced_labelling[of n x] by (elim allE[where x=j]) auto | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2837 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2838 | lemma reduced_labelling_Suc: "reduced (Suc n) x \<noteq> Suc n \<Longrightarrow> reduced (Suc n) x = reduced n x" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2839 | using reduced_labelling[of "Suc n" x] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2840 | by (intro reduced_labelling_unique[symmetric]) auto | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2841 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2842 | lemma complete_face_top: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2843 | assumes "\<forall>x\<in>f. \<forall>j\<le>n. x j = 0 \<longrightarrow> lab x j = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2844 | and "\<forall>x\<in>f. \<forall>j\<le>n. x j = p \<longrightarrow> lab x j = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2845 |     and eq: "(reduced (Suc n) \<circ> lab) ` f = {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2846 | 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)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2847 | proof (safe del: disjCI) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2848 | fix x j assume j: "j \<le> n" "\<forall>x\<in>f. x j = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2849 |   { fix x assume "x \<in> f" with assms j have "reduced (Suc n) (lab x) \<noteq> j"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2850 | by (intro reduced_labelling_zero) auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2851 | moreover have "j \<in> (reduced (Suc n) \<circ> lab) ` f" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2852 | using j eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2853 | ultimately show "x n = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2854 | by force | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2855 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2856 | fix x j assume j: "j \<le> n" "\<forall>x\<in>f. x j = p" and x: "x \<in> f" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2857 | have "j = n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2858 | proof (rule ccontr) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2859 | assume "\<not> ?thesis" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2860 |     { fix x assume "x \<in> f"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2861 | with assms j have "reduced (Suc n) (lab x) \<le> j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2862 | by (intro reduced_labelling_nonzero) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2863 | then have "reduced (Suc n) (lab x) \<noteq> n" | 
| 60420 | 2864 | using \<open>j \<noteq> n\<close> \<open>j \<le> n\<close> by simp } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2865 | moreover | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 2866 | have "n \<in> (reduced (Suc n) \<circ> lab) ` f" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2867 | using eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2868 | ultimately show False | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2869 | by force | 
| 53186 | 2870 | qed | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2871 | moreover have "j \<in> (reduced (Suc n) \<circ> lab) ` f" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2872 | using j eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2873 | ultimately show "x n = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2874 | using j x by auto | 
| 53688 | 2875 | qed auto | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2876 | |
| 60420 | 2877 | text \<open>Hence we get just about the nice induction.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2878 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2879 | lemma kuhn_induction: | 
| 53688 | 2880 | assumes "0 < p" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2881 | and lab_0: "\<forall>x. \<forall>j\<le>n. (\<forall>j. x j \<le> p) \<and> x j = 0 \<longrightarrow> lab x j = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2882 | and lab_1: "\<forall>x. \<forall>j\<le>n. (\<forall>j. x j \<le> p) \<and> x j = p \<longrightarrow> lab x j = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2883 |     and odd: "odd (card {s. ksimplex p n s \<and> (reduced n\<circ>lab) ` s = {..n}})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2884 |   shows "odd (card {s. ksimplex p (Suc n) s \<and> (reduced (Suc n)\<circ>lab) ` s = {..Suc n}})"
 | 
| 53248 | 2885 | proof - | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2886 | let ?rl = "reduced (Suc n) \<circ> lab" and ?ext = "\<lambda>f v. \<exists>j\<le>n. \<forall>x\<in>f. x j = v" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2887 | 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)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2888 |   have "\<forall>s. ksimplex p (Suc n) s \<longrightarrow> ?rl ` s \<subseteq> {..Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2889 | by (simp add: reduced_labelling subset_eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2890 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2891 |   have "{s. ksimplex p n s \<and> (reduced n \<circ> lab) ` s = {..n}} =
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2892 |         {f. \<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> f = s - {a} \<and> ?rl ` f = {..n} \<and> ?ext f}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2893 | proof (intro set_eqI, safe del: disjCI equalityI disjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2894 |     fix s assume s: "ksimplex p n s" and rl: "(reduced n \<circ> lab) ` s = {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2895 | from s obtain u b where "kuhn_simplex p n u b s" by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2896 | then interpret kuhn_simplex p n u b s . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2897 | have all_eq_p: "\<forall>x\<in>s. x n = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2898 | by (auto simp: out_eq_p) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2899 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2900 |     { fix x assume "x \<in> s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2901 | with lab_1[rule_format, of n x] all_eq_p s_le_p[of x] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2902 | have "?rl x \<le> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2903 | by (auto intro!: reduced_labelling_nonzero) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2904 | then have "?rl x = reduced n (lab x)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2905 | by (auto intro!: reduced_labelling_Suc) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2906 |     then have "?rl ` s = {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2907 | using rl by (simp cong: image_cong) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2908 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2909 |     obtain t a where "ksimplex p (Suc n) t" "a \<in> t" "s = t - {a}"
 | 
| 60420 | 2910 | using s unfolding simplex_top_face[OF \<open>0 < p\<close> all_eq_p] by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2911 | ultimately | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2912 |     show "\<exists>t a. ksimplex p (Suc n) t \<and> a \<in> t \<and> s = t - {a} \<and> ?rl ` s = {..n} \<and> ?ext s"
 | 
| 53688 | 2913 | by auto | 
| 53248 | 2914 | next | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2915 |     fix x s a assume s: "ksimplex p (Suc n) s" and rl: "?rl ` (s - {a}) = {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2916 |       and a: "a \<in> s" and "?ext (s - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2917 | from s obtain u b where "kuhn_simplex p (Suc n) u b s" by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2918 | then interpret kuhn_simplex p "Suc n" u b s . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2919 | have all_eq_p: "\<forall>x\<in>s. x (Suc n) = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2920 | by (auto simp: out_eq_p) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2921 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2922 |     { fix x assume "x \<in> s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2923 |       then have "?rl x \<in> ?rl ` (s - {a})"
 | 
| 53248 | 2924 | by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2925 | then have "?rl x \<le> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2926 | unfolding rl by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2927 | then have "?rl x = reduced n (lab x)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2928 | by (auto intro!: reduced_labelling_Suc) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2929 |     then show rl': "(reduced n\<circ>lab) ` (s - {a}) = {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2930 | unfolding rl[symmetric] by (intro image_cong) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2931 | |
| 60420 | 2932 |     from \<open>?ext (s - {a})\<close>
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2933 |     have all_eq_p: "\<forall>x\<in>s - {a}. x n = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2934 | proof (elim disjE exE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2935 |       fix j assume "j \<le> n" "\<forall>x\<in>s - {a}. x j = 0"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2936 | with lab_0[rule_format, of j] all_eq_p s_le_p | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2937 |       have "\<And>x. x \<in> s - {a} \<Longrightarrow> reduced (Suc n) (lab x) \<noteq> j"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2938 | by (intro reduced_labelling_zero) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2939 |       moreover have "j \<in> ?rl ` (s - {a})"
 | 
| 60420 | 2940 | using \<open>j \<le> n\<close> unfolding rl by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2941 | ultimately show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2942 | by force | 
| 53248 | 2943 | next | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2944 |       fix j assume "j \<le> n" and eq_p: "\<forall>x\<in>s - {a}. x j = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2945 | show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2946 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2947 | assume "j = n" with eq_p show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2948 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2949 | assume "j \<noteq> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2950 |         { fix x assume x: "x \<in> s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2951 | have "reduced n (lab x) \<le> j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2952 | proof (rule reduced_labelling_nonzero) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2953 | show "lab x j \<noteq> 0" | 
| 60420 | 2954 | using lab_1[rule_format, of j x] x s_le_p[of x] eq_p \<open>j \<le> n\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2955 | show "j < n" | 
| 60420 | 2956 | using \<open>j \<le> n\<close> \<open>j \<noteq> n\<close> by simp | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2957 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2958 | then have "reduced n (lab x) \<noteq> n" | 
| 60420 | 2959 | using \<open>j \<le> n\<close> \<open>j \<noteq> n\<close> by simp } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2960 |         moreover have "n \<in> (reduced n\<circ>lab) ` (s - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2961 | unfolding rl' by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2962 | ultimately show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2963 | by force | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2964 | qed | 
| 53248 | 2965 | qed | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2966 |     show "ksimplex p n (s - {a})"
 | 
| 60420 | 2967 | unfolding simplex_top_face[OF \<open>0 < p\<close> all_eq_p] using s a by auto | 
| 53248 | 2968 | qed | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2969 | ultimately show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2970 | using assms by (intro kuhn_simplex_lemma) auto | 
| 53248 | 2971 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2972 | |
| 60420 | 2973 | text \<open>And so we get the final combinatorial result.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2974 | |
| 53688 | 2975 | lemma ksimplex_0: "ksimplex p 0 s \<longleftrightarrow> s = {(\<lambda>x. p)}"
 | 
| 53248 | 2976 | proof | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2977 |   assume "ksimplex p 0 s" then show "s = {(\<lambda>x. p)}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2978 | by (blast dest: kuhn_simplex.ksimplex_0 elim: ksimplex.cases) | 
| 53248 | 2979 | next | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2980 |   assume s: "s = {(\<lambda>x. p)}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2981 | show "ksimplex p 0 s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2982 | proof (intro ksimplex, unfold_locales) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2983 |     show "(\<lambda>_. p) \<in> {..<0::nat} \<rightarrow> {..<p}" by auto
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2984 |     show "bij_betw id {..<0} {..<0}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2985 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2986 | qed (auto simp: s) | 
| 53248 | 2987 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2988 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2989 | lemma kuhn_combinatorial: | 
| 53688 | 2990 | assumes "0 < p" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2991 | and "\<forall>x j. (\<forall>j. x j \<le> p) \<and> j < n \<and> x j = 0 \<longrightarrow> lab x j = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2992 | and "\<forall>x j. (\<forall>j. x j \<le> p) \<and> j < n \<and> x j = p \<longrightarrow> lab x j = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2993 |   shows "odd (card {s. ksimplex p n s \<and> (reduced n\<circ>lab) ` s = {..n}})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2994 | (is "odd (card (?M n))") | 
| 53248 | 2995 | using assms | 
| 2996 | proof (induct n) | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2997 | case 0 then show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2998 | by (simp add: ksimplex_0 cong: conj_cong) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 2999 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3000 | case (Suc n) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3001 | then have "odd (card (?M n))" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3002 | by force | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3003 | with Suc show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3004 | using kuhn_induction[of p n] by (auto simp: comp_def) | 
| 53248 | 3005 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3006 | |
| 53248 | 3007 | lemma kuhn_lemma: | 
| 53688 | 3008 | fixes n p :: nat | 
| 3009 | assumes "0 < p" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3010 | and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. label x i = (0::nat) \<or> label x i = 1)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3011 | and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = 0 \<longrightarrow> label x i = 0)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3012 | and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = p \<longrightarrow> label x i = 1)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3013 | obtains q where "\<forall>i<n. q i < p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3014 | 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" | 
| 53248 | 3015 | proof - | 
| 60580 | 3016 | let ?rl = "reduced n \<circ> label" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3017 |   let ?A = "{s. ksimplex p n s \<and> ?rl ` s = {..n}}"
 | 
| 53248 | 3018 | have "odd (card ?A)" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3019 | using assms by (intro kuhn_combinatorial[of p n label]) auto | 
| 53688 | 3020 |   then have "?A \<noteq> {}"
 | 
| 60580 | 3021 | by fastforce | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3022 |   then obtain s b u where "kuhn_simplex p n b u s" and rl: "?rl ` s = {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3023 | by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3024 | interpret kuhn_simplex p n b u s by fact | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3025 | |
| 53248 | 3026 | show ?thesis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3027 | proof (intro that[of b] allI impI) | 
| 60580 | 3028 | fix i | 
| 3029 | assume "i < n" | |
| 3030 | then show "b i < p" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3031 | using base by auto | 
| 53248 | 3032 | next | 
| 60580 | 3033 | fix i | 
| 3034 | assume "i < n" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3035 |     then have "i \<in> {.. n}" "Suc i \<in> {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3036 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3037 | then obtain u v where u: "u \<in> s" "Suc i = ?rl u" and v: "v \<in> s" "i = ?rl v" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3038 | unfolding rl[symmetric] by blast | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3039 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3040 | have "label u i \<noteq> label v i" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3041 | using reduced_labelling [of n "label u"] reduced_labelling [of n "label v"] | 
| 60420 | 3042 | u(2)[symmetric] v(2)[symmetric] \<open>i < n\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3043 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3044 | moreover | 
| 60580 | 3045 | 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 | 
| 3046 | 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>] | |
| 3047 | by auto | |
| 3048 | ultimately show "\<exists>r s. (\<forall>j<n. b j \<le> r j \<and> r j \<le> b j + 1) \<and> | |
| 3049 | (\<forall>j<n. b j \<le> s j \<and> s j \<le> b j + 1) \<and> label r i \<noteq> label s i" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3050 | by blast | 
| 53248 | 3051 | qed | 
| 3052 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3053 | |
| 68617 | 3054 | subsubsection \<open>Main result for the unit cube\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3055 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3056 | lemma kuhn_labelling_lemma': | 
| 53688 | 3057 | assumes "(\<forall>x::nat\<Rightarrow>real. P x \<longrightarrow> P (f x))" | 
| 3058 | and "\<forall>x. P x \<longrightarrow> (\<forall>i::nat. Q i \<longrightarrow> 0 \<le> x i \<and> x i \<le> 1)" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3059 | shows "\<exists>l. (\<forall>x i. l x i \<le> (1::nat)) \<and> | 
| 53688 | 3060 | (\<forall>x i. P x \<and> Q i \<and> x i = 0 \<longrightarrow> l x i = 0) \<and> | 
| 3061 | (\<forall>x i. P x \<and> Q i \<and> x i = 1 \<longrightarrow> l x i = 1) \<and> | |
| 3062 | (\<forall>x i. P x \<and> Q i \<and> l x i = 0 \<longrightarrow> x i \<le> f x i) \<and> | |
| 3063 | (\<forall>x i. P x \<and> Q i \<and> l x i = 1 \<longrightarrow> f x i \<le> x i)" | |
| 53185 | 3064 | proof - | 
| 53688 | 3065 | have and_forall_thm: "\<And>P Q. (\<forall>x. P x) \<and> (\<forall>x. Q x) \<longleftrightarrow> (\<forall>x. P x \<and> Q x)" | 
| 3066 | by auto | |
| 3067 | 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" | |
| 53185 | 3068 | by auto | 
| 3069 | show ?thesis | |
| 3070 | unfolding and_forall_thm | |
| 3071 | apply (subst choice_iff[symmetric])+ | |
| 53688 | 3072 | apply rule | 
| 3073 | apply rule | |
| 3074 | proof - | |
| 60580 | 3075 | fix x x' | 
| 53688 | 3076 | let ?R = "\<lambda>y::nat. | 
| 60580 | 3077 | (P x \<and> Q x' \<and> x x' = 0 \<longrightarrow> y = 0) \<and> | 
| 3078 | (P x \<and> Q x' \<and> x x' = 1 \<longrightarrow> y = 1) \<and> | |
| 3079 | (P x \<and> Q x' \<and> y = 0 \<longrightarrow> x x' \<le> (f x) x') \<and> | |
| 3080 | (P x \<and> Q x' \<and> y = 1 \<longrightarrow> (f x) x' \<le> x x')" | |
| 3081 | have "0 \<le> f x x' \<and> f x x' \<le> 1" if "P x" "Q x'" | |
| 3082 | using assms(2)[rule_format,of "f x" x'] that | |
| 3083 | apply (drule_tac assms(1)[rule_format]) | |
| 3084 | apply auto | |
| 3085 | done | |
| 53688 | 3086 | then have "?R 0 \<or> ?R 1" | 
| 3087 | by auto | |
| 60580 | 3088 | then show "\<exists>y\<le>1. ?R y" | 
| 53688 | 3089 | by auto | 
| 53185 | 3090 | qed | 
| 3091 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3092 | |
| 68617 | 3093 | subsection \<open>Brouwer's fixed point theorem\<close> | 
| 3094 | ||
| 68621 | 3095 | text \<open>We start proving Brouwer's fixed point theorem for the unit cube = \<open>cbox 0 One\<close>.\<close> | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3096 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3097 | lemma brouwer_cube: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3098 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | 
| 68621 | 3099 | assumes "continuous_on (cbox 0 One) f" | 
| 3100 | and "f ` cbox 0 One \<subseteq> cbox 0 One" | |
| 3101 | shows "\<exists>x\<in>cbox 0 One. f x = x" | |
| 53185 | 3102 | proof (rule ccontr) | 
| 63040 | 3103 |   define n where "n = DIM('a)"
 | 
| 53185 | 3104 | have n: "1 \<le> n" "0 < n" "n \<noteq> 0" | 
| 68022 | 3105 | unfolding n_def by (auto simp: Suc_le_eq DIM_positive) | 
| 53674 | 3106 | assume "\<not> ?thesis" | 
| 68621 | 3107 | then have *: "\<not> (\<exists>x\<in>cbox 0 One. f x - x = 0)" | 
| 53674 | 3108 | by auto | 
| 55522 | 3109 | obtain d where | 
| 68621 | 3110 | d: "d > 0" "\<And>x. x \<in> cbox 0 One \<Longrightarrow> d \<le> norm (f x - x)" | 
| 3111 | apply (rule brouwer_compactness_lemma[OF compact_cbox _ *]) | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56273diff
changeset | 3112 | apply (rule continuous_intros assms)+ | 
| 55522 | 3113 | apply blast | 
| 53185 | 3114 | done | 
| 68621 | 3115 | have *: "\<forall>x. x \<in> cbox 0 One \<longrightarrow> f x \<in> cbox 0 One" | 
| 3116 | "\<forall>x. x \<in> (cbox 0 One::'a set) \<longrightarrow> (\<forall>i\<in>Basis. True \<longrightarrow> 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)" | |
| 53185 | 3117 | using assms(2)[unfolded image_subset_iff Ball_def] | 
| 68621 | 3118 | unfolding cbox_def | 
| 55522 | 3119 | by auto | 
| 68022 | 3120 | obtain label :: "'a \<Rightarrow> 'a \<Rightarrow> nat" where label [rule_format]: | 
| 55522 | 3121 | "\<forall>x. \<forall>i\<in>Basis. label x i \<le> 1" | 
| 68621 | 3122 | "\<forall>x. \<forall>i\<in>Basis. x \<in> cbox 0 One \<and> x \<bullet> i = 0 \<longrightarrow> label x i = 0" | 
| 3123 | "\<forall>x. \<forall>i\<in>Basis. x \<in> cbox 0 One \<and> x \<bullet> i = 1 \<longrightarrow> label x i = 1" | |
| 3124 | "\<forall>x. \<forall>i\<in>Basis. x \<in> cbox 0 One \<and> label x i = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i" | |
| 3125 | "\<forall>x. \<forall>i\<in>Basis. x \<in> cbox 0 One \<and> label x i = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i" | |
| 68022 | 3126 | using kuhn_labelling_lemma[OF *] by auto | 
| 53185 | 3127 | note label = this [rule_format] | 
| 68621 | 3128 | have lem1: "\<forall>x\<in>cbox 0 One. \<forall>y\<in>cbox 0 One. \<forall>i\<in>Basis. label x i \<noteq> label y i \<longrightarrow> | 
| 61945 | 3129 | \<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)" | 
| 53185 | 3130 | proof safe | 
| 3131 | fix x y :: 'a | |
| 68621 | 3132 | assume x: "x \<in> cbox 0 One" and y: "y \<in> cbox 0 One" | 
| 53185 | 3133 | fix i | 
| 3134 | assume i: "label x i \<noteq> label y i" "i \<in> Basis" | |
| 3135 | have *: "\<And>x y fx fy :: real. x \<le> fx \<and> fy \<le> y \<or> fx \<le> x \<and> y \<le> fy \<Longrightarrow> | |
| 61945 | 3136 | \<bar>fx - x\<bar> \<le> \<bar>fy - fx\<bar> + \<bar>y - x\<bar>" by auto | 
| 3137 | have "\<bar>(f x - x) \<bullet> i\<bar> \<le> \<bar>(f y - f x)\<bullet>i\<bar> + \<bar>(y - x)\<bullet>i\<bar>" | |
| 68022 | 3138 | proof (cases "label x i = 0") | 
| 3139 | case True | |
| 3140 | then have fxy: "\<not> f y \<bullet> i \<le> y \<bullet> i \<Longrightarrow> f x \<bullet> i \<le> x \<bullet> i" | |
| 3141 | by (metis True i label(1) label(5) le_antisym less_one not_le_imp_less y) | |
| 3142 | show ?thesis | |
| 3143 | unfolding inner_simps | |
| 3144 | by (rule *) (auto simp: True i label x y fxy) | |
| 53185 | 3145 | next | 
| 68022 | 3146 | case False | 
| 3147 | then show ?thesis | |
| 3148 | using label [OF \<open>i \<in> Basis\<close>] i(1) x y | |
| 3149 | apply (auto simp: inner_diff_left le_Suc_eq) | |
| 3150 | by (metis "*") | |
| 53185 | 3151 | qed | 
| 3152 | also have "\<dots> \<le> norm (f y - f x) + norm (y - x)" | |
| 68022 | 3153 | by (simp add: add_mono i(2) norm_bound_Basis_le) | 
| 53185 | 3154 | finally show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)" | 
| 3155 | unfolding inner_simps . | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3156 | qed | 
| 68621 | 3157 | have "\<exists>e>0. \<forall>x\<in>cbox 0 One. \<forall>y\<in>cbox 0 One. \<forall>z\<in>cbox 0 One. \<forall>i\<in>Basis. | 
| 68022 | 3158 | norm (x - z) < e \<longrightarrow> norm (y - z) < e \<longrightarrow> label x i \<noteq> label y i \<longrightarrow> | 
| 61945 | 3159 | \<bar>(f(z) - z)\<bullet>i\<bar> < d / (real n)" | 
| 53185 | 3160 | proof - | 
| 53688 | 3161 | have d': "d / real n / 8 > 0" | 
| 56541 | 3162 | using d(1) by (simp add: n_def DIM_positive) | 
| 68621 | 3163 | have *: "uniformly_continuous_on (cbox 0 One) f" | 
| 3164 | by (rule compact_uniformly_continuous[OF assms(1) compact_cbox]) | |
| 55522 | 3165 | obtain e where e: | 
| 3166 | "e > 0" | |
| 68621 | 3167 | "\<And>x x'. x \<in> cbox 0 One \<Longrightarrow> | 
| 3168 | x' \<in> cbox 0 One \<Longrightarrow> | |
| 55522 | 3169 | norm (x' - x) < e \<Longrightarrow> | 
| 3170 | norm (f x' - f x) < d / real n / 8" | |
| 3171 | using *[unfolded uniformly_continuous_on_def,rule_format,OF d'] | |
| 3172 | unfolding dist_norm | |
| 3173 | by blast | |
| 53185 | 3174 | show ?thesis | 
| 68022 | 3175 | proof (intro exI conjI ballI impI) | 
| 53185 | 3176 | show "0 < min (e / 2) (d / real n / 8)" | 
| 3177 | using d' e by auto | |
| 3178 | fix x y z i | |
| 53688 | 3179 | assume as: | 
| 68621 | 3180 | "x \<in> cbox 0 One" "y \<in> cbox 0 One" "z \<in> cbox 0 One" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36587diff
changeset | 3181 | "norm (x - z) < min (e / 2) (d / real n / 8)" | 
| 53688 | 3182 | "norm (y - z) < min (e / 2) (d / real n / 8)" | 
| 3183 | "label x i \<noteq> label y i" | |
| 3184 | assume i: "i \<in> Basis" | |
| 61945 | 3185 | have *: "\<And>z fz x fx n1 n2 n3 n4 d4 d :: real. \<bar>fx - x\<bar> \<le> n1 + n2 \<Longrightarrow> | 
| 3186 | \<bar>fx - fz\<bar> \<le> n3 \<Longrightarrow> \<bar>x - z\<bar> \<le> n4 \<Longrightarrow> | |
| 53185 | 3187 | n1 < d4 \<Longrightarrow> n2 < 2 * d4 \<Longrightarrow> n3 < d4 \<Longrightarrow> n4 < d4 \<Longrightarrow> | 
| 61945 | 3188 | (8 * d4 = d) \<Longrightarrow> \<bar>fz - z\<bar> < d" | 
| 53688 | 3189 | by auto | 
| 3190 | show "\<bar>(f z - z) \<bullet> i\<bar> < d / real n" | |
| 3191 | unfolding inner_simps | |
| 53185 | 3192 | proof (rule *) | 
| 3193 | show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y -f x) + norm (y - x)" | |
| 68022 | 3194 | using as(1) as(2) as(6) i lem1 by blast | 
| 3195 | show "norm (f x - f z) < d / real n / 8" | |
| 3196 | using d' e as by auto | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3197 | 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)" | 
| 55522 | 3198 | unfolding inner_diff_left[symmetric] | 
| 53688 | 3199 | by (rule Basis_le_norm[OF i])+ | 
| 3200 | have tria: "norm (y - x) \<le> norm (y - z) + norm (x - z)" | |
| 53185 | 3201 | using dist_triangle[of y x z, unfolded dist_norm] | 
| 53688 | 3202 | unfolding norm_minus_commute | 
| 3203 | by auto | |
| 53185 | 3204 | also have "\<dots> < e / 2 + e / 2" | 
| 68022 | 3205 | using as(4) as(5) by auto | 
| 53185 | 3206 | finally show "norm (f y - f x) < d / real n / 8" | 
| 68022 | 3207 | using as(1) as(2) e(2) by auto | 
| 53185 | 3208 | have "norm (y - z) + norm (x - z) < d / real n / 8 + d / real n / 8" | 
| 68022 | 3209 | using as(4) as(5) by auto | 
| 3210 | with tria show "norm (y - x) < 2 * (d / real n / 8)" | |
| 53688 | 3211 | by auto | 
| 68022 | 3212 | qed (use as in auto) | 
| 53185 | 3213 | qed | 
| 3214 | qed | |
| 55522 | 3215 | then | 
| 3216 | obtain e where e: | |
| 3217 | "e > 0" | |
| 68621 | 3218 | "\<And>x y z i. x \<in> cbox 0 One \<Longrightarrow> | 
| 3219 | y \<in> cbox 0 One \<Longrightarrow> | |
| 3220 | z \<in> cbox 0 One \<Longrightarrow> | |
| 55522 | 3221 | i \<in> Basis \<Longrightarrow> | 
| 3222 | norm (x - z) < e \<and> norm (y - z) < e \<and> label x i \<noteq> label y i \<Longrightarrow> | |
| 3223 | \<bar>(f z - z) \<bullet> i\<bar> < d / real n" | |
| 3224 | by blast | |
| 3225 | obtain p :: nat where p: "1 + real n / e \<le> real p" | |
| 3226 | using real_arch_simple .. | |
| 53185 | 3227 | have "1 + real n / e > 0" | 
| 56541 | 3228 | using e(1) n by (simp add: add_pos_pos) | 
| 53688 | 3229 | then have "p > 0" | 
| 3230 | using p by auto | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3231 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3232 |   obtain b :: "nat \<Rightarrow> 'a" where b: "bij_betw b {..< n} Basis"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3233 | by atomize_elim (auto simp: n_def intro!: finite_same_card_bij) | 
| 63040 | 3234 |   define b' where "b' = inv_into {..< n} b"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3235 |   then have b': "bij_betw b' Basis {..< n}"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3236 | using bij_betw_inv_into[OF b] by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3237 |   then have b'_Basis: "\<And>i. i \<in> Basis \<Longrightarrow> b' i \<in> {..< n}"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3238 | unfolding bij_betw_def by (auto simp: set_eq_iff) | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3239 | have bb'[simp]:"\<And>i. i \<in> Basis \<Longrightarrow> b (b' i) = i" | 
| 53688 | 3240 | unfolding b'_def | 
| 3241 | using b | |
| 3242 | by (auto simp: f_inv_into_f bij_betw_def) | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3243 | have b'b[simp]:"\<And>i. i < n \<Longrightarrow> b' (b i) = i" | 
| 53688 | 3244 | unfolding b'_def | 
| 3245 | using b | |
| 3246 | by (auto simp: inv_into_f_eq bij_betw_def) | |
| 3247 | have *: "\<And>x :: nat. x = 0 \<or> x = 1 \<longleftrightarrow> x \<le> 1" | |
| 3248 | by auto | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3249 | have b'': "\<And>j. j < n \<Longrightarrow> b j \<in> Basis" | 
| 53185 | 3250 | using b unfolding bij_betw_def by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3251 | have q1: "0 < p" "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3252 | (\<forall>i<n. (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0 \<or> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3253 | (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)" | 
| 53688 | 3254 | unfolding * | 
| 60420 | 3255 | using \<open>p > 0\<close> \<open>n > 0\<close> | 
| 53688 | 3256 | using label(1)[OF b''] | 
| 3257 | by auto | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3258 |   { fix x :: "nat \<Rightarrow> nat" and i assume "\<forall>i<n. x i \<le> p" "i < n" "x i = p \<or> x i = 0"
 | 
| 68621 | 3259 | then have "(\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<in> (cbox 0 One::'a set)" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3260 | using b'_Basis | 
| 68621 | 3261 | by (auto simp: cbox_def inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3262 | note cube = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3263 | have q2: "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = 0 \<longrightarrow> | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3264 | (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0)" | 
| 68022 | 3265 | unfolding o_def using cube \<open>p > 0\<close> by (intro allI impI label(2)) (auto simp: b'') | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 3266 | have q3: "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = p \<longrightarrow> | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3267 | (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)" | 
| 68022 | 3268 | using cube \<open>p > 0\<close> unfolding o_def by (intro allI impI label(3)) (auto simp: b'') | 
| 55522 | 3269 | obtain q where q: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3270 | "\<forall>i<n. q i < p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3271 | "\<forall>i<n. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3272 | \<exists>r s. (\<forall>j<n. q j \<le> r j \<and> r j \<le> q j + 1) \<and> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3273 | (\<forall>j<n. q j \<le> s j \<and> s j \<le> q j + 1) \<and> | 
| 55522 | 3274 | (label (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i) \<circ> b) i \<noteq> | 
| 3275 | (label (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i) \<circ> b) i" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3276 | by (rule kuhn_lemma[OF q1 q2 q3]) | 
| 63040 | 3277 | define z :: 'a where "z = (\<Sum>i\<in>Basis. (real (q (b' i)) / real p) *\<^sub>R i)" | 
| 61945 | 3278 | have "\<exists>i\<in>Basis. d / real n \<le> \<bar>(f z - z)\<bullet>i\<bar>" | 
| 53185 | 3279 | proof (rule ccontr) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3280 |     have "\<forall>i\<in>Basis. q (b' i) \<in> {0..p}"
 | 
| 53688 | 3281 | using q(1) b' | 
| 3282 | by (auto intro: less_imp_le simp: bij_betw_def) | |
| 68621 | 3283 | then have "z \<in> cbox 0 One" | 
| 3284 | unfolding z_def cbox_def | |
| 53688 | 3285 | using b'_Basis | 
| 68022 | 3286 | by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) | 
| 53688 | 3287 | then have d_fz_z: "d \<le> norm (f z - z)" | 
| 3288 | by (rule d) | |
| 3289 | assume "\<not> ?thesis" | |
| 53674 | 3290 | then have as: "\<forall>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar> < d / real n" | 
| 60420 | 3291 | using \<open>n > 0\<close> | 
| 68022 | 3292 | by (auto simp: not_le inner_diff) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3293 | have "norm (f z - z) \<le> (\<Sum>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar>)" | 
| 53688 | 3294 | unfolding inner_diff_left[symmetric] | 
| 3295 | by (rule norm_le_l1) | |
| 53185 | 3296 | also have "\<dots> < (\<Sum>(i::'a) \<in> Basis. d / real n)" | 
| 68022 | 3297 | by (meson as finite_Basis nonempty_Basis sum_strict_mono) | 
| 53185 | 3298 | also have "\<dots> = d" | 
| 68022 | 3299 | using DIM_positive[where 'a='a] by (auto simp: n_def) | 
| 53688 | 3300 | finally show False | 
| 3301 | using d_fz_z by auto | |
| 53185 | 3302 | qed | 
| 55522 | 3303 | then obtain i where i: "i \<in> Basis" "d / real n \<le> \<bar>(f z - z) \<bullet> i\<bar>" .. | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3304 | have *: "b' i < n" | 
| 55522 | 3305 | using i and b'[unfolded bij_betw_def] | 
| 53688 | 3306 | by auto | 
| 55522 | 3307 | obtain r s where rs: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3308 | "\<And>j. j < n \<Longrightarrow> q j \<le> r j \<and> r j \<le> q j + 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3309 | "\<And>j. j < n \<Longrightarrow> q j \<le> s j \<and> s j \<le> q j + 1" | 
| 55522 | 3310 | "(label (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i) \<circ> b) (b' i) \<noteq> | 
| 3311 | (label (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i) \<circ> b) (b' i)" | |
| 3312 | using q(2)[rule_format,OF *] by blast | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 3313 | have b'_im: "\<And>i. i \<in> Basis \<Longrightarrow> b' i < n" | 
| 53185 | 3314 | using b' unfolding bij_betw_def by auto | 
| 63040 | 3315 | define r' ::'a where "r' = (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i)" | 
| 53185 | 3316 | have "\<And>i. i \<in> Basis \<Longrightarrow> r (b' i) \<le> p" | 
| 3317 | apply (rule order_trans) | |
| 3318 | apply (rule rs(1)[OF b'_im,THEN conjunct2]) | |
| 53252 | 3319 | using q(1)[rule_format,OF b'_im] | 
| 68022 | 3320 | apply (auto simp: Suc_le_eq) | 
| 53185 | 3321 | done | 
| 68621 | 3322 | then have "r' \<in> cbox 0 One" | 
| 3323 | unfolding r'_def cbox_def | |
| 53688 | 3324 | using b'_Basis | 
| 68022 | 3325 | by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) | 
| 63040 | 3326 | define s' :: 'a where "s' = (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i)" | 
| 53688 | 3327 | have "\<And>i. i \<in> Basis \<Longrightarrow> s (b' i) \<le> p" | 
| 68022 | 3328 | using b'_im q(1) rs(2) by fastforce | 
| 68621 | 3329 | then have "s' \<in> cbox 0 One" | 
| 3330 | unfolding s'_def cbox_def | |
| 68022 | 3331 | using b'_Basis by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) | 
| 68621 | 3332 | have "z \<in> cbox 0 One" | 
| 3333 | unfolding z_def cbox_def | |
| 60420 | 3334 | using b'_Basis q(1)[rule_format,OF b'_im] \<open>p > 0\<close> | 
| 68022 | 3335 | by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le) | 
| 53688 | 3336 |   {
 | 
| 3337 | have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)" | |
| 68022 | 3338 | by (rule sum_mono) (use rs(1)[OF b'_im] in force) | 
| 53688 | 3339 | also have "\<dots> < e * real p" | 
| 60420 | 3340 | using p \<open>e > 0\<close> \<open>p > 0\<close> | 
| 68022 | 3341 | by (auto simp: field_simps n_def) | 
| 53185 | 3342 | finally have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) < e * real p" . | 
| 3343 | } | |
| 3344 | moreover | |
| 53688 | 3345 |   {
 | 
| 3346 | have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)" | |
| 68022 | 3347 | by (rule sum_mono) (use rs(2)[OF b'_im] in force) | 
| 53688 | 3348 | also have "\<dots> < e * real p" | 
| 60420 | 3349 | using p \<open>e > 0\<close> \<open>p > 0\<close> | 
| 68022 | 3350 | by (auto simp: field_simps n_def) | 
| 53185 | 3351 | finally have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) < e * real p" . | 
| 3352 | } | |
| 3353 | ultimately | |
| 53688 | 3354 | have "norm (r' - z) < e" and "norm (s' - z) < e" | 
| 53185 | 3355 | unfolding r'_def s'_def z_def | 
| 60420 | 3356 | using \<open>p > 0\<close> | 
| 53185 | 3357 | apply (rule_tac[!] le_less_trans[OF norm_le_l1]) | 
| 68022 | 3358 | apply (auto simp: field_simps sum_divide_distrib[symmetric] inner_diff_left) | 
| 53185 | 3359 | done | 
| 53674 | 3360 | then have "\<bar>(f z - z) \<bullet> i\<bar> < d / real n" | 
| 53688 | 3361 | using rs(3) i | 
| 3362 | unfolding r'_def[symmetric] s'_def[symmetric] o_def bb' | |
| 68621 | 3363 | by (intro e(2)[OF \<open>r'\<in>cbox 0 One\<close> \<open>s'\<in>cbox 0 One\<close> \<open>z\<in>cbox 0 One\<close>]) auto | 
| 53688 | 3364 | then show False | 
| 3365 | using i by auto | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3366 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3367 | |
| 68617 | 3368 | text \<open>Next step is to prove it for nonempty interiors.\<close> | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3369 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3370 | lemma brouwer_weak: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3371 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | 
| 68022 | 3372 | assumes "compact S" | 
| 3373 | and "convex S" | |
| 3374 |     and "interior S \<noteq> {}"
 | |
| 3375 | and "continuous_on S f" | |
| 3376 | and "f ` S \<subseteq> S" | |
| 3377 | obtains x where "x \<in> S" and "f x = x" | |
| 53185 | 3378 | proof - | 
| 68621 | 3379 | let ?U = "cbox 0 One :: 'a set" | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3380 | have "\<Sum>Basis /\<^sub>R 2 \<in> interior ?U" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3381 | proof (rule interiorI) | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3382 |     let ?I = "(\<Inter>i\<in>Basis. {x::'a. 0 < x \<bullet> i} \<inter> {x. x \<bullet> i < 1})"
 | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3383 | show "open ?I" | 
| 63332 | 3384 | by (intro open_INT finite_Basis ballI open_Int, auto intro: open_Collect_less simp: continuous_on_inner continuous_on_const continuous_on_id) | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3385 | show "\<Sum>Basis /\<^sub>R 2 \<in> ?I" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3386 | by simp | 
| 68621 | 3387 | show "?I \<subseteq> cbox 0 One" | 
| 3388 | unfolding cbox_def by force | |
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3389 | qed | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3390 |   then have *: "interior ?U \<noteq> {}" by fast
 | 
| 68022 | 3391 | have *: "?U homeomorphic S" | 
| 68621 | 3392 | using homeomorphic_convex_compact[OF convex_box(1) compact_cbox * assms(2,1,3)] . | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3393 | have "\<forall>f. continuous_on ?U f \<and> f ` ?U \<subseteq> ?U \<longrightarrow> | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3394 | (\<exists>x\<in>?U. f x = x)" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36587diff
changeset | 3395 | using brouwer_cube by auto | 
| 53674 | 3396 | then show ?thesis | 
| 53185 | 3397 | unfolding homeomorphic_fixpoint_property[OF *] | 
| 53252 | 3398 | using assms | 
| 68022 | 3399 | by (auto intro: that) | 
| 53185 | 3400 | qed | 
| 3401 | ||
| 68617 | 3402 | text \<open>Then the particular case for closed balls.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3403 | |
| 53185 | 3404 | lemma brouwer_ball: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3405 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | 
| 53674 | 3406 | assumes "e > 0" | 
| 3407 | and "continuous_on (cball a e) f" | |
| 53688 | 3408 | and "f ` cball a e \<subseteq> cball a e" | 
| 53674 | 3409 | obtains x where "x \<in> cball a e" and "f x = x" | 
| 53185 | 3410 | using brouwer_weak[OF compact_cball convex_cball, of a e f] | 
| 3411 | unfolding interior_cball ball_eq_empty | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3412 | using assms by auto | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3413 | |
| 68617 | 3414 | text \<open>And finally we prove Brouwer's fixed point theorem in its general version.\<close> | 
| 3415 | ||
| 3416 | theorem brouwer: | |
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3417 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | 
| 68022 | 3418 |   assumes S: "compact S" "convex S" "S \<noteq> {}"
 | 
| 3419 | and contf: "continuous_on S f" | |
| 3420 | and fim: "f ` S \<subseteq> S" | |
| 3421 | obtains x where "x \<in> S" and "f x = x" | |
| 53185 | 3422 | proof - | 
| 68022 | 3423 | have "\<exists>e>0. S \<subseteq> cball 0 e" | 
| 3424 | using compact_imp_bounded[OF \<open>compact S\<close>] unfolding bounded_pos | |
| 3425 | by auto | |
| 3426 | then obtain e where e: "e > 0" "S \<subseteq> cball 0 e" | |
| 55522 | 3427 | by blast | 
| 68022 | 3428 | have "\<exists>x\<in> cball 0 e. (f \<circ> closest_point S) x = x" | 
| 3429 | proof (rule_tac brouwer_ball[OF e(1)]) | |
| 3430 | show "continuous_on (cball 0 e) (f \<circ> closest_point S)" | |
| 3431 | apply (rule continuous_on_compose) | |
| 3432 | using S compact_eq_bounded_closed continuous_on_closest_point apply blast | |
| 3433 | by (meson S contf closest_point_in_set compact_imp_closed continuous_on_subset image_subsetI) | |
| 3434 | show "(f \<circ> closest_point S) ` cball 0 e \<subseteq> cball 0 e" | |
| 3435 | by clarsimp (metis S fim closest_point_exists(1) compact_eq_bounded_closed e(2) image_subset_iff mem_cball_0 subsetCE) | |
| 3436 | qed (use assms in auto) | |
| 3437 | then obtain x where x: "x \<in> cball 0 e" "(f \<circ> closest_point S) x = x" .. | |
| 3438 | have "x \<in> S" | |
| 3439 | by (metis closest_point_in_set comp_apply compact_imp_closed fim image_eqI S(1) S(3) subset_iff x(2)) | |
| 3440 | then have *: "closest_point S x = x" | |
| 3441 | by (rule closest_point_self) | |
| 53185 | 3442 | show thesis | 
| 68022 | 3443 | proof | 
| 3444 | show "closest_point S x \<in> S" | |
| 3445 | by (simp add: "*" \<open>x \<in> S\<close>) | |
| 3446 | show "f (closest_point S x) = closest_point S x" | |
| 3447 | using "*" x(2) by auto | |
| 3448 | qed | |
| 53185 | 3449 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3450 | |
| 68617 | 3451 | subsection \<open>Applications\<close> | 
| 3452 | ||
| 60420 | 3453 | text \<open>So we get the no-retraction theorem.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3454 | |
| 68617 | 3455 | corollary no_retraction_cball: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 3456 | fixes a :: "'a::euclidean_space" | 
| 53674 | 3457 | assumes "e > 0" | 
| 3458 | shows "\<not> (frontier (cball a e) retract_of (cball a e))" | |
| 53185 | 3459 | proof | 
| 60580 | 3460 | assume *: "frontier (cball a e) retract_of (cball a e)" | 
| 3461 | have **: "\<And>xa. a - (2 *\<^sub>R a - xa) = - (a - xa)" | |
| 53185 | 3462 | using scaleR_left_distrib[of 1 1 a] by auto | 
| 68022 | 3463 |   obtain x where x: "x \<in> {x. norm (a - x) = e}" "2 *\<^sub>R a - x = x"
 | 
| 3464 | proof (rule retract_fixpoint_property[OF *, of "\<lambda>x. scaleR 2 a - x"]) | |
| 3465 | show "continuous_on (frontier (cball a e)) ((-) (2 *\<^sub>R a))" | |
| 3466 | by (intro continuous_intros) | |
| 3467 | show "(-) (2 *\<^sub>R a) ` frontier (cball a e) \<subseteq> frontier (cball a e)" | |
| 3468 | by clarsimp (metis "**" dist_norm norm_minus_cancel) | |
| 3469 | qed (auto simp: dist_norm intro: brouwer_ball[OF assms]) | |
| 53674 | 3470 | then have "scaleR 2 a = scaleR 1 x + scaleR 1 x" | 
| 68022 | 3471 | by (auto simp: algebra_simps) | 
| 53674 | 3472 | then have "a = x" | 
| 53688 | 3473 | unfolding scaleR_left_distrib[symmetric] | 
| 3474 | by auto | |
| 53674 | 3475 | then show False | 
| 3476 | using x assms by auto | |
| 53185 | 3477 | qed | 
| 3478 | ||
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3479 | corollary contractible_sphere: | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3480 | fixes a :: "'a::euclidean_space" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3481 | shows "contractible(sphere a r) \<longleftrightarrow> r \<le> 0" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3482 | proof (cases "0 < r") | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3483 | case True | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3484 | then show ?thesis | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3485 | unfolding contractible_def nullhomotopic_from_sphere_extension | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3486 | using no_retraction_cball [OF True, of a] | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3487 | by (auto simp: retract_of_def retraction_def) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3488 | next | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3489 | case False | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3490 | then show ?thesis | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3491 | unfolding contractible_def nullhomotopic_from_sphere_extension | 
| 68022 | 3492 | using continuous_on_const less_eq_real_def by auto | 
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3493 | qed | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 3494 | |
| 68617 | 3495 | corollary connected_sphere_eq: | 
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3496 | fixes a :: "'a :: euclidean_space" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3497 |   shows "connected(sphere a r) \<longleftrightarrow> 2 \<le> DIM('a) \<or> r \<le> 0"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3498 | (is "?lhs = ?rhs") | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3499 | proof (cases r "0::real" rule: linorder_cases) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3500 | case less | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3501 | then show ?thesis by auto | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3502 | next | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3503 | case equal | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3504 | then show ?thesis by auto | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3505 | next | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3506 | case greater | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3507 | show ?thesis | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3508 | proof | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3509 | assume L: ?lhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3510 |     have "False" if 1: "DIM('a) = 1"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3511 | proof - | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3512 |       obtain x y where xy: "sphere a r = {x,y}" "x \<noteq> y"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3513 | using sphere_1D_doubleton [OF 1 greater] | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3514 | by (metis dist_self greater insertI1 less_add_same_cancel1 mem_sphere mult_2 not_le zero_le_dist) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3515 | then have "finite (sphere a r)" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3516 | by auto | 
| 68022 | 3517 | with L \<open>r > 0\<close> xy show "False" | 
| 3518 | using connected_finite_iff_sing by auto | |
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3519 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3520 | with greater show ?rhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3521 | by (metis DIM_ge_Suc0 One_nat_def Suc_1 le_antisym not_less_eq_eq) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3522 | next | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3523 | assume ?rhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3524 | then show ?lhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3525 | using connected_sphere greater by auto | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3526 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3527 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3528 | |
| 68617 | 3529 | corollary path_connected_sphere_eq: | 
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3530 | fixes a :: "'a :: euclidean_space" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3531 |   shows "path_connected(sphere a r) \<longleftrightarrow> 2 \<le> DIM('a) \<or> r \<le> 0"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3532 | (is "?lhs = ?rhs") | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3533 | proof | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3534 | assume ?lhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3535 | then show ?rhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3536 | using connected_sphere_eq path_connected_imp_connected by blast | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3537 | next | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3538 | assume R: ?rhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3539 | then show ?lhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3540 | by (auto simp: contractible_imp_path_connected contractible_sphere path_connected_sphere) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3541 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3542 | |
| 64122 | 3543 | proposition frontier_subset_retraction: | 
| 3544 | fixes S :: "'a::euclidean_space set" | |
| 3545 | assumes "bounded S" and fros: "frontier S \<subseteq> T" | |
| 3546 | and contf: "continuous_on (closure S) f" | |
| 3547 | and fim: "f ` S \<subseteq> T" | |
| 3548 | and fid: "\<And>x. x \<in> T \<Longrightarrow> f x = x" | |
| 3549 | shows "S \<subseteq> T" | |
| 3550 | proof (rule ccontr) | |
| 3551 | assume "\<not> S \<subseteq> T" | |
| 3552 | then obtain a where "a \<in> S" "a \<notin> T" by blast | |
| 3553 | define g where "g \<equiv> \<lambda>z. if z \<in> closure S then f z else z" | |
| 3554 | have "continuous_on (closure S \<union> closure(-S)) g" | |
| 3555 | unfolding g_def | |
| 3556 | apply (rule continuous_on_cases) | |
| 3557 | using fros fid frontier_closures | |
| 3558 | apply (auto simp: contf continuous_on_id) | |
| 3559 | done | |
| 3560 | moreover have "closure S \<union> closure(- S) = UNIV" | |
| 3561 | using closure_Un by fastforce | |
| 3562 | ultimately have contg: "continuous_on UNIV g" by metis | |
| 3563 | obtain B where "0 < B" and B: "closure S \<subseteq> ball a B" | |
| 3564 | using \<open>bounded S\<close> bounded_subset_ballD by blast | |
| 3565 | have notga: "g x \<noteq> a" for x | |
| 3566 | unfolding g_def using fros fim \<open>a \<notin> T\<close> | |
| 3567 | apply (auto simp: frontier_def) | |
| 3568 | using fid interior_subset apply fastforce | |
| 3569 | by (simp add: \<open>a \<in> S\<close> closure_def) | |
| 3570 | define h where "h \<equiv> (\<lambda>y. a + (B / norm(y - a)) *\<^sub>R (y - a)) \<circ> g" | |
| 3571 | have "\<not> (frontier (cball a B) retract_of (cball a B))" | |
| 3572 | by (metis no_retraction_cball \<open>0 < B\<close>) | |
| 3573 | then have "\<And>k. \<not> retraction (cball a B) (frontier (cball a B)) k" | |
| 3574 | by (simp add: retract_of_def) | |
| 3575 | moreover have "retraction (cball a B) (frontier (cball a B)) h" | |
| 3576 | unfolding retraction_def | |
| 3577 | proof (intro conjI ballI) | |
| 3578 | show "frontier (cball a B) \<subseteq> cball a B" | |
| 68022 | 3579 | by force | 
| 64122 | 3580 | show "continuous_on (cball a B) h" | 
| 3581 | unfolding h_def | |
| 68022 | 3582 | by (intro continuous_intros) (use contg continuous_on_subset notga in auto) | 
| 64122 | 3583 | show "h ` cball a B \<subseteq> frontier (cball a B)" | 
| 3584 | using \<open>0 < B\<close> by (auto simp: h_def notga dist_norm) | |
| 3585 | show "\<And>x. x \<in> frontier (cball a B) \<Longrightarrow> h x = x" | |
| 3586 | apply (auto simp: h_def algebra_simps) | |
| 3587 | apply (simp add: vector_add_divide_simps notga) | |
| 3588 | by (metis (no_types, hide_lams) B add.commute dist_commute dist_norm g_def mem_ball not_less_iff_gr_or_eq subset_eq) | |
| 3589 | qed | |
| 3590 | ultimately show False by simp | |
| 3591 | qed | |
| 3592 | ||
| 68617 | 3593 | subsubsection \<open>Punctured affine hulls, etc\<close> | 
| 3594 | ||
| 3595 | lemma rel_frontier_deformation_retract_of_punctured_convex: | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3596 | fixes S :: "'a::euclidean_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3597 | assumes "convex S" "convex T" "bounded S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3598 | and arelS: "a \<in> rel_interior S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3599 | and relS: "rel_frontier S \<subseteq> T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3600 | and affS: "T \<subseteq> affine hull S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3601 |   obtains r where "homotopic_with (\<lambda>x. True) (T - {a}) (T - {a}) id r"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3602 |                   "retraction (T - {a}) (rel_frontier S) r"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3603 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3604 | have "\<exists>d. 0 < d \<and> (a + d *\<^sub>R l) \<in> rel_frontier S \<and> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3605 | (\<forall>e. 0 \<le> e \<and> e < d \<longrightarrow> (a + e *\<^sub>R l) \<in> rel_interior S)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3606 | if "(a + l) \<in> affine hull S" "l \<noteq> 0" for l | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3607 | apply (rule ray_to_rel_frontier [OF \<open>bounded S\<close> arelS]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3608 | apply (rule that)+ | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3609 | by metis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3610 | then obtain dd | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3611 | 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" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3612 | and dd2: "\<And>l e. \<lbrakk>(a + l) \<in> affine hull S; e < dd l; 0 \<le> e; l \<noteq> 0\<rbrakk> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3613 | \<Longrightarrow> (a + e *\<^sub>R l) \<in> rel_interior S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3614 | by metis+ | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3615 | have aaffS: "a \<in> affine hull S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3616 | by (meson arelS subsetD hull_inc rel_interior_subset) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3617 |   have "((\<lambda>z. z - a) ` (affine hull S - {a})) = ((\<lambda>z. z - a) ` (affine hull S)) - {0}"
 | 
| 68022 | 3618 | by auto | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3619 |   moreover have "continuous_on (((\<lambda>z. z - a) ` (affine hull S)) - {0}) (\<lambda>x. dd x *\<^sub>R x)"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3620 | proof (rule continuous_on_compact_surface_projection) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3621 | show "compact (rel_frontier ((\<lambda>z. z - a) ` S))" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3622 | by (simp add: \<open>bounded S\<close> bounded_translation_minus compact_rel_frontier_bounded) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3623 | have releq: "rel_frontier ((\<lambda>z. z - a) ` S) = (\<lambda>z. z - a) ` rel_frontier S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3624 | using rel_frontier_translation [of "-a"] add.commute by simp | 
| 68022 | 3625 |     also have "\<dots> \<subseteq> (\<lambda>z. z - a) ` (affine hull S) - {0}"
 | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3626 | using rel_frontier_affine_hull arelS rel_frontier_def by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3627 |     finally show "rel_frontier ((\<lambda>z. z - a) ` S) \<subseteq> (\<lambda>z. z - a) ` (affine hull S) - {0}" .
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3628 | show "cone ((\<lambda>z. z - a) ` (affine hull S))" | 
| 68022 | 3629 | by (rule subspace_imp_cone) | 
| 3630 | (use aaffS in \<open>simp add: subspace_affine image_comp o_def affine_translation_aux [of a]\<close>) | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3631 | show "(0 < k \<and> k *\<^sub>R x \<in> rel_frontier ((\<lambda>z. z - a) ` S)) \<longleftrightarrow> (dd x = k)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3632 |          if x: "x \<in> (\<lambda>z. z - a) ` (affine hull S) - {0}" for k x
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3633 | proof | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3634 | show "dd x = k \<Longrightarrow> 0 < k \<and> k *\<^sub>R x \<in> rel_frontier ((\<lambda>z. z - a) ` S)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3635 | using dd1 [of x] that image_iff by (fastforce simp add: releq) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3636 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3637 | assume k: "0 < k \<and> k *\<^sub>R x \<in> rel_frontier ((\<lambda>z. z - a) ` S)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3638 | have False if "dd x < k" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3639 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3640 | have "k \<noteq> 0" "a + k *\<^sub>R x \<in> closure S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3641 | using k closure_translation [of "-a"] | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3642 | by (auto simp: rel_frontier_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3643 | then have segsub: "open_segment a (a + k *\<^sub>R x) \<subseteq> rel_interior S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3644 | by (metis rel_interior_closure_convex_segment [OF \<open>convex S\<close> arelS]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3645 | have "x \<noteq> 0" and xaffS: "a + x \<in> affine hull S" | 
| 68022 | 3646 | using x by auto | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3647 | then have "0 < dd x" and inS: "a + dd x *\<^sub>R x \<in> rel_frontier S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3648 | using dd1 by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3649 | moreover have "a + dd x *\<^sub>R x \<in> open_segment a (a + k *\<^sub>R x)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3650 | using k \<open>x \<noteq> 0\<close> \<open>0 < dd x\<close> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3651 | apply (simp add: in_segment) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3652 | apply (rule_tac x = "dd x / k" in exI) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3653 | apply (simp add: field_simps that) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3654 | apply (simp add: vector_add_divide_simps algebra_simps) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3655 | apply (metis (no_types) \<open>k \<noteq> 0\<close> divide_inverse_commute inverse_eq_divide mult.left_commute right_inverse) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3656 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3657 | ultimately show ?thesis | 
| 68022 | 3658 | using segsub by (auto simp: rel_frontier_def) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3659 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3660 | moreover have False if "k < dd x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3661 | using x k that rel_frontier_def | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3662 | by (fastforce simp: algebra_simps releq dest!: dd2) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3663 | ultimately show "dd x = k" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3664 | by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3665 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3666 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3667 |   ultimately have *: "continuous_on ((\<lambda>z. z - a) ` (affine hull S - {a})) (\<lambda>x. dd x *\<^sub>R x)"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3668 | by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3669 |   have "continuous_on (affine hull S - {a}) ((\<lambda>x. a + dd x *\<^sub>R x) \<circ> (\<lambda>z. z - a))"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3670 | by (intro * continuous_intros continuous_on_compose) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3671 |   with affS have contdd: "continuous_on (T - {a}) ((\<lambda>x. a + dd x *\<^sub>R x) \<circ> (\<lambda>z. z - a))"
 | 
| 68022 | 3672 | by (blast intro: continuous_on_subset) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3673 | show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3674 | proof | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3675 |     show "homotopic_with (\<lambda>x. True) (T - {a}) (T - {a}) id (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3676 | proof (rule homotopic_with_linear) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3677 |       show "continuous_on (T - {a}) id"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3678 | by (intro continuous_intros continuous_on_compose) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3679 |       show "continuous_on (T - {a}) (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3680 | using contdd by (simp add: o_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3681 |       show "closed_segment (id x) (a + dd (x - a) *\<^sub>R (x - a)) \<subseteq> T - {a}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3682 |            if "x \<in> T - {a}" for x
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3683 | proof (clarsimp simp: in_segment, intro conjI) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3684 | fix u::real assume u: "0 \<le> u" "u \<le> 1" | 
| 68022 | 3685 | have "a + dd (x - a) *\<^sub>R (x - a) \<in> T" | 
| 3686 | by (metis DiffD1 DiffD2 add.commute add.right_neutral affS dd1 diff_add_cancel relS singletonI subsetCE that) | |
| 3687 | then show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \<in> T" | |
| 3688 | using convexD [OF \<open>convex T\<close>] that u by simp | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3689 | have iff: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + d *\<^sub>R (x - a)) = a \<longleftrightarrow> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3690 | (1 - u + u * d) *\<^sub>R (x - a) = 0" for d | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3691 | by (auto simp: algebra_simps) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3692 | have "x \<in> T" "x \<noteq> a" using that by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3693 | then have axa: "a + (x - a) \<in> affine hull S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3694 | by (metis (no_types) add.commute affS diff_add_cancel set_rev_mp) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3695 | then have "\<not> dd (x - a) \<le> 0 \<and> a + dd (x - a) *\<^sub>R (x - a) \<in> rel_frontier S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3696 | using \<open>x \<noteq> a\<close> dd1 by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3697 | 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" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3698 | apply (auto simp: iff) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3699 | using less_eq_real_def mult_le_0_iff not_less u by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3700 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3701 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3702 |     show "retraction (T - {a}) (rel_frontier S) (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3703 | proof (simp add: retraction_def, intro conjI ballI) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3704 |       show "rel_frontier S \<subseteq> T - {a}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3705 | using arelS relS rel_frontier_def by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3706 |       show "continuous_on (T - {a}) (\<lambda>x. a + dd (x - a) *\<^sub>R (x - a))"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3707 | using contdd by (simp add: o_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3708 |       show "(\<lambda>x. a + dd (x - a) *\<^sub>R (x - a)) ` (T - {a}) \<subseteq> rel_frontier S"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3709 | apply (auto simp: rel_frontier_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3710 | apply (metis Diff_subset add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 rel_frontier_def subset_iff) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3711 | by (metis DiffE add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 rel_frontier_def rev_subsetD) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3712 | show "a + dd (x - a) *\<^sub>R (x - a) = x" if x: "x \<in> rel_frontier S" for x | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3713 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3714 | have "x \<noteq> a" | 
| 68022 | 3715 | using that arelS by (auto simp: rel_frontier_def) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3716 | have False if "dd (x - a) < 1" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3717 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3718 | have "x \<in> closure S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3719 | using x by (auto simp: rel_frontier_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3720 | then have segsub: "open_segment a x \<subseteq> rel_interior S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3721 | by (metis rel_interior_closure_convex_segment [OF \<open>convex S\<close> arelS]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3722 | have xaffS: "x \<in> affine hull S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3723 | using affS relS x by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3724 | then have "0 < dd (x - a)" and inS: "a + dd (x - a) *\<^sub>R (x - a) \<in> rel_frontier S" | 
| 68022 | 3725 | using dd1 by (auto simp: \<open>x \<noteq> a\<close>) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3726 | moreover have "a + dd (x - a) *\<^sub>R (x - a) \<in> open_segment a x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3727 | using \<open>x \<noteq> a\<close> \<open>0 < dd (x - a)\<close> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3728 | apply (simp add: in_segment) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3729 | apply (rule_tac x = "dd (x - a)" in exI) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3730 | apply (simp add: algebra_simps that) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3731 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3732 | ultimately show ?thesis | 
| 68022 | 3733 | using segsub by (auto simp: rel_frontier_def) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3734 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3735 | moreover have False if "1 < dd (x - a)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3736 | using x that dd2 [of "x - a" 1] \<open>x \<noteq> a\<close> closure_affine_hull | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3737 | by (auto simp: rel_frontier_def) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67399diff
changeset | 3738 | ultimately have "dd (x - a) = 1" \<comment> \<open>similar to another proof above\<close> | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3739 | by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3740 | with that show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3741 | by (simp add: rel_frontier_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3742 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3743 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3744 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3745 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3746 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3747 | corollary rel_frontier_retract_of_punctured_affine_hull: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3748 | fixes S :: "'a::euclidean_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3749 | assumes "bounded S" "convex S" "a \<in> rel_interior S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3750 |     shows "rel_frontier S retract_of (affine hull S - {a})"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3751 | apply (rule rel_frontier_deformation_retract_of_punctured_convex [of S "affine hull S" a]) | 
| 68022 | 3752 | apply (auto simp: affine_imp_convex rel_frontier_affine_hull retract_of_def assms) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3753 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3754 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3755 | corollary rel_boundary_retract_of_punctured_affine_hull: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3756 | fixes S :: "'a::euclidean_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3757 | assumes "compact S" "convex S" "a \<in> rel_interior S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3758 |     shows "(S - rel_interior S) retract_of (affine hull S - {a})"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3759 | by (metis assms closure_closed compact_eq_bounded_closed rel_frontier_def | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3760 | rel_frontier_retract_of_punctured_affine_hull) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3761 | |
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3762 | lemma homotopy_eqv_rel_frontier_punctured_convex: | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3763 | fixes S :: "'a::euclidean_space set" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3764 | assumes "convex S" "bounded S" "a \<in> rel_interior S" "convex T" "rel_frontier S \<subseteq> T" "T \<subseteq> affine hull S" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3765 |   shows "(rel_frontier S) homotopy_eqv (T - {a})"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3766 | apply (rule rel_frontier_deformation_retract_of_punctured_convex [of S T]) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3767 | using assms | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3768 | apply auto | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3769 | apply (subst homotopy_eqv_sym) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3770 | using deformation_retract_imp_homotopy_eqv by blast | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3771 | |
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3772 | lemma homotopy_eqv_rel_frontier_punctured_affine_hull: | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3773 | fixes S :: "'a::euclidean_space set" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3774 | assumes "convex S" "bounded S" "a \<in> rel_interior S" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3775 |     shows "(rel_frontier S) homotopy_eqv (affine hull S - {a})"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3776 | apply (rule homotopy_eqv_rel_frontier_punctured_convex) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3777 | using assms rel_frontier_affine_hull by force+ | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 3778 | |
| 64394 | 3779 | lemma path_connected_sphere_gen: | 
| 3780 | assumes "convex S" "bounded S" "aff_dim S \<noteq> 1" | |
| 3781 | shows "path_connected(rel_frontier S)" | |
| 3782 | proof (cases "rel_interior S = {}")
 | |
| 3783 | case True | |
| 3784 | then show ?thesis | |
| 3785 | by (simp add: \<open>convex S\<close> convex_imp_path_connected rel_frontier_def) | |
| 3786 | next | |
| 3787 | case False | |
| 3788 | then show ?thesis | |
| 3789 | by (metis aff_dim_affine_hull affine_affine_hull affine_imp_convex all_not_in_conv assms path_connected_punctured_convex rel_frontier_retract_of_punctured_affine_hull retract_of_path_connected) | |
| 3790 | qed | |
| 3791 | ||
| 3792 | lemma connected_sphere_gen: | |
| 3793 | assumes "convex S" "bounded S" "aff_dim S \<noteq> 1" | |
| 3794 | shows "connected(rel_frontier S)" | |
| 3795 | by (simp add: assms path_connected_imp_connected path_connected_sphere_gen) | |
| 3796 | ||
| 68617 | 3797 | subsubsection\<open>Borsuk-style characterization of separation\<close> | 
| 63301 | 3798 | |
| 3799 | lemma continuous_on_Borsuk_map: | |
| 3800 | "a \<notin> s \<Longrightarrow> continuous_on s (\<lambda>x. inverse(norm (x - a)) *\<^sub>R (x - a))" | |
| 3801 | by (rule continuous_intros | force)+ | |
| 3802 | ||
| 3803 | lemma Borsuk_map_into_sphere: | |
| 3804 | "(\<lambda>x. inverse(norm (x - a)) *\<^sub>R (x - a)) ` s \<subseteq> sphere 0 1 \<longleftrightarrow> (a \<notin> s)" | |
| 3805 | by auto (metis eq_iff_diff_eq_0 left_inverse norm_eq_zero) | |
| 3806 | ||
| 3807 | lemma Borsuk_maps_homotopic_in_path_component: | |
| 3808 | assumes "path_component (- s) a b" | |
| 3809 | shows "homotopic_with (\<lambda>x. True) s (sphere 0 1) | |
| 3810 | (\<lambda>x. inverse(norm(x - a)) *\<^sub>R (x - a)) | |
| 3811 | (\<lambda>x. inverse(norm(x - b)) *\<^sub>R (x - b))" | |
| 3812 | proof - | |
| 3813 | obtain g where "path g" "path_image g \<subseteq> -s" "pathstart g = a" "pathfinish g = b" | |
| 3814 | using assms by (auto simp: path_component_def) | |
| 3815 | then show ?thesis | |
| 3816 | apply (simp add: path_def path_image_def pathstart_def pathfinish_def homotopic_with_def) | |
| 68022 | 3817 | apply (rule_tac x = "\<lambda>z. inverse(norm(snd z - (g \<circ> fst)z)) *\<^sub>R (snd z - (g \<circ> fst)z)" in exI) | 
| 63301 | 3818 | apply (intro conjI continuous_intros) | 
| 3819 | apply (rule continuous_intros | erule continuous_on_subset | fastforce simp: divide_simps sphere_def)+ | |
| 3820 | done | |
| 3821 | qed | |
| 3822 | ||
| 3823 | lemma non_extensible_Borsuk_map: | |
| 3824 | fixes a :: "'a :: euclidean_space" | |
| 3825 | assumes "compact s" and cin: "c \<in> components(- s)" and boc: "bounded c" and "a \<in> c" | |
| 3826 | shows "~ (\<exists>g. continuous_on (s \<union> c) g \<and> | |
| 3827 | g ` (s \<union> c) \<subseteq> sphere 0 1 \<and> | |
| 3828 | (\<forall>x \<in> s. g x = inverse(norm(x - a)) *\<^sub>R (x - a)))" | |
| 3829 | proof - | |
| 3830 | have "closed s" using assms by (simp add: compact_imp_closed) | |
| 3831 | have "c \<subseteq> -s" | |
| 3832 | using assms by (simp add: in_components_subset) | |
| 3833 | with \<open>a \<in> c\<close> have "a \<notin> s" by blast | |
| 3834 | then have ceq: "c = connected_component_set (- s) a" | |
| 3835 | by (metis \<open>a \<in> c\<close> cin components_iff connected_component_eq) | |
| 3836 | then have "bounded (s \<union> connected_component_set (- s) a)" | |
| 3837 | using \<open>compact s\<close> boc compact_imp_bounded by auto | |
| 3838 | with bounded_subset_ballD obtain r where "0 < r" and r: "(s \<union> connected_component_set (- s) a) \<subseteq> ball a r" | |
| 3839 | by blast | |
| 3840 |   { fix g
 | |
| 3841 | assume "continuous_on (s \<union> c) g" | |
| 3842 | "g ` (s \<union> c) \<subseteq> sphere 0 1" | |
| 3843 | and [simp]: "\<And>x. x \<in> s \<Longrightarrow> g x = (x - a) /\<^sub>R norm (x - a)" | |
| 3844 | then have [simp]: "\<And>x. x \<in> s \<union> c \<Longrightarrow> norm (g x) = 1" | |
| 3845 | by force | |
| 3846 | have cb_eq: "cball a r = (s \<union> connected_component_set (- s) a) \<union> | |
| 3847 | (cball a r - connected_component_set (- s) a)" | |
| 3848 | using ball_subset_cball [of a r] r by auto | |
| 3849 | have cont1: "continuous_on (s \<union> connected_component_set (- s) a) | |
| 3850 | (\<lambda>x. a + r *\<^sub>R g x)" | |
| 3851 | apply (rule continuous_intros)+ | |
| 3852 | using \<open>continuous_on (s \<union> c) g\<close> ceq by blast | |
| 3853 | have cont2: "continuous_on (cball a r - connected_component_set (- s) a) | |
| 3854 | (\<lambda>x. a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" | |
| 3855 | by (rule continuous_intros | force simp: \<open>a \<notin> s\<close>)+ | |
| 3856 | have 1: "continuous_on (cball a r) | |
| 3857 | (\<lambda>x. if connected_component (- s) a x | |
| 3858 | then a + r *\<^sub>R g x | |
| 3859 | else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" | |
| 3860 | apply (subst cb_eq) | |
| 3861 | apply (rule continuous_on_cases [OF _ _ cont1 cont2]) | |
| 3862 | using ceq cin | |
| 3863 | apply (auto intro: closed_Un_complement_component | |
| 3864 | simp: \<open>closed s\<close> open_Compl open_connected_component) | |
| 3865 | done | |
| 3866 | have 2: "(\<lambda>x. a + r *\<^sub>R g x) ` (cball a r \<inter> connected_component_set (- s) a) | |
| 3867 | \<subseteq> sphere a r " | |
| 3868 | using \<open>0 < r\<close> by (force simp: dist_norm ceq) | |
| 3869 | have "retraction (cball a r) (sphere a r) | |
| 3870 | (\<lambda>x. if x \<in> connected_component_set (- s) a | |
| 3871 | then a + r *\<^sub>R g x | |
| 3872 | else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" | |
| 3873 | using \<open>0 < r\<close> | |
| 3874 | apply (simp add: retraction_def dist_norm 1 2, safe) | |
| 3875 | apply (force simp: dist_norm abs_if mult_less_0_iff divide_simps \<open>a \<notin> s\<close>) | |
| 3876 | using r | |
| 3877 | by (auto simp: dist_norm norm_minus_commute) | |
| 3878 | then have False | |
| 3879 | using no_retraction_cball | |
| 3880 | [OF \<open>0 < r\<close>, of a, unfolded retract_of_def, simplified, rule_format, | |
| 3881 | of "\<lambda>x. if x \<in> connected_component_set (- s) a | |
| 3882 | then a + r *\<^sub>R g x | |
| 3883 | else a + r *\<^sub>R inverse(norm(x - a)) *\<^sub>R (x - a)"] | |
| 3884 | by blast | |
| 3885 | } | |
| 3886 | then show ?thesis | |
| 3887 | by blast | |
| 3888 | qed | |
| 3889 | ||
| 68617 | 3890 | subsubsection \<open>We continue with ANRs and ENRs\<close> | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3891 | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3892 | lemma ENR_rel_frontier_convex: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3893 | fixes S :: "'a::euclidean_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3894 | assumes "bounded S" "convex S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3895 | shows "ENR(rel_frontier S)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3896 | proof (cases "S = {}")
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3897 | case True then show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3898 | by simp | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3899 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3900 | case False | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3901 |   with assms have "rel_interior S \<noteq> {}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3902 | by (simp add: rel_interior_eq_empty) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3903 | then obtain a where a: "a \<in> rel_interior S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3904 | by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3905 |   have ahS: "affine hull S - {a} \<subseteq> {x. closest_point (affine hull S) x \<noteq> a}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3906 | by (auto simp: closest_point_self) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3907 |   have "rel_frontier S retract_of affine hull S - {a}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3908 | by (simp add: assms a rel_frontier_retract_of_punctured_affine_hull) | 
| 68022 | 3909 |   also have "\<dots> retract_of {x. closest_point (affine hull S) x \<noteq> a}"
 | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3910 | apply (simp add: retract_of_def retraction_def ahS) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3911 | apply (rule_tac x="closest_point (affine hull S)" in exI) | 
| 68022 | 3912 | apply (auto simp: False closest_point_self affine_imp_convex closest_point_in_set continuous_on_closest_point) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3913 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3914 |   finally have "rel_frontier S retract_of {x. closest_point (affine hull S) x \<noteq> a}" .
 | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3915 |   moreover have "openin (subtopology euclidean UNIV) (UNIV \<inter> closest_point (affine hull S) -` (- {a}))"
 | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3916 | apply (rule continuous_openin_preimage_gen) | 
| 68022 | 3917 | apply (auto simp: False affine_imp_convex continuous_on_closest_point) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3918 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3919 | ultimately show ?thesis | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3920 | unfolding ENR_def | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3921 |     apply (rule_tac x = "closest_point (affine hull S) -` (- {a})" in exI)
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3922 | apply (simp add: vimage_def) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3923 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3924 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3925 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3926 | lemma ANR_rel_frontier_convex: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3927 | fixes S :: "'a::euclidean_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3928 | assumes "bounded S" "convex S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3929 | shows "ANR(rel_frontier S)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 3930 | by (simp add: ENR_imp_ANR ENR_rel_frontier_convex assms) | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3931 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3932 | lemma ENR_closedin_Un_local: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3933 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3934 | shows "\<lbrakk>ENR S; ENR T; ENR(S \<inter> T); | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3935 | closedin (subtopology euclidean (S \<union> T)) S; closedin (subtopology euclidean (S \<union> T)) T\<rbrakk> | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3936 | \<Longrightarrow> ENR(S \<union> T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3937 | by (simp add: ENR_ANR ANR_closed_Un_local locally_compact_closedin_Un) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3938 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3939 | lemma ENR_closed_Un: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3940 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3941 | shows "\<lbrakk>closed S; closed T; ENR S; ENR T; ENR(S \<inter> T)\<rbrakk> \<Longrightarrow> ENR(S \<union> T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3942 | by (auto simp: closed_subset ENR_closedin_Un_local) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3943 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3944 | lemma absolute_retract_Un: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3945 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3946 | shows "\<lbrakk>S retract_of UNIV; T retract_of UNIV; (S \<inter> T) retract_of UNIV\<rbrakk> | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3947 | \<Longrightarrow> (S \<union> T) retract_of UNIV" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3948 | by (meson AR_closed_Un_local_aux closed_subset retract_of_UNIV retract_of_imp_subset) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3949 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3950 | lemma retract_from_Un_Int: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3951 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3952 | assumes clS: "closedin (subtopology euclidean (S \<union> T)) S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3953 | and clT: "closedin (subtopology euclidean (S \<union> T)) T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3954 | and Un: "(S \<union> T) retract_of U" and Int: "(S \<inter> T) retract_of T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3955 | shows "S retract_of U" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3956 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3957 | obtain r where r: "continuous_on T r" "r ` T \<subseteq> S \<inter> T" "\<forall>x\<in>S \<inter> T. r x = x" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3958 | using Int by (auto simp: retraction_def retract_of_def) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3959 | have "S retract_of S \<union> T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3960 | unfolding retraction_def retract_of_def | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3961 | proof (intro exI conjI) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3962 | show "continuous_on (S \<union> T) (\<lambda>x. if x \<in> S then x else r x)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3963 | apply (rule continuous_on_cases_local [OF clS clT]) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3964 | using r by (auto simp: continuous_on_id) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3965 | qed (use r in auto) | 
| 68022 | 3966 | also have "\<dots> retract_of U" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3967 | by (rule Un) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3968 | finally show ?thesis . | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3969 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3970 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3971 | lemma AR_from_Un_Int_local: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3972 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3973 | assumes clS: "closedin (subtopology euclidean (S \<union> T)) S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3974 | and clT: "closedin (subtopology euclidean (S \<union> T)) T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3975 | and Un: "AR(S \<union> T)" and Int: "AR(S \<inter> T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3976 | shows "AR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3977 | apply (rule AR_retract_of_AR [OF Un]) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3978 | by (meson AR_imp_retract clS clT closedin_closed_subset local.Int retract_from_Un_Int retract_of_refl sup_ge2) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3979 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3980 | lemma AR_from_Un_Int_local': | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3981 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3982 | assumes "closedin (subtopology euclidean (S \<union> T)) S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3983 | and "closedin (subtopology euclidean (S \<union> T)) T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3984 | and "AR(S \<union> T)" "AR(S \<inter> T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3985 | shows "AR T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3986 | using AR_from_Un_Int_local [of T S] assms by (simp add: Un_commute Int_commute) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3987 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3988 | lemma AR_from_Un_Int: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3989 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3990 | assumes clo: "closed S" "closed T" and Un: "AR(S \<union> T)" and Int: "AR(S \<inter> T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3991 | shows "AR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3992 | by (metis AR_from_Un_Int_local [OF _ _ Un Int] Un_commute clo closed_closedin closedin_closed_subset inf_sup_absorb subtopology_UNIV top_greatest) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3993 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3994 | lemma ANR_from_Un_Int_local: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3995 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3996 | assumes clS: "closedin (subtopology euclidean (S \<union> T)) S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3997 | and clT: "closedin (subtopology euclidean (S \<union> T)) T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3998 | and Un: "ANR(S \<union> T)" and Int: "ANR(S \<inter> T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 3999 | shows "ANR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4000 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4001 | obtain V where clo: "closedin (subtopology euclidean (S \<union> T)) (S \<inter> T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4002 | and ope: "openin (subtopology euclidean (S \<union> T)) V" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4003 | and ret: "S \<inter> T retract_of V" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4004 | using ANR_imp_neighbourhood_retract [OF Int] by (metis clS clT closedin_Int) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4005 | then obtain r where r: "continuous_on V r" and rim: "r ` V \<subseteq> S \<inter> T" and req: "\<forall>x\<in>S \<inter> T. r x = x" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4006 | by (auto simp: retraction_def retract_of_def) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4007 | have Vsub: "V \<subseteq> S \<union> T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4008 | by (meson ope openin_contains_cball) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4009 | have Vsup: "S \<inter> T \<subseteq> V" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4010 | by (simp add: retract_of_imp_subset ret) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4011 | then have eq: "S \<union> V = ((S \<union> T) - T) \<union> V" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4012 | by auto | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4013 | have eq': "S \<union> V = S \<union> (V \<inter> T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4014 | using Vsub by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4015 | have "continuous_on (S \<union> V \<inter> T) (\<lambda>x. if x \<in> S then x else r x)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4016 | proof (rule continuous_on_cases_local) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4017 | show "closedin (subtopology euclidean (S \<union> V \<inter> T)) S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4018 | using clS closedin_subset_trans inf.boundedE by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4019 | show "closedin (subtopology euclidean (S \<union> V \<inter> T)) (V \<inter> T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4020 | using clT Vsup by (auto simp: closedin_closed) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4021 | show "continuous_on (V \<inter> T) r" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4022 | by (meson Int_lower1 continuous_on_subset r) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4023 | qed (use req continuous_on_id in auto) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4024 | with rim have "S retract_of S \<union> V" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4025 | unfolding retraction_def retract_of_def | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4026 | apply (rule_tac x="\<lambda>x. if x \<in> S then x else r x" in exI) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4027 | apply (auto simp: eq') | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4028 | done | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4029 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4030 | using ANR_neighborhood_retract [OF Un] | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4031 | using \<open>S \<union> V = S \<union> T - T \<union> V\<close> clT ope by fastforce | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4032 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4033 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4034 | lemma ANR_from_Un_Int: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4035 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4036 | assumes clo: "closed S" "closed T" and Un: "ANR(S \<union> T)" and Int: "ANR(S \<inter> T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4037 | shows "ANR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4038 | by (metis ANR_from_Un_Int_local [OF _ _ Un Int] Un_commute clo closed_closedin closedin_closed_subset inf_sup_absorb subtopology_UNIV top_greatest) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4039 | |
| 68617 | 4040 | lemma ANR_finite_Union_convex_closed: | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4041 | fixes \<T> :: "'a::euclidean_space set set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4042 | assumes \<T>: "finite \<T>" and clo: "\<And>C. C \<in> \<T> \<Longrightarrow> closed C" and con: "\<And>C. C \<in> \<T> \<Longrightarrow> convex C" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4043 | shows "ANR(\<Union>\<T>)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4044 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4045 | have "ANR(\<Union>\<T>)" if "card \<T> < n" for n | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4046 | using assms that | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4047 | proof (induction n arbitrary: \<T>) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4048 | case 0 then show ?case by simp | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4049 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4050 | case (Suc n) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4051 | have "ANR(\<Union>\<U>)" if "finite \<U>" "\<U> \<subseteq> \<T>" for \<U> | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4052 | using that | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4053 | proof (induction \<U>) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4054 | case empty | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4055 | then show ?case by simp | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4056 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4057 | case (insert C \<U>) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4058 | have "ANR (C \<union> \<Union>\<U>)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4059 | proof (rule ANR_closed_Un) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4060 | show "ANR (C \<inter> \<Union>\<U>)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4061 | unfolding Int_Union | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4062 | proof (rule Suc) | 
| 67399 | 4063 | show "finite ((\<inter>) C ` \<U>)" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4064 | by (simp add: insert.hyps(1)) | 
| 67399 | 4065 | show "\<And>Ca. Ca \<in> (\<inter>) C ` \<U> \<Longrightarrow> closed Ca" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4066 | by (metis (no_types, hide_lams) Suc.prems(2) closed_Int subsetD imageE insert.prems insertI1 insertI2) | 
| 67399 | 4067 | show "\<And>Ca. Ca \<in> (\<inter>) C ` \<U> \<Longrightarrow> convex Ca" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4068 | by (metis (mono_tags, lifting) Suc.prems(3) convex_Int imageE insert.prems insert_subset subsetCE) | 
| 67399 | 4069 | show "card ((\<inter>) C ` \<U>) < n" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4070 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4071 | have "card \<T> \<le> n" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4072 | by (meson Suc.prems(4) not_less not_less_eq) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4073 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4074 | by (metis Suc.prems(1) card_image_le card_seteq insert.hyps insert.prems insert_subset le_trans not_less) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4075 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4076 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4077 | show "closed (\<Union>\<U>)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4078 | using Suc.prems(2) insert.hyps(1) insert.prems by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4079 | qed (use Suc.prems convex_imp_ANR insert.prems insert.IH in auto) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4080 | then show ?case | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4081 | by simp | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4082 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4083 | then show ?case | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4084 | using Suc.prems(1) by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4085 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4086 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4087 | by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4088 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4089 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4090 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4091 | lemma finite_imp_ANR: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4092 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4093 | assumes "finite S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4094 | shows "ANR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4095 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4096 |   have "ANR(\<Union>x \<in> S. {x})"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4097 | by (blast intro: ANR_finite_Union_convex_closed assms) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4098 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4099 | by simp | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4100 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4101 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4102 | lemma ANR_insert: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4103 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4104 | assumes "ANR S" "closed S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4105 | shows "ANR(insert a S)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4106 | by (metis ANR_closed_Un ANR_empty ANR_singleton Diff_disjoint Diff_insert_absorb assms closed_singleton insert_absorb insert_is_Un) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4107 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4108 | lemma ANR_path_component_ANR: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4109 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4110 | shows "ANR S \<Longrightarrow> ANR(path_component_set S x)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4111 | using ANR_imp_locally_path_connected ANR_openin openin_path_component_locally_path_connected by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4112 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4113 | lemma ANR_connected_component_ANR: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4114 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4115 | shows "ANR S \<Longrightarrow> ANR(connected_component_set S x)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4116 | by (metis ANR_openin openin_connected_component_locally_connected ANR_imp_locally_connected) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4117 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4118 | lemma ANR_component_ANR: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4119 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4120 | assumes "ANR S" "c \<in> components S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4121 | shows "ANR c" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4122 | by (metis ANR_connected_component_ANR assms componentsE) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4123 | |
| 68617 | 4124 | subsubsection\<open>Original ANR material, now for ENRs\<close> | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4125 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4126 | lemma ENR_bounded: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4127 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4128 | assumes "bounded S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4129 | shows "ENR S \<longleftrightarrow> (\<exists>U. open U \<and> bounded U \<and> S retract_of U)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4130 | (is "?lhs = ?rhs") | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4131 | proof | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4132 | obtain r where "0 < r" and r: "S \<subseteq> ball 0 r" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4133 | using bounded_subset_ballD assms by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4134 | assume ?lhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4135 | then show ?rhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4136 | apply (clarsimp simp: ENR_def) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4137 | apply (rule_tac x="ball 0 r \<inter> U" in exI, auto) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4138 | using r retract_of_imp_subset retract_of_subset by fastforce | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4139 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4140 | assume ?rhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4141 | then show ?lhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4142 | using ENR_def by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4143 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4144 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4145 | lemma absolute_retract_imp_AR_gen: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4146 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4147 |   assumes "S retract_of T" "convex T" "T \<noteq> {}" "S homeomorphic S'" "closedin (subtopology euclidean U) S'"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4148 | shows "S' retract_of U" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4149 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4150 | have "AR T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4151 | by (simp add: assms convex_imp_AR) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4152 | then have "AR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4153 | using AR_retract_of_AR assms by auto | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4154 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4155 | using assms AR_imp_absolute_retract by metis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4156 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4157 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4158 | lemma absolute_retract_imp_AR: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4159 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4160 | assumes "S retract_of UNIV" "S homeomorphic S'" "closed S'" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4161 | shows "S' retract_of UNIV" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4162 | using AR_imp_absolute_retract_UNIV assms retract_of_UNIV by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4163 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4164 | lemma homeomorphic_compact_arness: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4165 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4166 | assumes "S homeomorphic S'" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4167 | shows "compact S \<and> S retract_of UNIV \<longleftrightarrow> compact S' \<and> S' retract_of UNIV" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4168 | using assms homeomorphic_compactness | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4169 | apply auto | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4170 | apply (meson assms compact_AR homeomorphic_AR_iff_AR homeomorphic_compactness)+ | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4171 | done | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4172 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4173 | lemma absolute_retract_from_Un_Int: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4174 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4175 | assumes "(S \<union> T) retract_of UNIV" "(S \<inter> T) retract_of UNIV" "closed S" "closed T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4176 | shows "S retract_of UNIV" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4177 | using AR_from_Un_Int assms retract_of_UNIV by auto | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4178 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4179 | lemma ENR_from_Un_Int_gen: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4180 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4181 | assumes "closedin (subtopology euclidean (S \<union> T)) S" "closedin (subtopology euclidean (S \<union> T)) T" "ENR(S \<union> T)" "ENR(S \<inter> T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4182 | shows "ENR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4183 | apply (simp add: ENR_ANR) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4184 | using ANR_from_Un_Int_local ENR_ANR assms locally_compact_closedin by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4185 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4186 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4187 | lemma ENR_from_Un_Int: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4188 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4189 | assumes "closed S" "closed T" "ENR(S \<union> T)" "ENR(S \<inter> T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4190 | shows "ENR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4191 | by (meson ENR_from_Un_Int_gen assms closed_subset sup_ge1 sup_ge2) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4192 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4193 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4194 | lemma ENR_finite_Union_convex_closed: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4195 | fixes \<T> :: "'a::euclidean_space set set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4196 | assumes \<T>: "finite \<T>" and clo: "\<And>C. C \<in> \<T> \<Longrightarrow> closed C" and con: "\<And>C. C \<in> \<T> \<Longrightarrow> convex C" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4197 | shows "ENR(\<Union> \<T>)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4198 | by (simp add: ENR_ANR ANR_finite_Union_convex_closed \<T> clo closed_Union closed_imp_locally_compact con) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4199 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4200 | lemma finite_imp_ENR: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4201 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4202 | shows "finite S \<Longrightarrow> ENR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4203 | by (simp add: ENR_ANR finite_imp_ANR finite_imp_closed closed_imp_locally_compact) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4204 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4205 | lemma ENR_insert: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4206 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4207 | assumes "closed S" "ENR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4208 | shows "ENR(insert a S)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4209 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4210 |   have "ENR ({a} \<union> S)"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4211 | by (metis ANR_insert ENR_ANR Un_commute Un_insert_right assms closed_imp_locally_compact closed_insert sup_bot_right) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4212 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4213 | by auto | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4214 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4215 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4216 | lemma ENR_path_component_ENR: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4217 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4218 | assumes "ENR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4219 | shows "ENR(path_component_set S x)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4220 | by (metis ANR_imp_locally_path_connected ENR_empty ENR_imp_ANR ENR_openin assms | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4221 | locally_path_connected_2 openin_subtopology_self path_component_eq_empty) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4222 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4223 | (*UNUSED | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4224 | lemma ENR_Times: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4225 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4226 | assumes "ENR S" "ENR T" shows "ENR(S \<times> T)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4227 | using assms apply (simp add: ENR_ANR ANR_Times) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4228 | thm locally_compact_Times | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4229 | oops | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4230 | SIMP_TAC[ENR_ANR; ANR_PCROSS; LOCALLY_COMPACT_PCROSS]);; | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4231 | *) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4232 | |
| 68617 | 4233 | subsubsection\<open>Finally, spheres are ANRs and ENRs\<close> | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4234 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4235 | lemma absolute_retract_homeomorphic_convex_compact: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4236 | fixes S :: "'a::euclidean_space set" and U :: "'b::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4237 |   assumes "S homeomorphic U" "S \<noteq> {}" "S \<subseteq> T" "convex U" "compact U"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4238 | shows "S retract_of T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4239 | by (metis UNIV_I assms compact_AR convex_imp_AR homeomorphic_AR_iff_AR homeomorphic_compactness homeomorphic_empty(1) retract_of_subset subsetI) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4240 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4241 | lemma frontier_retract_of_punctured_universe: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4242 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4243 | assumes "convex S" "bounded S" "a \<in> interior S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4244 |   shows "(frontier S) retract_of (- {a})"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4245 | using rel_frontier_retract_of_punctured_affine_hull | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4246 | by (metis Compl_eq_Diff_UNIV affine_hull_nonempty_interior assms empty_iff rel_frontier_frontier rel_interior_nonempty_interior) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4247 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4248 | lemma sphere_retract_of_punctured_universe_gen: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4249 | fixes a :: "'a::euclidean_space" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4250 | assumes "b \<in> ball a r" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4251 |   shows  "sphere a r retract_of (- {b})"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4252 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4253 |   have "frontier (cball a r) retract_of (- {b})"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4254 | apply (rule frontier_retract_of_punctured_universe) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4255 | using assms by auto | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4256 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4257 | by simp | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4258 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4259 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4260 | lemma sphere_retract_of_punctured_universe: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4261 | fixes a :: "'a::euclidean_space" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4262 | assumes "0 < r" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4263 |   shows "sphere a r retract_of (- {a})"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4264 | by (simp add: assms sphere_retract_of_punctured_universe_gen) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4265 | |
| 68617 | 4266 | lemma ENR_sphere: | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4267 | fixes a :: "'a::euclidean_space" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4268 | shows "ENR(sphere a r)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4269 | proof (cases "0 < r") | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4270 | case True | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4271 |   then have "sphere a r retract_of -{a}"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4272 | by (simp add: sphere_retract_of_punctured_universe) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4273 | with open_delete show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4274 | by (auto simp: ENR_def) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4275 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4276 | case False | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4277 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4278 | using finite_imp_ENR | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4279 | by (metis finite_insert infinite_imp_nonempty less_linear sphere_eq_empty sphere_trivial) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4280 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4281 | |
| 68617 | 4282 | corollary%unimportant ANR_sphere: | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4283 | fixes a :: "'a::euclidean_space" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4284 | shows "ANR(sphere a r)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4285 | by (simp add: ENR_imp_ANR ENR_sphere) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4286 | |
| 68617 | 4287 | subsubsection\<open>Spheres are connected, etc\<close> | 
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4288 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4289 | lemma locally_path_connected_sphere_gen: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4290 | fixes S :: "'a::euclidean_space set" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4291 | assumes "bounded S" and "convex S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4292 | shows "locally path_connected (rel_frontier S)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4293 | proof (cases "rel_interior S = {}")
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4294 | case True | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4295 | with assms show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4296 | by (simp add: rel_interior_eq_empty) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4297 | next | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4298 | case False | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4299 | then obtain a where a: "a \<in> rel_interior S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4300 | by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4301 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4302 | proof (rule retract_of_locally_path_connected) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4303 |     show "locally path_connected (affine hull S - {a})"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4304 | by (meson convex_affine_hull convex_imp_locally_path_connected locally_open_subset openin_delete openin_subtopology_self) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4305 |     show "rel_frontier S retract_of affine hull S - {a}"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4306 | using a assms rel_frontier_retract_of_punctured_affine_hull by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4307 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4308 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4309 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4310 | lemma locally_connected_sphere_gen: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4311 | fixes S :: "'a::euclidean_space set" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4312 | assumes "bounded S" and "convex S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4313 | shows "locally connected (rel_frontier S)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4314 | by (simp add: ANR_imp_locally_connected ANR_rel_frontier_convex assms) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4315 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4316 | lemma locally_path_connected_sphere: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4317 | fixes a :: "'a::euclidean_space" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4318 | shows "locally path_connected (sphere a r)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4319 | using ENR_imp_locally_path_connected ENR_sphere by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4320 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4321 | lemma locally_connected_sphere: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4322 | fixes a :: "'a::euclidean_space" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4323 | shows "locally connected(sphere a r)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4324 | using ANR_imp_locally_connected ANR_sphere by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4325 | |
| 68617 | 4326 | subsubsection\<open>Borsuk homotopy extension theorem\<close> | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4327 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4328 | text\<open>It's only this late so we can use the concept of retraction, | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4329 | saying that the domain sets or range set are ENRs.\<close> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4330 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4331 | theorem Borsuk_homotopy_extension_homotopic: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4332 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4333 | assumes cloTS: "closedin (subtopology euclidean T) S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4334 | and anr: "(ANR S \<and> ANR T) \<or> ANR U" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4335 | and contf: "continuous_on T f" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4336 | and "f ` T \<subseteq> U" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4337 | and "homotopic_with (\<lambda>x. True) S U f g" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4338 | obtains g' where "homotopic_with (\<lambda>x. True) T U f g'" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4339 | "continuous_on T g'" "image g' T \<subseteq> U" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4340 | "\<And>x. x \<in> S \<Longrightarrow> g' x = g x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4341 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4342 | have "S \<subseteq> T" using assms closedin_imp_subset by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4343 |   obtain h where conth: "continuous_on ({0..1} \<times> S) h"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4344 |              and him: "h ` ({0..1} \<times> S) \<subseteq> U"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4345 | and [simp]: "\<And>x. h(0, x) = f x" "\<And>x. h(1::real, x) = g x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4346 | using assms by (auto simp: homotopic_with_def) | 
| 68022 | 4347 | define h' where "h' \<equiv> \<lambda>z. if snd z \<in> S then h z else (f \<circ> snd) z" | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4348 |   define B where "B \<equiv> {0::real} \<times> T \<union> {0..1} \<times> S"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4349 |   have clo0T: "closedin (subtopology euclidean ({0..1} \<times> T)) ({0::real} \<times> T)"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4350 | by (simp add: closedin_subtopology_refl closedin_Times) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4351 |   moreover have cloT1S: "closedin (subtopology euclidean ({0..1} \<times> T)) ({0..1} \<times> S)"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4352 | by (simp add: closedin_subtopology_refl closedin_Times assms) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4353 |   ultimately have clo0TB:"closedin (subtopology euclidean ({0..1} \<times> T)) B"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4354 | by (auto simp: B_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4355 |   have cloBS: "closedin (subtopology euclidean B) ({0..1} \<times> S)"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4356 | by (metis (no_types) Un_subset_iff B_def closedin_subset_trans [OF cloT1S] clo0TB closedin_imp_subset closedin_self) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4357 |   moreover have cloBT: "closedin (subtopology euclidean B) ({0} \<times> T)"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4358 | using \<open>S \<subseteq> T\<close> closedin_subset_trans [OF clo0T] | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4359 | by (metis B_def Un_upper1 clo0TB closedin_closed inf_le1) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4360 |   moreover have "continuous_on ({0} \<times> T) (f \<circ> snd)"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4361 | apply (rule continuous_intros)+ | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4362 | apply (simp add: contf) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4363 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4364 | ultimately have conth': "continuous_on B h'" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4365 |     apply (simp add: h'_def B_def Un_commute [of "{0} \<times> T"])
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4366 | apply (auto intro!: continuous_on_cases_local conth) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4367 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4368 | have "image h' B \<subseteq> U" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4369 | using \<open>f ` T \<subseteq> U\<close> him by (auto simp: h'_def B_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4370 |   obtain V k where "B \<subseteq> V" and opeTV: "openin (subtopology euclidean ({0..1} \<times> T)) V"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4371 | and contk: "continuous_on V k" and kim: "k ` V \<subseteq> U" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4372 | and keq: "\<And>x. x \<in> B \<Longrightarrow> k x = h' x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4373 | using anr | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4374 | proof | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4375 | assume ST: "ANR S \<and> ANR T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4376 |     have eq: "({0} \<times> T \<inter> {0..1} \<times> S) = {0::real} \<times> S"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4377 | using \<open>S \<subseteq> T\<close> by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4378 | have "ANR B" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4379 | apply (simp add: B_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4380 | apply (rule ANR_closed_Un_local) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4381 | apply (metis cloBT B_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4382 | apply (metis Un_commute cloBS B_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4383 | apply (simp_all add: ANR_Times convex_imp_ANR ANR_singleton ST eq) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4384 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4385 | note Vk = that | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4386 |     have *: thesis if "openin (subtopology euclidean ({0..1::real} \<times> T)) V"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4387 | "retraction V B r" for V r | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4388 | using that | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4389 | apply (clarsimp simp add: retraction_def) | 
| 68022 | 4390 | apply (rule Vk [of V "h' \<circ> r"], assumption+) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4391 | apply (metis continuous_on_compose conth' continuous_on_subset) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4392 | using \<open>h' ` B \<subseteq> U\<close> apply force+ | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4393 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4394 | show thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4395 | apply (rule ANR_imp_neighbourhood_retract [OF \<open>ANR B\<close> clo0TB]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4396 | apply (auto simp: ANR_Times ANR_singleton ST retract_of_def *) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4397 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4398 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4399 | assume "ANR U" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4400 | with ANR_imp_absolute_neighbourhood_extensor \<open>h' ` B \<subseteq> U\<close> clo0TB conth' that | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4401 | show ?thesis by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4402 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4403 |   define S' where "S' \<equiv> {x. \<exists>u::real. u \<in> {0..1} \<and> (u, x::'a) \<in> {0..1} \<times> T - V}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4404 | have "closedin (subtopology euclidean T) S'" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4405 | unfolding S'_def | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4406 | apply (rule closedin_compact_projection, blast) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4407 | using closedin_self opeTV by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4408 |   have S'_def: "S' = {x. \<exists>u::real.  (u, x::'a) \<in> {0..1} \<times> T - V}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4409 | by (auto simp: S'_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4410 | have cloTS': "closedin (subtopology euclidean T) S'" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4411 | using S'_def \<open>closedin (subtopology euclidean T) S'\<close> by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4412 |   have "S \<inter> S' = {}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4413 | using S'_def B_def \<open>B \<subseteq> V\<close> by force | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4414 | obtain a :: "'a \<Rightarrow> real" where conta: "continuous_on T a" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4415 | and "\<And>x. x \<in> T \<Longrightarrow> a x \<in> closed_segment 1 0" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4416 | and a1: "\<And>x. x \<in> S \<Longrightarrow> a x = 1" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4417 | and a0: "\<And>x. x \<in> S' \<Longrightarrow> a x = 0" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4418 |     apply (rule Urysohn_local [OF cloTS cloTS' \<open>S \<inter> S' = {}\<close>, of 1 0], blast)
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4419 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4420 |   then have ain: "\<And>x. x \<in> T \<Longrightarrow> a x \<in> {0..1}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4421 | using closed_segment_eq_real_ivl by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4422 | have inV: "(u * a t, t) \<in> V" if "t \<in> T" "0 \<le> u" "u \<le> 1" for t u | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4423 | proof (rule ccontr) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4424 | assume "(u * a t, t) \<notin> V" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4425 | with ain [OF \<open>t \<in> T\<close>] have "a t = 0" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4426 | apply simp | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4427 | apply (rule a0) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4428 | by (metis (no_types, lifting) Diff_iff S'_def SigmaI atLeastAtMost_iff mem_Collect_eq mult_le_one mult_nonneg_nonneg that) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4429 | show False | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4430 | using B_def \<open>(u * a t, t) \<notin> V\<close> \<open>B \<subseteq> V\<close> \<open>a t = 0\<close> that by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4431 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4432 | show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4433 | proof | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4434 | show hom: "homotopic_with (\<lambda>x. True) T U f (\<lambda>x. k (a x, x))" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4435 | proof (simp add: homotopic_with, intro exI conjI) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4436 |       show "continuous_on ({0..1} \<times> T) (k \<circ> (\<lambda>z. (fst z *\<^sub>R (a \<circ> snd) z, snd z)))"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4437 | apply (intro continuous_on_compose continuous_intros) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4438 | apply (rule continuous_on_subset [OF conta], force) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4439 | apply (rule continuous_on_subset [OF contk]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4440 | apply (force intro: inV) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4441 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4442 |       show "(k \<circ> (\<lambda>z. (fst z *\<^sub>R (a \<circ> snd) z, snd z))) ` ({0..1} \<times> T) \<subseteq> U"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4443 | using inV kim by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4444 | show "\<forall>x\<in>T. (k \<circ> (\<lambda>z. (fst z *\<^sub>R (a \<circ> snd) z, snd z))) (0, x) = f x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4445 | by (simp add: B_def h'_def keq) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4446 | 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)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4447 | by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4448 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4449 | show "continuous_on T (\<lambda>x. k (a x, x))" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4450 | using hom homotopic_with_imp_continuous by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4451 | show "(\<lambda>x. k (a x, x)) ` T \<subseteq> U" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4452 | proof clarify | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4453 | fix t | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4454 | assume "t \<in> T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4455 | show "k (a t, t) \<in> U" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4456 | by (metis \<open>t \<in> T\<close> image_subset_iff inV kim not_one_le_zero linear mult_cancel_right1) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4457 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4458 | show "\<And>x. x \<in> S \<Longrightarrow> k (a x, x) = g x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4459 | by (simp add: B_def a1 h'_def keq) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4460 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4461 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4462 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4463 | |
| 68617 | 4464 | corollary%unimportant nullhomotopic_into_ANR_extension: | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4465 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4466 | assumes "closed S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4467 | and contf: "continuous_on S f" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4468 | and "ANR T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4469 | and fim: "f ` S \<subseteq> T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4470 |       and "S \<noteq> {}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4471 | shows "(\<exists>c. homotopic_with (\<lambda>x. True) S T f (\<lambda>x. c)) \<longleftrightarrow> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4472 | (\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> T \<and> (\<forall>x \<in> S. g x = f x))" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4473 | (is "?lhs = ?rhs") | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4474 | proof | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4475 | assume ?lhs | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4476 | then obtain c where c: "homotopic_with (\<lambda>x. True) S T (\<lambda>x. c) f" | 
| 68022 | 4477 | by (blast intro: homotopic_with_symD) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4478 | have "closedin (subtopology euclidean UNIV) S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4479 | using \<open>closed S\<close> closed_closedin by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4480 | then obtain g where "continuous_on UNIV g" "range g \<subseteq> T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4481 | "\<And>x. x \<in> S \<Longrightarrow> g x = f x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4482 | apply (rule Borsuk_homotopy_extension_homotopic [OF _ _ continuous_on_const _ c, where T=UNIV]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4483 |     using \<open>ANR T\<close> \<open>S \<noteq> {}\<close> c homotopic_with_imp_subset1 apply fastforce+
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4484 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4485 | then show ?rhs by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4486 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4487 | assume ?rhs | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4488 | then obtain g where "continuous_on UNIV g" "range g \<subseteq> T" "\<And>x. x\<in>S \<Longrightarrow> g x = f x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4489 | by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4490 | then obtain c where "homotopic_with (\<lambda>h. True) UNIV T g (\<lambda>x. c)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4491 | using nullhomotopic_from_contractible [of UNIV g T] contractible_UNIV by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4492 | then show ?lhs | 
| 68022 | 4493 | apply (rule_tac x=c in exI) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4494 | apply (rule homotopic_with_eq [of _ _ _ g "\<lambda>x. c"]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4495 | apply (rule homotopic_with_subset_left) | 
| 68022 | 4496 | apply (auto simp: \<open>\<And>x. x \<in> S \<Longrightarrow> g x = f x\<close>) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4497 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4498 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4499 | |
| 68617 | 4500 | corollary%unimportant nullhomotopic_into_rel_frontier_extension: | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4501 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4502 | assumes "closed S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4503 | and contf: "continuous_on S f" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4504 | and "convex T" "bounded T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4505 | and fim: "f ` S \<subseteq> rel_frontier T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4506 |       and "S \<noteq> {}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4507 | shows "(\<exists>c. homotopic_with (\<lambda>x. True) S (rel_frontier T) f (\<lambda>x. c)) \<longleftrightarrow> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4508 | (\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> rel_frontier T \<and> (\<forall>x \<in> S. g x = f x))" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4509 | by (simp add: nullhomotopic_into_ANR_extension assms ANR_rel_frontier_convex) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4510 | |
| 68617 | 4511 | corollary%unimportant nullhomotopic_into_sphere_extension: | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4512 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4513 | assumes "closed S" and contf: "continuous_on S f" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4514 |       and "S \<noteq> {}" and fim: "f ` S \<subseteq> sphere a r"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4515 | shows "((\<exists>c. homotopic_with (\<lambda>x. True) S (sphere a r) f (\<lambda>x. c)) \<longleftrightarrow> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4516 | (\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> sphere a r \<and> (\<forall>x \<in> S. g x = f x)))" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4517 | (is "?lhs = ?rhs") | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4518 | proof (cases "r = 0") | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4519 | case True with fim show ?thesis | 
| 68022 | 4520 | apply auto | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4521 | using fim continuous_on_const apply fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4522 | by (metis contf contractible_sing nullhomotopic_into_contractible) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4523 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4524 | case False | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4525 | then have eq: "sphere a r = rel_frontier (cball a r)" by simp | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4526 | show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4527 | using fim unfolding eq | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4528 | apply (rule nullhomotopic_into_rel_frontier_extension [OF \<open>closed S\<close> contf convex_cball bounded_cball]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4529 |     apply (rule \<open>S \<noteq> {}\<close>)
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4530 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4531 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4532 | |
| 68617 | 4533 | proposition%unimportant Borsuk_map_essential_bounded_component: | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4534 | fixes a :: "'a :: euclidean_space" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4535 | assumes "compact S" and "a \<notin> S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4536 | shows "bounded (connected_component_set (- S) a) \<longleftrightarrow> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4537 | ~(\<exists>c. homotopic_with (\<lambda>x. True) S (sphere 0 1) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4538 | (\<lambda>x. inverse(norm(x - a)) *\<^sub>R (x - a)) (\<lambda>x. c))" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4539 | (is "?lhs = ?rhs") | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4540 | proof (cases "S = {}")
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4541 | case True then show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4542 | by simp | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4543 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4544 | case False | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4545 | have "closed S" "bounded S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4546 | using \<open>compact S\<close> compact_eq_bounded_closed by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4547 | have s01: "(\<lambda>x. (x - a) /\<^sub>R norm (x - a)) ` S \<subseteq> sphere 0 1" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4548 | using \<open>a \<notin> S\<close> by clarsimp (metis dist_eq_0_iff dist_norm mult.commute right_inverse) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4549 | have aincc: "a \<in> connected_component_set (- S) a" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4550 | by (simp add: \<open>a \<notin> S\<close>) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4551 | obtain r where "r>0" and r: "S \<subseteq> ball 0 r" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4552 | using bounded_subset_ballD \<open>bounded S\<close> by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4553 | have "~ ?rhs \<longleftrightarrow> ~ ?lhs" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4554 | proof | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4555 | assume notr: "~ ?rhs" | 
| 63497 
ef794d2e3754
HOL-Multivariate_Analysis: add amssymb to document generation; reintroduce \nexists (cf d51a0a772094)
 hoelzl parents: 
63493diff
changeset | 4556 | have nog: "\<nexists>g. continuous_on (S \<union> connected_component_set (- S) a) g \<and> | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4557 | g ` (S \<union> connected_component_set (- S) a) \<subseteq> sphere 0 1 \<and> | 
| 63497 
ef794d2e3754
HOL-Multivariate_Analysis: add amssymb to document generation; reintroduce \nexists (cf d51a0a772094)
 hoelzl parents: 
63493diff
changeset | 4558 | (\<forall>x\<in>S. g x = (x - a) /\<^sub>R norm (x - a))" | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4559 | if "bounded (connected_component_set (- S) a)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4560 | apply (rule non_extensible_Borsuk_map [OF \<open>compact S\<close> componentsI _ aincc]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4561 | using \<open>a \<notin> S\<close> that by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4562 | obtain g where "range g \<subseteq> sphere 0 1" "continuous_on UNIV g" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4563 | "\<And>x. x \<in> S \<Longrightarrow> g x = (x - a) /\<^sub>R norm (x - a)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4564 | using notr | 
| 68022 | 4565 | by (auto simp: nullhomotopic_into_sphere_extension | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4566 | [OF \<open>closed S\<close> continuous_on_Borsuk_map [OF \<open>a \<notin> S\<close>] False s01]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4567 | with \<open>a \<notin> S\<close> show "~ ?lhs" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4568 | apply (clarsimp simp: Borsuk_map_into_sphere [of a S, symmetric] dest!: nog) | 
| 68022 | 4569 | apply (drule_tac x=g in spec) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4570 | using continuous_on_subset by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4571 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4572 | assume "~ ?lhs" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4573 | then obtain b where b: "b \<in> connected_component_set (- S) a" and "r \<le> norm b" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4574 | using bounded_iff linear by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4575 | then have bnot: "b \<notin> ball 0 r" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4576 | by simp | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4577 | have "homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. (x - a) /\<^sub>R norm (x - a)) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4578 | (\<lambda>x. (x - b) /\<^sub>R norm (x - b))" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4579 | apply (rule Borsuk_maps_homotopic_in_path_component) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4580 | using \<open>closed S\<close> b open_Compl open_path_connected_component apply fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4581 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4582 | moreover | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4583 | obtain c where "homotopic_with (\<lambda>x. True) (ball 0 r) (sphere 0 1) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4584 | (\<lambda>x. inverse (norm (x - b)) *\<^sub>R (x - b)) (\<lambda>x. c)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4585 | proof (rule nullhomotopic_from_contractible) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4586 | show "contractible (ball (0::'a) r)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4587 | by (metis convex_imp_contractible convex_ball) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4588 | show "continuous_on (ball 0 r) (\<lambda>x. inverse(norm (x - b)) *\<^sub>R (x - b))" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4589 | by (rule continuous_on_Borsuk_map [OF bnot]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4590 | show "(\<lambda>x. (x - b) /\<^sub>R norm (x - b)) ` ball 0 r \<subseteq> sphere 0 1" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4591 | using bnot Borsuk_map_into_sphere by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4592 | qed blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4593 | ultimately have "homotopic_with (\<lambda>x. True) S (sphere 0 1) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4594 | (\<lambda>x. (x - a) /\<^sub>R norm (x - a)) (\<lambda>x. c)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4595 | by (meson homotopic_with_subset_left homotopic_with_trans r) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4596 | then show "~ ?rhs" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4597 | by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4598 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4599 | then show ?thesis by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4600 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4601 | |
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4602 | lemma homotopic_Borsuk_maps_in_bounded_component: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4603 | fixes a :: "'a :: euclidean_space" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4604 | assumes "compact S" and "a \<notin> S"and "b \<notin> S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4605 | and boc: "bounded (connected_component_set (- S) a)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4606 | and hom: "homotopic_with (\<lambda>x. True) S (sphere 0 1) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4607 | (\<lambda>x. (x - a) /\<^sub>R norm (x - a)) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4608 | (\<lambda>x. (x - b) /\<^sub>R norm (x - b))" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4609 | shows "connected_component (- S) a b" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4610 | proof (rule ccontr) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4611 | assume notcc: "\<not> connected_component (- S) a b" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4612 | let ?T = "S \<union> connected_component_set (- S) a" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4613 | have "\<nexists>g. continuous_on (S \<union> connected_component_set (- S) a) g \<and> | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4614 | g ` (S \<union> connected_component_set (- S) a) \<subseteq> sphere 0 1 \<and> | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4615 | (\<forall>x\<in>S. g x = (x - a) /\<^sub>R norm (x - a))" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4616 | by (simp add: \<open>a \<notin> S\<close> componentsI non_extensible_Borsuk_map [OF \<open>compact S\<close> _ boc]) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4617 | moreover obtain g where "continuous_on (S \<union> connected_component_set (- S) a) g" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4618 | "g ` (S \<union> connected_component_set (- S) a) \<subseteq> sphere 0 1" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4619 | "\<And>x. x \<in> S \<Longrightarrow> g x = (x - a) /\<^sub>R norm (x - a)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4620 | proof (rule Borsuk_homotopy_extension_homotopic) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4621 | show "closedin (subtopology euclidean ?T) S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4622 | by (simp add: \<open>compact S\<close> closed_subset compact_imp_closed) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4623 | show "continuous_on ?T (\<lambda>x. (x - b) /\<^sub>R norm (x - b))" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4624 | by (simp add: \<open>b \<notin> S\<close> notcc continuous_on_Borsuk_map) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4625 | show "(\<lambda>x. (x - b) /\<^sub>R norm (x - b)) ` ?T \<subseteq> sphere 0 1" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4626 | by (simp add: \<open>b \<notin> S\<close> notcc Borsuk_map_into_sphere) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4627 | show "homotopic_with (\<lambda>x. True) S (sphere 0 1) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4628 | (\<lambda>x. (x - b) /\<^sub>R norm (x - b)) (\<lambda>x. (x - a) /\<^sub>R norm (x - a))" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4629 | by (simp add: hom homotopic_with_symD) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4630 | qed (auto simp: ANR_sphere intro: that) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4631 | ultimately show False by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4632 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4633 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4634 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4635 | lemma Borsuk_maps_homotopic_in_connected_component_eq: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4636 | fixes a :: "'a :: euclidean_space" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4637 |   assumes S: "compact S" "a \<notin> S" "b \<notin> S" and 2: "2 \<le> DIM('a)"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4638 | shows "(homotopic_with (\<lambda>x. True) S (sphere 0 1) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4639 | (\<lambda>x. (x - a) /\<^sub>R norm (x - a)) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4640 | (\<lambda>x. (x - b) /\<^sub>R norm (x - b)) \<longleftrightarrow> | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4641 | connected_component (- S) a b)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4642 | (is "?lhs = ?rhs") | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4643 | proof | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4644 | assume L: ?lhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4645 | show ?rhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4646 | proof (cases "bounded(connected_component_set (- S) a)") | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4647 | case True | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4648 | show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4649 | by (rule homotopic_Borsuk_maps_in_bounded_component [OF S True L]) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4650 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4651 | case not_bo_a: False | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4652 | show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4653 | proof (cases "bounded(connected_component_set (- S) b)") | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4654 | case True | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4655 | show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4656 | using homotopic_Borsuk_maps_in_bounded_component [OF S] | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4657 | by (simp add: L True assms connected_component_sym homotopic_Borsuk_maps_in_bounded_component homotopic_with_sym) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4658 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4659 | case False | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4660 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4661 | using cobounded_unique_unbounded_component [of "-S" a b] \<open>compact S\<close> not_bo_a | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4662 | by (auto simp: compact_eq_bounded_closed assms connected_component_eq_eq) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4663 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4664 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4665 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4666 | assume R: ?rhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4667 | then have "path_component (- S) a b" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4668 | using assms(1) compact_eq_bounded_closed open_Compl open_path_connected_component_set by fastforce | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4669 | then show ?lhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4670 | by (simp add: Borsuk_maps_homotopic_in_path_component) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4671 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4672 | |
| 68617 | 4673 | subsubsection\<open>More extension theorems\<close> | 
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4674 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4675 | lemma extension_from_clopen: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4676 | assumes ope: "openin (subtopology euclidean S) T" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4677 | and clo: "closedin (subtopology euclidean S) T" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4678 |       and contf: "continuous_on T f" and fim: "f ` T \<subseteq> U" and null: "U = {} \<Longrightarrow> S = {}"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4679 | obtains g where "continuous_on S g" "g ` S \<subseteq> U" "\<And>x. x \<in> T \<Longrightarrow> g x = f x" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4680 | proof (cases "U = {}")
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4681 | case True | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4682 | then show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4683 | by (simp add: null that) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4684 | next | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4685 | case False | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4686 | then obtain a where "a \<in> U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4687 | by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4688 | let ?g = "\<lambda>x. if x \<in> T then f x else a" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4689 | have Seq: "S = T \<union> (S - T)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4690 | using clo closedin_imp_subset by fastforce | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4691 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4692 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4693 | have "continuous_on (T \<union> (S - T)) ?g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4694 | apply (rule continuous_on_cases_local) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4695 | using Seq clo ope by (auto simp: contf continuous_on_const intro: continuous_on_cases_local) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4696 | with Seq show "continuous_on S ?g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4697 | by metis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4698 | show "?g ` S \<subseteq> U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4699 | using \<open>a \<in> U\<close> fim by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4700 | show "\<And>x. x \<in> T \<Longrightarrow> ?g x = f x" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4701 | by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4702 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4703 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4704 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4705 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4706 | lemma extension_from_component: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4707 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4708 | assumes S: "locally connected S \<or> compact S" and "ANR U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4709 | and C: "C \<in> components S" and contf: "continuous_on C f" and fim: "f ` C \<subseteq> U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4710 | obtains g where "continuous_on S g" "g ` S \<subseteq> U" "\<And>x. x \<in> C \<Longrightarrow> g x = f x" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4711 | proof - | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4712 | obtain T g where ope: "openin (subtopology euclidean S) T" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4713 | and clo: "closedin (subtopology euclidean S) T" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4714 | and "C \<subseteq> T" and contg: "continuous_on T g" and gim: "g ` T \<subseteq> U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4715 | and gf: "\<And>x. x \<in> C \<Longrightarrow> g x = f x" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4716 | using S | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4717 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4718 | assume "locally connected S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4719 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4720 | by (metis C \<open>locally connected S\<close> openin_components_locally_connected closedin_component contf fim order_refl that) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4721 | next | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4722 | assume "compact S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4723 | then obtain W g where "C \<subseteq> W" and opeW: "openin (subtopology euclidean S) W" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4724 | and contg: "continuous_on W g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4725 | and gim: "g ` W \<subseteq> U" and gf: "\<And>x. x \<in> C \<Longrightarrow> g x = f x" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4726 | using ANR_imp_absolute_neighbourhood_extensor [of U C f S] C \<open>ANR U\<close> closedin_component contf fim by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4727 | then obtain V where "open V" and V: "W = S \<inter> V" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4728 | by (auto simp: openin_open) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4729 | moreover have "locally compact S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4730 | by (simp add: \<open>compact S\<close> closed_imp_locally_compact compact_imp_closed) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4731 | ultimately obtain K where opeK: "openin (subtopology euclidean S) K" and "compact K" "C \<subseteq> K" "K \<subseteq> V" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4732 | by (metis C Int_subset_iff \<open>C \<subseteq> W\<close> \<open>compact S\<close> compact_components Sura_Bura_clopen_subset) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4733 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4734 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4735 | show "closedin (subtopology euclidean S) K" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4736 | by (meson \<open>compact K\<close> \<open>compact S\<close> closedin_compact_eq opeK openin_imp_subset) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4737 | show "continuous_on K g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4738 | by (metis Int_subset_iff V \<open>K \<subseteq> V\<close> contg continuous_on_subset opeK openin_subtopology subset_eq) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4739 | show "g ` K \<subseteq> U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4740 | using V \<open>K \<subseteq> V\<close> gim opeK openin_imp_subset by fastforce | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4741 | qed (use opeK gf \<open>C \<subseteq> K\<close> in auto) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4742 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4743 | obtain h where "continuous_on S h" "h ` S \<subseteq> U" "\<And>x. x \<in> T \<Longrightarrow> h x = g x" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4744 | using extension_from_clopen | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4745 | by (metis C bot.extremum_uniqueI clo contg gim fim image_is_empty in_components_nonempty ope) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4746 | then show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4747 | by (metis \<open>C \<subseteq> T\<close> gf subset_eq that) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4748 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4749 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4750 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4751 | lemma tube_lemma: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4752 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4753 |   assumes "compact S" and S: "S \<noteq> {}" "(\<lambda>x. (x,a)) ` S \<subseteq> U" 
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4754 | and ope: "openin (subtopology euclidean (S \<times> T)) U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4755 | obtains V where "openin (subtopology euclidean T) V" "a \<in> V" "S \<times> V \<subseteq> U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4756 | proof - | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4757 |   let ?W = "{y. \<exists>x. x \<in> S \<and> (x, y) \<in> (S \<times> T - U)}"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4758 | have "U \<subseteq> S \<times> T" "closedin (subtopology euclidean (S \<times> T)) (S \<times> T - U)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4759 | using ope by (auto simp: openin_closedin_eq) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4760 | then have "closedin (subtopology euclidean T) ?W" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4761 | using \<open>compact S\<close> closedin_compact_projection by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4762 | moreover have "a \<in> T - ?W" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4763 | using \<open>U \<subseteq> S \<times> T\<close> S by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4764 | moreover have "S \<times> (T - ?W) \<subseteq> U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4765 | by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4766 | ultimately show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4767 | by (metis (no_types, lifting) Sigma_cong closedin_def that topspace_euclidean_subtopology) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4768 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4769 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4770 | lemma tube_lemma_gen: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4771 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4772 |   assumes "compact S" "S \<noteq> {}" "T \<subseteq> T'" "S \<times> T \<subseteq> U"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4773 | and ope: "openin (subtopology euclidean (S \<times> T')) U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4774 | obtains V where "openin (subtopology euclidean T') V" "T \<subseteq> V" "S \<times> V \<subseteq> U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4775 | proof - | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4776 | have "\<And>x. x \<in> T \<Longrightarrow> \<exists>V. openin (subtopology euclidean T') V \<and> x \<in> V \<and> S \<times> V \<subseteq> U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4777 | using assms by (auto intro: tube_lemma [OF \<open>compact S\<close>]) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4778 | then obtain F where F: "\<And>x. x \<in> T \<Longrightarrow> openin (subtopology euclidean T') (F x) \<and> x \<in> F x \<and> S \<times> F x \<subseteq> U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4779 | by metis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4780 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4781 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4782 | show "openin (subtopology euclidean T') (UNION T F)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4783 | using F by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4784 | show "T \<subseteq> UNION T F" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4785 | using F by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4786 | show "S \<times> UNION T F \<subseteq> U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4787 | using F by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4788 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4789 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4790 | |
| 68617 | 4791 | proposition%unimportant homotopic_neighbourhood_extension: | 
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4792 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4793 | assumes contf: "continuous_on S f" and fim: "f ` S \<subseteq> U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4794 | and contg: "continuous_on S g" and gim: "g ` S \<subseteq> U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4795 | and clo: "closedin (subtopology euclidean S) T" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4796 | and "ANR U" and hom: "homotopic_with (\<lambda>x. True) T U f g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4797 | obtains V where "T \<subseteq> V" "openin (subtopology euclidean S) V" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4798 | "homotopic_with (\<lambda>x. True) V U f g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4799 | proof - | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4800 | have "T \<subseteq> S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4801 | using clo closedin_imp_subset by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4802 |   obtain h where conth: "continuous_on ({0..1::real} \<times> T) h"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4803 |              and him: "h ` ({0..1} \<times> T) \<subseteq> U"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4804 | and h0: "\<And>x. h(0, x) = f x" and h1: "\<And>x. h(1, x) = g x" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4805 | using hom by (auto simp: homotopic_with_def) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4806 |   define h' where "h' \<equiv> \<lambda>z. if fst z \<in> {0} then f(snd z)
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4807 |                              else if fst z \<in> {1} then g(snd z)
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4808 | else h z" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4809 |   let ?S0 = "{0::real} \<times> S" and ?S1 = "{1::real} \<times> S"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4810 |   have "continuous_on(?S0 \<union> (?S1 \<union> {0..1} \<times> T)) h'"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4811 | unfolding h'_def | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4812 | proof (intro continuous_on_cases_local) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4813 |     show "closedin (subtopology euclidean (?S0 \<union> (?S1 \<union> {0..1} \<times> T))) ?S0"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4814 |          "closedin (subtopology euclidean (?S1 \<union> {0..1} \<times> T)) ?S1"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4815 |       using \<open>T \<subseteq> S\<close> by (force intro: closedin_Times closedin_subset_trans [of "{0..1} \<times> S"])+
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4816 |     show "closedin (subtopology euclidean (?S0 \<union> (?S1 \<union> {0..1} \<times> T))) (?S1 \<union> {0..1} \<times> T)"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4817 |          "closedin (subtopology euclidean (?S1 \<union> {0..1} \<times> T)) ({0..1} \<times> T)"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4818 |       using \<open>T \<subseteq> S\<close> by (force intro: clo closedin_Times closedin_subset_trans [of "{0..1} \<times> S"])+
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4819 | show "continuous_on (?S0) (\<lambda>x. f (snd x))" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4820 | by (intro continuous_intros continuous_on_compose2 [OF contf]) auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4821 | show "continuous_on (?S1) (\<lambda>x. g (snd x))" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4822 | by (intro continuous_intros continuous_on_compose2 [OF contg]) auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4823 | qed (use h0 h1 conth in auto) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4824 |   then have "continuous_on ({0,1} \<times> S \<union> ({0..1} \<times> T)) h'"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4825 | by (metis Sigma_Un_distrib1 Un_assoc insert_is_Un) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4826 |   moreover have "h' ` ({0,1} \<times> S \<union> {0..1} \<times> T) \<subseteq> U"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4827 | using fim gim him \<open>T \<subseteq> S\<close> unfolding h'_def by force | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4828 |   moreover have "closedin (subtopology euclidean ({0..1::real} \<times> S)) ({0,1} \<times> S \<union> {0..1::real} \<times> T)"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4829 | by (intro closedin_Times closedin_Un clo) (simp_all add: closed_subset) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4830 | ultimately | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4831 |   obtain W k where W: "({0,1} \<times> S) \<union> ({0..1} \<times> T) \<subseteq> W"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4832 |                and opeW: "openin (subtopology euclidean ({0..1} \<times> S)) W"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4833 | and contk: "continuous_on W k" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4834 | and kim: "k ` W \<subseteq> U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4835 |                and kh': "\<And>x. x \<in> ({0,1} \<times> S) \<union> ({0..1} \<times> T) \<Longrightarrow> k x = h' x"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4836 |     by (metis ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR U\<close>, of "({0,1} \<times> S) \<union> ({0..1} \<times> T)" h' "{0..1} \<times> S"])
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4837 | obtain T' where opeT': "openin (subtopology euclidean S) T'" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4838 |               and "T \<subseteq> T'" and TW: "{0..1} \<times> T' \<subseteq> W"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4839 |     using tube_lemma_gen [of "{0..1::real}" T S W] W \<open>T \<subseteq> S\<close> opeW by auto
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4840 | moreover have "homotopic_with (\<lambda>x. True) T' U f g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4841 | proof (simp add: homotopic_with, intro exI conjI) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4842 |     show "continuous_on ({0..1} \<times> T') k"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4843 | using TW continuous_on_subset contk by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4844 |     show "k ` ({0..1} \<times> T') \<subseteq> U"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4845 | using TW kim by fastforce | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4846 | have "T' \<subseteq> S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4847 | by (meson opeT' subsetD openin_imp_subset) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4848 | then show "\<forall>x\<in>T'. k (0, x) = f x" "\<forall>x\<in>T'. k (1, x) = g x" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4849 | by (auto simp: kh' h'_def) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4850 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4851 | ultimately show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4852 | by (blast intro: that) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4853 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4854 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4855 | text\<open> Homotopy on a union of closed-open sets.\<close> | 
| 68617 | 4856 | proposition%unimportant homotopic_on_clopen_Union: | 
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4857 | fixes \<F> :: "'a::euclidean_space set set" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4858 | assumes "\<And>S. S \<in> \<F> \<Longrightarrow> closedin (subtopology euclidean (\<Union>\<F>)) S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4859 | and "\<And>S. S \<in> \<F> \<Longrightarrow> openin (subtopology euclidean (\<Union>\<F>)) S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4860 | and "\<And>S. S \<in> \<F> \<Longrightarrow> homotopic_with (\<lambda>x. True) S T f g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4861 | shows "homotopic_with (\<lambda>x. True) (\<Union>\<F>) T f g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4862 | proof - | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4863 | obtain \<V> where "\<V> \<subseteq> \<F>" "countable \<V>" and eqU: "\<Union>\<V> = \<Union>\<F>" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4864 | using Lindelof_openin assms by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4865 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4866 |   proof (cases "\<V> = {}")
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4867 | case True | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4868 | then show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4869 | by (metis Union_empty eqU homotopic_on_empty) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4870 | next | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4871 | case False | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4872 | then obtain V :: "nat \<Rightarrow> 'a set" where V: "range V = \<V>" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4873 | using range_from_nat_into \<open>countable \<V>\<close> by metis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4874 | with \<open>\<V> \<subseteq> \<F>\<close> have clo: "\<And>n. closedin (subtopology euclidean (\<Union>\<F>)) (V n)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4875 | and ope: "\<And>n. openin (subtopology euclidean (\<Union>\<F>)) (V n)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4876 | and hom: "\<And>n. homotopic_with (\<lambda>x. True) (V n) T f g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4877 | using assms by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4878 |     then obtain h where conth: "\<And>n. continuous_on ({0..1::real} \<times> V n) (h n)"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4879 |                   and him: "\<And>n. h n ` ({0..1} \<times> V n) \<subseteq> T" 
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4880 | and h0: "\<And>n. \<And>x. x \<in> V n \<Longrightarrow> h n (0, x) = f x" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4881 | and h1: "\<And>n. \<And>x. x \<in> V n \<Longrightarrow> h n (1, x) = g x" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4882 | by (simp add: homotopic_with) metis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4883 | have wop: "b \<in> V x \<Longrightarrow> \<exists>k. b \<in> V k \<and> (\<forall>j<k. b \<notin> V j)" for b x | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4884 | using nat_less_induct [where P = "\<lambda>i. b \<notin> V i"] by meson | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4885 |     obtain \<zeta> where cont: "continuous_on ({0..1} \<times> UNION UNIV V) \<zeta>"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4886 |               and eq: "\<And>x i. \<lbrakk>x \<in> {0..1} \<times> UNION UNIV V \<inter>
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4887 |                                    {0..1} \<times> (V i - (\<Union>m<i. V m))\<rbrakk> \<Longrightarrow> \<zeta> x = h i x"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4888 | proof (rule pasting_lemma_exists) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4889 |       show "{0..1} \<times> UNION UNIV V \<subseteq> (\<Union>i. {0..1::real} \<times> (V i - (\<Union>m<i. V m)))"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4890 | by (force simp: Ball_def dest: wop) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4891 |       show "openin (subtopology euclidean ({0..1} \<times> UNION UNIV V)) 
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4892 |                    ({0..1::real} \<times> (V i - (\<Union>m<i. V m)))" for i
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4893 | proof (intro openin_Times openin_subtopology_self openin_diff) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4894 | show "openin (subtopology euclidean (UNION UNIV V)) (V i)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4895 | using ope V eqU by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4896 | show "closedin (subtopology euclidean (UNION UNIV V)) (\<Union>m<i. V m)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4897 | using V clo eqU by (force intro: closedin_Union) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4898 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4899 |       show "continuous_on ({0..1} \<times> (V i - (\<Union>m<i. V m))) (h i)" for i
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4900 | by (rule continuous_on_subset [OF conth]) auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4901 |       show "\<And>i j x. x \<in> {0..1} \<times> UNION UNIV V \<inter>
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4902 |                     {0..1} \<times> (V i - (\<Union>m<i. V m)) \<inter> {0..1} \<times> (V j - (\<Union>m<j. V m))
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4903 | \<Longrightarrow> h i x = h j x" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4904 | by clarsimp (metis lessThan_iff linorder_neqE_nat) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4905 | qed auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4906 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4907 | proof (simp add: homotopic_with eqU [symmetric], intro exI conjI ballI) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4908 |       show "continuous_on ({0..1} \<times> \<Union>\<V>) \<zeta>"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4909 | using V eqU by (blast intro!: continuous_on_subset [OF cont]) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4910 |       show "\<zeta>` ({0..1} \<times> \<Union>\<V>) \<subseteq> T"
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4911 | proof clarsimp | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4912 | fix t :: real and y :: "'a" and X :: "'a set" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4913 | assume "y \<in> X" "X \<in> \<V>" and t: "0 \<le> t" "t \<le> 1" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4914 | then obtain k where "y \<in> V k" and j: "\<forall>j<k. y \<notin> V j" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4915 | by (metis image_iff V wop) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4916 | with him t show "\<zeta>(t, y) \<in> T" | 
| 68022 | 4917 | by (subst eq) force+ | 
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4918 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4919 | fix X y | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4920 | assume "X \<in> \<V>" "y \<in> X" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4921 | then obtain k where "y \<in> V k" and j: "\<forall>j<k. y \<notin> V j" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4922 | by (metis image_iff V wop) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4923 | then show "\<zeta>(0, y) = f y" and "\<zeta>(1, y) = g y" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4924 | by (subst eq [where i=k]; force simp: h0 h1)+ | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4925 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4926 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4927 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4928 | |
| 68617 | 4929 | lemma homotopic_on_components_eq: | 
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4930 | fixes S :: "'a :: euclidean_space set" and T :: "'b :: euclidean_space set" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4931 | assumes S: "locally connected S \<or> compact S" and "ANR T" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4932 | shows "homotopic_with (\<lambda>x. True) S T f g \<longleftrightarrow> | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4933 | (continuous_on S f \<and> f ` S \<subseteq> T \<and> continuous_on S g \<and> g ` S \<subseteq> T) \<and> | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4934 | (\<forall>C \<in> components S. homotopic_with (\<lambda>x. True) C T f g)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4935 | (is "?lhs \<longleftrightarrow> ?C \<and> ?rhs") | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4936 | proof - | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4937 | have "continuous_on S f" "f ` S \<subseteq> T" "continuous_on S g" "g ` S \<subseteq> T" if ?lhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4938 | using homotopic_with_imp_continuous homotopic_with_imp_subset1 homotopic_with_imp_subset2 that by blast+ | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4939 | moreover have "?lhs \<longleftrightarrow> ?rhs" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4940 | if contf: "continuous_on S f" and fim: "f ` S \<subseteq> T" and contg: "continuous_on S g" and gim: "g ` S \<subseteq> T" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4941 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4942 | assume ?lhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4943 | with that show ?rhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4944 | by (simp add: homotopic_with_subset_left in_components_subset) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4945 | next | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4946 | assume R: ?rhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4947 | have "\<exists>U. C \<subseteq> U \<and> closedin (subtopology euclidean S) U \<and> | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4948 | openin (subtopology euclidean S) U \<and> | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4949 | homotopic_with (\<lambda>x. True) U T f g" if C: "C \<in> components S" for C | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4950 | proof - | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4951 | have "C \<subseteq> S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4952 | by (simp add: in_components_subset that) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4953 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4954 | using S | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4955 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4956 | assume "locally connected S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4957 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4958 | proof (intro exI conjI) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4959 | show "closedin (subtopology euclidean S) C" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4960 | by (simp add: closedin_component that) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4961 | show "openin (subtopology euclidean S) C" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4962 | by (simp add: \<open>locally connected S\<close> openin_components_locally_connected that) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4963 | show "homotopic_with (\<lambda>x. True) C T f g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4964 | by (simp add: R that) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4965 | qed auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4966 | next | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4967 | assume "compact S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4968 | have hom: "homotopic_with (\<lambda>x. True) C T f g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4969 | using R that by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4970 | obtain U where "C \<subseteq> U" and opeU: "openin (subtopology euclidean S) U" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4971 | and hom: "homotopic_with (\<lambda>x. True) U T f g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4972 | using homotopic_neighbourhood_extension [OF contf fim contg gim _ \<open>ANR T\<close> hom] | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4973 | \<open>C \<in> components S\<close> closedin_component by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4974 | then obtain V where "open V" and V: "U = S \<inter> V" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4975 | by (auto simp: openin_open) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4976 | moreover have "locally compact S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4977 | by (simp add: \<open>compact S\<close> closed_imp_locally_compact compact_imp_closed) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4978 | ultimately obtain K where opeK: "openin (subtopology euclidean S) K" and "compact K" "C \<subseteq> K" "K \<subseteq> V" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4979 | by (metis C Int_subset_iff Sura_Bura_clopen_subset \<open>C \<subseteq> U\<close> \<open>compact S\<close> compact_components) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4980 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4981 | proof (intro exI conjI) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4982 | show "closedin (subtopology euclidean S) K" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4983 | by (meson \<open>compact K\<close> \<open>compact S\<close> closedin_compact_eq opeK openin_imp_subset) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4984 | show "homotopic_with (\<lambda>x. True) K T f g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4985 | using V \<open>K \<subseteq> V\<close> hom homotopic_with_subset_left opeK openin_imp_subset by fastforce | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4986 | qed (use opeK \<open>C \<subseteq> K\<close> in auto) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4987 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4988 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4989 | then obtain \<phi> where \<phi>: "\<And>C. C \<in> components S \<Longrightarrow> C \<subseteq> \<phi> C" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4990 | and clo\<phi>: "\<And>C. C \<in> components S \<Longrightarrow> closedin (subtopology euclidean S) (\<phi> C)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4991 | and ope\<phi>: "\<And>C. C \<in> components S \<Longrightarrow> openin (subtopology euclidean S) (\<phi> C)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4992 | and hom\<phi>: "\<And>C. C \<in> components S \<Longrightarrow> homotopic_with (\<lambda>x. True) (\<phi> C) T f g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4993 | by metis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4994 | have Seq: "S = UNION (components S) \<phi>" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4995 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4996 | show "S \<subseteq> UNION (components S) \<phi>" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4997 | by (metis Sup_mono Union_components \<phi> imageI) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4998 | show "UNION (components S) \<phi> \<subseteq> S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4999 | using ope\<phi> openin_imp_subset by fastforce | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5000 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5001 | show ?lhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5002 | apply (subst Seq) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5003 | apply (rule homotopic_on_clopen_Union) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5004 | using Seq clo\<phi> ope\<phi> hom\<phi> by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5005 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5006 | ultimately show ?thesis by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5007 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5008 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5009 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5010 | lemma cohomotopically_trivial_on_components: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5011 | fixes S :: "'a :: euclidean_space set" and T :: "'b :: euclidean_space set" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5012 | assumes S: "locally connected S \<or> compact S" and "ANR T" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5013 | shows | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5014 | "(\<forall>f g. continuous_on S f \<longrightarrow> f ` S \<subseteq> T \<longrightarrow> continuous_on S g \<longrightarrow> g ` S \<subseteq> T \<longrightarrow> | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5015 | homotopic_with (\<lambda>x. True) S T f g) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5016 | \<longleftrightarrow> | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5017 | (\<forall>C\<in>components S. | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5018 | \<forall>f g. continuous_on C f \<longrightarrow> f ` C \<subseteq> T \<longrightarrow> continuous_on C g \<longrightarrow> g ` C \<subseteq> T \<longrightarrow> | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5019 | homotopic_with (\<lambda>x. True) C T f g)" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5020 | (is "?lhs = ?rhs") | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5021 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5022 | assume L[rule_format]: ?lhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5023 | show ?rhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5024 | proof clarify | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5025 | fix C f g | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5026 | assume contf: "continuous_on C f" and fim: "f ` C \<subseteq> T" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5027 | and contg: "continuous_on C g" and gim: "g ` C \<subseteq> T" and C: "C \<in> components S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5028 | obtain f' where contf': "continuous_on S f'" and f'im: "f' ` S \<subseteq> T" and f'f: "\<And>x. x \<in> C \<Longrightarrow> f' x = f x" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5029 | using extension_from_component [OF S \<open>ANR T\<close> C contf fim] by metis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5030 | obtain g' where contg': "continuous_on S g'" and g'im: "g' ` S \<subseteq> T" and g'g: "\<And>x. x \<in> C \<Longrightarrow> g' x = g x" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5031 | using extension_from_component [OF S \<open>ANR T\<close> C contg gim] by metis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5032 | have "homotopic_with (\<lambda>x. True) C T f' g'" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5033 | using L [OF contf' f'im contg' g'im] homotopic_with_subset_left C in_components_subset by fastforce | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5034 | then show "homotopic_with (\<lambda>x. True) C T f g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5035 | using f'f g'g homotopic_with_eq by force | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5036 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5037 | next | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5038 | assume R [rule_format]: ?rhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5039 | show ?lhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5040 | proof clarify | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5041 | fix f g | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5042 | assume contf: "continuous_on S f" and fim: "f ` S \<subseteq> T" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5043 | and contg: "continuous_on S g" and gim: "g ` S \<subseteq> T" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5044 | moreover have "homotopic_with (\<lambda>x. True) C T f g" if "C \<in> components S" for C | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5045 | using R [OF that] | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5046 | by (meson contf contg continuous_on_subset fim gim image_mono in_components_subset order.trans that) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5047 | ultimately show "homotopic_with (\<lambda>x. True) S T f g" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5048 | by (subst homotopic_on_components_eq [OF S \<open>ANR T\<close>]) auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5049 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5050 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5051 | |
| 68617 | 5052 | subsubsection\<open>The complement of a set and path-connectedness\<close> | 
| 64122 | 5053 | |
| 5054 | text\<open>Complement in dimension N > 1 of set homeomorphic to any interval in | |
| 5055 | any dimension is (path-)connected. This naively generalizes the argument | |
| 5056 | in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer fixed point theorem", | |
| 5057 | American Mathematical Monthly 1984.\<close> | |
| 5058 | ||
| 5059 | lemma unbounded_components_complement_absolute_retract: | |
| 5060 | fixes S :: "'a::euclidean_space set" | |
| 5061 | assumes C: "C \<in> components(- S)" and S: "compact S" "AR S" | |
| 5062 | shows "\<not> bounded C" | |
| 5063 | proof - | |
| 5064 | obtain y where y: "C = connected_component_set (- S) y" and "y \<notin> S" | |
| 5065 | using C by (auto simp: components_def) | |
| 5066 | have "open(- S)" | |
| 5067 | using S by (simp add: closed_open compact_eq_bounded_closed) | |
| 5068 | have "S retract_of UNIV" | |
| 5069 | using S compact_AR by blast | |
| 5070 | then obtain r where contr: "continuous_on UNIV r" and ontor: "range r \<subseteq> S" | |
| 5071 | and r: "\<And>x. x \<in> S \<Longrightarrow> r x = x" | |
| 5072 | by (auto simp: retract_of_def retraction_def) | |
| 5073 | show ?thesis | |
| 5074 | proof | |
| 5075 | assume "bounded C" | |
| 5076 | have "connected_component_set (- S) y \<subseteq> S" | |
| 5077 | proof (rule frontier_subset_retraction) | |
| 5078 | show "bounded (connected_component_set (- S) y)" | |
| 5079 | using \<open>bounded C\<close> y by blast | |
| 5080 | show "frontier (connected_component_set (- S) y) \<subseteq> S" | |
| 5081 | using C \<open>compact S\<close> compact_eq_bounded_closed frontier_of_components_closed_complement y by blast | |
| 5082 | show "continuous_on (closure (connected_component_set (- S) y)) r" | |
| 5083 | by (blast intro: continuous_on_subset [OF contr]) | |
| 5084 | qed (use ontor r in auto) | |
| 5085 | with \<open>y \<notin> S\<close> show False by force | |
| 5086 | qed | |
| 5087 | qed | |
| 5088 | ||
| 5089 | lemma connected_complement_absolute_retract: | |
| 5090 | fixes S :: "'a::euclidean_space set" | |
| 5091 |   assumes S: "compact S" "AR S" and 2: "2 \<le> DIM('a)"
 | |
| 5092 | shows "connected(- S)" | |
| 5093 | proof - | |
| 5094 | have "S retract_of UNIV" | |
| 5095 | using S compact_AR by blast | |
| 5096 | show ?thesis | |
| 5097 | apply (clarsimp simp: connected_iff_connected_component_eq) | |
| 5098 | apply (rule cobounded_unique_unbounded_component [OF _ 2]) | |
| 5099 | apply (simp add: \<open>compact S\<close> compact_imp_bounded) | |
| 5100 | apply (meson ComplI S componentsI unbounded_components_complement_absolute_retract)+ | |
| 5101 | done | |
| 5102 | qed | |
| 5103 | ||
| 5104 | lemma path_connected_complement_absolute_retract: | |
| 5105 | fixes S :: "'a::euclidean_space set" | |
| 5106 |   assumes "compact S" "AR S" "2 \<le> DIM('a)"
 | |
| 5107 | shows "path_connected(- S)" | |
| 5108 | using connected_complement_absolute_retract [OF assms] | |
| 5109 | using \<open>compact S\<close> compact_eq_bounded_closed connected_open_path_connected by blast | |
| 5110 | ||
| 5111 | theorem connected_complement_homeomorphic_convex_compact: | |
| 5112 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | |
| 5113 |   assumes hom: "S homeomorphic T" and T: "convex T" "compact T" and 2: "2 \<le> DIM('a)"
 | |
| 5114 | shows "connected(- S)" | |
| 5115 | proof (cases "S = {}")
 | |
| 5116 | case True | |
| 5117 | then show ?thesis | |
| 5118 | by (simp add: connected_UNIV) | |
| 5119 | next | |
| 5120 | case False | |
| 5121 | show ?thesis | |
| 5122 | proof (rule connected_complement_absolute_retract) | |
| 5123 | show "compact S" | |
| 5124 | using \<open>compact T\<close> hom homeomorphic_compactness by auto | |
| 5125 | show "AR S" | |
| 5126 | by (meson AR_ANR False \<open>convex T\<close> convex_imp_ANR convex_imp_contractible hom homeomorphic_ANR_iff_ANR homeomorphic_contractible_eq) | |
| 5127 | qed (rule 2) | |
| 5128 | qed | |
| 5129 | ||
| 5130 | corollary path_connected_complement_homeomorphic_convex_compact: | |
| 5131 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | |
| 5132 |   assumes hom: "S homeomorphic T" "convex T" "compact T" "2 \<le> DIM('a)"
 | |
| 5133 | shows "path_connected(- S)" | |
| 5134 | using connected_complement_homeomorphic_convex_compact [OF assms] | |
| 5135 | using \<open>compact T\<close> compact_eq_bounded_closed connected_open_path_connected hom homeomorphic_compactness by blast | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 5136 | |
| 68017 | 5137 | lemma path_connected_complement_homeomorphic_interval: | 
| 5138 | fixes S :: "'a::euclidean_space set" | |
| 5139 |   assumes "S homeomorphic cbox a b" "2 \<le> DIM('a)"
 | |
| 5140 | shows "path_connected(-S)" | |
| 5141 | using assms compact_cbox convex_box(1) path_connected_complement_homeomorphic_convex_compact by blast | |
| 5142 | ||
| 5143 | lemma connected_complement_homeomorphic_interval: | |
| 5144 | fixes S :: "'a::euclidean_space set" | |
| 5145 |   assumes "S homeomorphic cbox a b" "2 \<le> DIM('a)"
 | |
| 5146 | shows "connected(-S)" | |
| 5147 | using assms path_connected_complement_homeomorphic_interval path_connected_imp_connected by blast | |
| 5148 | ||
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 5149 | end |