| author | wenzelm | 
| Thu, 25 Jan 2018 14:13:55 +0100 | |
| changeset 67505 | ceb324e34c14 | 
| parent 67443 | 3abf6a722518 | 
| child 67673 | c8caefb20564 | 
| 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 | (* ========================================================================= *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 6 | (* Results connected with topological dimension. *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 7 | (* *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 8 | (* 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 | 9 | (* 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 | 10 | (* 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 | 11 | (* *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 12 | (* 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 | 13 | (* 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 | 14 | (* 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 | 15 | (* *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 16 | (* (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 | 17 | (* ========================================================================= *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 18 | |
| 60420 | 19 | section \<open>Results connected with topological dimension.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 20 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 21 | theory Brouwer_Fixpoint | 
| 63129 | 22 | 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 | 23 | begin | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 24 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 25 | lemma bij_betw_singleton_eq: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 26 | 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 | 27 | 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 | 28 | shows "f a = g a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 29 | proof - | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 30 |   have "f ` (A - {a}) = g ` (A - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 31 | by (intro image_cong) (simp_all add: eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 32 |   then have "B - {f a} = B - {g a}"
 | 
| 60303 | 33 | 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 | 34 | 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 | 35 | 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 | 36 | ultimately show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 37 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 38 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 39 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 40 | lemma swap_image: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 41 |   "Fun.swap i j f ` A = (if i \<in> A then (if j \<in> A then f ` A else f ` ((A - {i}) \<union> {j}))
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 42 |                                   else (if j \<in> A then f ` ((A - {j}) \<union> {i}) else f ` A))"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 43 | apply (auto simp: Fun.swap_def image_iff) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 44 | apply metis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 45 | apply (metis member_remove remove_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 46 | apply (metis member_remove remove_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 47 | done | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 48 | |
| 63365 | 49 | lemmas swap_apply1 = swap_apply(1) | 
| 50 | lemmas swap_apply2 = swap_apply(2) | |
| 51 | lemmas lessThan_empty_iff = Iio_eq_empty_iff_nat | |
| 52 | lemmas Zero_notin_Suc = zero_notin_Suc_image | |
| 53 | lemmas atMost_Suc_eq_insert_0 = Iic_Suc_eq_insert_0 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 54 | |
| 64267 | 55 | lemma sum_union_disjoint': | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 56 | assumes "finite A" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 57 | and "finite B" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 58 |     and "A \<inter> B = {}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 59 | and "A \<union> B = C" | 
| 64267 | 60 | shows "sum g C = sum g A + sum g B" | 
| 61 | using sum.union_disjoint[OF assms(1-3)] and assms(4) by auto | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 62 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 63 | lemma pointwise_minimal_pointwise_maximal: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 64 | fixes s :: "(nat \<Rightarrow> nat) set" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 65 | assumes "finite s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 66 |     and "s \<noteq> {}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 67 | 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 | 68 | 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 | 69 | 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 | 70 | using assms | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 71 | proof (induct s rule: finite_ne_induct) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 72 | case (insert b s) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 73 | assume *: "\<forall>x\<in>insert b s. \<forall>y\<in>insert b s. x \<le> y \<or> y \<le> x" | 
| 63540 | 74 | then obtain u l where "l \<in> s" "\<forall>b\<in>s. l \<le> b" "u \<in> s" "\<forall>b\<in>s. b \<le> u" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 75 | using insert by auto | 
| 63540 | 76 | with * show "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. a \<le> x" "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. x \<le> a" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 77 | 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 | 78 | 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 | 79 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 80 | lemma brouwer_compactness_lemma: | 
| 56226 | 81 | fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" | 
| 53674 | 82 | assumes "compact s" | 
| 83 | and "continuous_on s f" | |
| 53688 | 84 | and "\<not> (\<exists>x\<in>s. f x = 0)" | 
| 53674 | 85 | obtains d where "0 < d" and "\<forall>x\<in>s. d \<le> norm (f x)" | 
| 53185 | 86 | proof (cases "s = {}")
 | 
| 53674 | 87 | case True | 
| 53688 | 88 | show thesis | 
| 89 | by (rule that [of 1]) (auto simp: True) | |
| 53674 | 90 | next | 
| 49374 | 91 | case False | 
| 92 | have "continuous_on s (norm \<circ> f)" | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56273diff
changeset | 93 | by (rule continuous_intros continuous_on_norm assms(2))+ | 
| 53674 | 94 | with False obtain x where x: "x \<in> s" "\<forall>y\<in>s. (norm \<circ> f) x \<le> (norm \<circ> f) y" | 
| 95 | using continuous_attains_inf[OF assms(1), of "norm \<circ> f"] | |
| 96 | unfolding o_def | |
| 97 | by auto | |
| 98 | have "(norm \<circ> f) x > 0" | |
| 99 | using assms(3) and x(1) | |
| 100 | by auto | |
| 101 | then show ?thesis | |
| 102 | by (rule that) (insert x(2), auto simp: o_def) | |
| 49555 | 103 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 104 | |
| 49555 | 105 | 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 | 106 | fixes P Q :: "'a::euclidean_space \<Rightarrow> bool" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 107 | 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 | 108 | and "\<forall>x. P x \<longrightarrow> (\<forall>i\<in>Basis. Q i \<longrightarrow> 0 \<le> x\<bullet>i \<and> x\<bullet>i \<le> 1)" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 109 | 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 | 110 | (\<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 | 111 | (\<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 | 112 | (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (l x i = 0) \<longrightarrow> x\<bullet>i \<le> f x\<bullet>i) \<and> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 113 | (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (l x i = 1) \<longrightarrow> f x\<bullet>i \<le> x\<bullet>i)" | 
| 49374 | 114 | proof - | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 115 |   { fix x i
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 116 | 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 | 117 | (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 | 118 | (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 | 119 | (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 | 120 |     { assume "P x" "Q i" "i \<in> Basis" with assms have "0 \<le> f x \<bullet> i \<and> f x \<bullet> i \<le> 1" by auto }
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 121 | 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 | 122 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 123 | 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 | 124 | by (subst choice_iff[symmetric])+ blast | 
| 49374 | 125 | qed | 
| 126 | ||
| 53185 | 127 | |
| 60420 | 128 | subsection \<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 | 129 | |
| 53252 | 130 | lemma kuhn_counting_lemma: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 131 | fixes bnd compo compo' face S F | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 132 |   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 | 133 | 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 | 134 |     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 | 135 |     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 | 136 | 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 | 137 | 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 | 138 |     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 | 139 |   shows "odd (card {s\<in>S. compo s})"
 | 
| 53185 | 140 | proof - | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 141 | have "(\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) + (\<Sum>s | s \<in> S \<and> compo s. nF s) = (\<Sum>s\<in>S. nF s)" | 
| 64267 | 142 | 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 | 143 |   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 | 144 |                   (\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> \<not> bnd f}. face f s})"
 | 
| 64267 | 145 | unfolding sum.distrib[symmetric] | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 146 | by (subst card_Un_disjoint[symmetric]) | 
| 64267 | 147 | (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 | 148 |   also have "\<dots> = 1 * card {f\<in>F. compo' f \<and> bnd f} + 2 * card {f\<in>F. compo' f \<and> \<not> bnd f}"
 | 
| 67399 | 149 | 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 | 150 |   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 | 151 | using assms(6,8) by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 152 | 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 | 153 | (\<Sum>s | s \<in> S \<and> \<not> compo s \<and> nF s = 0. nF s) + (\<Sum>s | s \<in> S \<and> \<not> compo s \<and> nF s = 2. nF s)" | 
| 64267 | 154 | using assms(7) by (subst sum.union_disjoint[symmetric]) (fastforce intro!: sum.cong)+ | 
| 53688 | 155 | ultimately show ?thesis | 
| 156 | by auto | |
| 53186 | 157 | qed | 
| 158 | ||
| 60420 | 159 | subsection \<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 | 160 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 161 | lemma kuhn_complete_lemma: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 162 | assumes [simp]: "finite simplices" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 163 |     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 | 164 | 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 | 165 |     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 | 166 |     and bnd: "\<And>f s. s \<in> simplices \<Longrightarrow> face f s \<Longrightarrow> bnd f \<Longrightarrow> card {s\<in>simplices. face f s} = 1"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 167 |     and nbnd: "\<And>f s. s \<in> simplices \<Longrightarrow> face f s \<Longrightarrow> \<not> bnd f \<Longrightarrow> card {s\<in>simplices. face f s} = 2"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 168 |     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 | 169 |   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 | 170 | proof (rule kuhn_counting_lemma) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 171 | 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 | 172 | 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 | 173 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 174 |   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 | 175 |   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 | 176 | by (auto simp: face) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 177 | show "finite ?F" | 
| 60420 | 178 | 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 | 179 | |
| 60421 | 180 |   show "card {s \<in> simplices. face f s} = 1" if "f \<in> ?F" "bnd f" for f
 | 
| 60449 | 181 | using bnd that by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 182 | |
| 60421 | 183 |   show "card {s \<in> simplices. face f s} = 2" if "f \<in> ?F" "\<not> bnd f" for f
 | 
| 60449 | 184 | using nbnd that by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 185 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 186 |   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 | 187 | using odd_card by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 188 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 189 | fix s assume s[simp]: "s \<in> simplices" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 190 |   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 | 191 |   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 | 192 | using s by (fastforce simp: face) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 193 |   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 | 194 | by (auto intro!: card_image inj_onI) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 195 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 196 |   { assume rl: "rl ` s = {..Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 197 | then have inj_rl: "inj_on rl s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 198 | by (intro eq_card_imp_inj_on) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 199 | 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 | 200 | 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 | 201 |     ultimately have n: "{..n} = rl ` (s - {a})"
 | 
| 60303 | 202 | by (auto simp add: inj_on_image_set_diff Diff_subset rl) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 203 |     have "{a\<in>s. rl ` (s - {a}) = {..n}} = {a}"
 | 
| 60420 | 204 | using inj_rl \<open>a \<in> s\<close> by (auto simp add: 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 | 205 | then show "card ?S = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 206 | unfolding card_S by simp } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 207 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 208 |   { assume rl: "rl ` s \<noteq> {..Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 209 | show "card ?S = 0 \<or> card ?S = 2" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 210 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 211 |       assume *: "{..n} \<subseteq> rl ` s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 212 |       with rl rl_bd[OF s] have rl_s: "rl ` s = {..n}"
 | 
| 62390 | 213 | by (auto simp add: 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 | 214 | then have "\<not> inj_on rl s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 215 | by (intro pigeonhole) simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 216 | 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 | 217 | by (auto simp: inj_on_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 218 |       then have eq: "rl ` (s - {a}) = rl ` s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 219 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 220 |       with ab have inj: "inj_on rl (s - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 221 | by (intro eq_card_imp_inj_on) (auto simp add: rl_s card_Diff_singleton_if) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 222 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 223 |       { 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 | 224 |         then have "rl ` s - {rl x} = rl ` ((s - {a}) - {x})"
 | 
| 60303 | 225 | 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 | 226 |         also have "\<dots> = rl ` (s - {x})"
 | 
| 60420 | 227 |           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 | 228 | also assume "\<dots> = rl ` s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 229 | finally have False | 
| 60420 | 230 | using \<open>x\<in>s\<close> by auto } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 231 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 232 |       { 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 | 233 | 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 | 234 |       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 | 235 | unfolding rl_s[symmetric] by fastforce | 
| 60420 | 236 | 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 | 237 | unfolding card_S by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 238 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 239 |       assume "\<not> {..n} \<subseteq> rl ` s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 240 |       then have "\<And>x. rl ` (s - {x}) \<noteq> {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 241 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 242 | then show "card ?S = 0 \<or> card ?S = 2" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 243 | unfolding card_S by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 244 | qed } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 245 | qed fact | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 246 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 247 | locale kuhn_simplex = | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 248 | 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 | 249 |   assumes base: "base \<in> {..< n} \<rightarrow> {..< p}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 250 | 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 | 251 |   assumes upd: "bij_betw upd {..< n} {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 252 |   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 | 253 | begin | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 254 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 255 | 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 | 256 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 257 | lemma s_eq: "s = enum ` {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 258 | unfolding s_pre enum_def[abs_def] .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 259 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 260 | lemma upd_space: "i < n \<Longrightarrow> upd i < n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 261 | using upd by (auto dest!: bij_betwE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 262 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 263 | lemma s_space: "s \<subseteq> {..< n} \<rightarrow> {.. p}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 264 | proof - | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 265 |   { 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 | 266 | proof (induct i) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 267 | case 0 then show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 268 | 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 | 269 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 270 | case (Suc i) with base show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 271 | 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 | 272 | qed } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 273 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 274 | by (auto simp: s_eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 275 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 276 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 277 | lemma inj_upd: "inj_on upd {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 278 | using upd by (simp add: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 279 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 280 | lemma inj_enum: "inj_on enum {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 281 | proof - | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 282 |   { 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 | 283 |     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 | 284 |       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 | 285 | then have "enum x \<noteq> enum y" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 286 | by (auto simp add: enum_def fun_eq_iff) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 287 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 288 | by (auto simp: inj_on_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 289 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 290 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 291 | lemma enum_0: "enum 0 = base" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 292 | by (simp add: enum_def[abs_def]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 293 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 294 | lemma base_in_s: "base \<in> s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 295 | unfolding s_eq by (subst enum_0[symmetric]) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 296 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 297 | 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 | 298 | unfolding s_eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 299 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 300 | lemma one_step: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 301 | assumes a: "a \<in> s" "j < n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 302 | 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 | 303 | shows "a j \<noteq> p'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 304 | proof | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 305 | assume "a j = p'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 306 | 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 | 307 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 308 | 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 | 309 | unfolding s_eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 310 |   from this[of 0] this[of n] have "j \<notin> upd ` {..< n}"
 | 
| 62390 | 311 | by (auto simp: enum_def fun_eq_iff split: if_split_asm) | 
| 60420 | 312 | with upd \<open>j < n\<close> show False | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 313 | by (auto simp: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 314 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 315 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 316 | 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 | 317 | 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 | 318 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 319 | lemma upd_surj: "upd ` {..< n} = {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 320 | using upd by (auto simp: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 321 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 322 | 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 | 323 |   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 | 324 | by (auto simp: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 325 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 326 | 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 | 327 | 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 | 328 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 329 | 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 | 330 | 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 | 331 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 332 | 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 | 333 | 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 | 334 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 335 | lemma enum_strict_mono: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i < enum j \<longleftrightarrow> i < j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 336 | using enum_mono[of i j] enum_inj[of i j] by (auto simp add: le_less) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 337 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 338 | 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 | 339 | by (auto simp: s_eq enum_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 340 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 341 | 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 | 342 | 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 | 343 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 344 | 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 | 345 | 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 | 346 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 347 | 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 | 348 | 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 | 349 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 350 | 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 | 351 | 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 | 352 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 353 | 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 | 354 | 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 | 355 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 356 | lemma out_eq_p: "a \<in> s \<Longrightarrow> n \<le> j \<Longrightarrow> a j = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 357 | unfolding s_eq by (auto simp add: enum_eq_p) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 358 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 359 | 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 | 360 | 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 | 361 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 362 | 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 | 363 | unfolding s_eq by (auto simp: enum_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 364 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 365 | 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 | 366 | unfolding s_eq by (auto simp: enum_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 367 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 368 | 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 | 369 | using enum_in[of i] s_space by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 370 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 371 | 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 | 372 | 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 | 373 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 374 | lemma ksimplex_0: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 375 |   "n = 0 \<Longrightarrow> s = {(\<lambda>x. p)}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 376 | using s_eq enum_def base_out by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 377 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 378 | lemma replace_0: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 379 |   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 | 380 | shows "x \<le> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 381 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 382 | assume "x \<noteq> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 383 | have "a j \<noteq> 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 384 | using assms by (intro one_step[where a=a]) auto | 
| 60420 | 385 | with less[OF \<open>x\<in>s\<close> \<open>a\<in>s\<close>, of j] p[rule_format, of x] \<open>x \<in> s\<close> \<open>x \<noteq> a\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 386 | show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 387 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 388 | qed simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 389 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 390 | lemma replace_1: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 391 |   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 | 392 | shows "a \<le> x" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 393 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 394 | assume "x \<noteq> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 395 | have "a j \<noteq> p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 396 | using assms by (intro one_step[where a=a]) auto | 
| 60420 | 397 | 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 | 398 | have "a j < p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 399 | by (auto simp: less_le s_eq) | 
| 60420 | 400 | with less[OF \<open>a\<in>s\<close> \<open>x\<in>s\<close>, of j] p[rule_format, of x] \<open>x \<in> s\<close> \<open>x \<noteq> a\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 401 | show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 402 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 403 | qed simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 404 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 405 | end | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 406 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 407 | locale kuhn_simplex_pair = s: kuhn_simplex p n b_s u_s s + t: kuhn_simplex p n b_t u_t t | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 408 | 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 | 409 | begin | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 410 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 411 | lemma enum_eq: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 412 | 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 | 413 |   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 | 414 | shows "s.enum l = t.enum (l + d)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 415 | using l proof (induct l rule: dec_induct) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 416 | case base | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 417 |   then have s: "s.enum i \<in> t.enum ` {i + d .. j + d}" and t: "t.enum (i + d) \<in> s.enum ` {i .. j}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 418 | using eq by auto | 
| 60420 | 419 | 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 | 420 | by (auto simp: s.enum_mono) | 
| 60420 | 421 | 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 | 422 | by (auto simp: t.enum_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 423 | ultimately show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 424 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 425 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 426 | case (step l) | 
| 60420 | 427 | 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 | 428 | "s.enum l < s.enum (Suc l)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 429 | "t.enum (l + d) < t.enum (Suc l + d)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 430 | 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 | 431 | moreover have | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 432 |       "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 | 433 |       "t.enum (Suc l + d) \<in> s.enum ` {i .. j}"
 | 
| 60420 | 434 | 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 | 435 | ultimately have "s.enum (Suc l) = t.enum (Suc (l + d))" | 
| 60420 | 436 | 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 | 437 | 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 | 438 | (auto intro!: s.enum_in t.enum_in) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 439 | then show ?case by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 440 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 441 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 442 | lemma ksimplex_eq_bot: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 443 | 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 | 444 | 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 | 445 |   assumes eq: "s - {a} = t - {b}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 446 | shows "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 447 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 448 | 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 | 449 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 450 | assume "n \<noteq> 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 451 | 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 | 452 | "t.enum 0 = (t.enum (Suc 0)) (u_t 0 := t.enum (Suc 0) (u_t 0) - 1)" | 
| 60420 | 453 | 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 | 454 | 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 | 455 | 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 | 456 | 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 | 457 |   { fix j assume "0 < j" "j \<le> n"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 458 |     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 | 459 | 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 | 460 | ultimately have "s.enum j = t.enum j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 461 | 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 | 462 | note enum_eq = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 463 | then have "s.enum (Suc 0) = t.enum (Suc 0)" | 
| 60420 | 464 | using \<open>n \<noteq> 0\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 465 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 466 |   { fix j assume "Suc j < n"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 467 | 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 | 468 | have "u_s (Suc j) = u_t (Suc j)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 469 | using s.enum_Suc[of "Suc j"] t.enum_Suc[of "Suc j"] | 
| 62390 | 470 | 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 | 471 | 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 | 472 | by (auto simp: gr0_conv_Suc) | 
| 60420 | 473 | 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 | 474 | 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 | 475 | ultimately have "a = b" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 476 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 477 | with assms show "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 478 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 479 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 480 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 481 | lemma ksimplex_eq_top: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 482 | 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 | 483 | 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 | 484 |   assumes eq: "s - {a} = t - {b}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 485 | shows "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 486 | proof (cases n) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 487 | 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 | 488 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 489 | case (Suc n') | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 490 | 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 | 491 | "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 | 492 | 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 | 493 | 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 | 494 | 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 | 495 | 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 | 496 |   { fix j assume "j < n"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 497 |     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 | 498 | 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 | 499 | ultimately have "s.enum j = t.enum j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 500 | 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 | 501 | note enum_eq = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 502 | then have "s.enum n' = t.enum n'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 503 | using Suc by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 504 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 505 |   { fix j assume "j < n'"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 506 | 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 | 507 | have "u_s j = u_t j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 508 | using s.enum_Suc[of j] t.enum_Suc[of j] | 
| 62390 | 509 | 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 | 510 | 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 | 511 | by (auto simp: gr0_conv_Suc) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 512 | then have "u_t n' = u_s n'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 513 | 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 | 514 | ultimately have "a = b" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 515 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 516 | with assms show "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 517 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 518 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 519 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 520 | end | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 521 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 522 | inductive ksimplex for p n :: nat where | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 523 | 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 | 524 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 525 | lemma finite_ksimplexes: "finite {s. ksimplex p n s}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 526 | proof (rule finite_subset) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 527 |   { 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 | 528 | 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 | 529 | then interpret kuhn_simplex p n b u s . | 
| 60420 | 530 | 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 | 531 |     have "a \<in> (\<lambda>f x. if n \<le> x then p else f x) ` ({..< n} \<rightarrow>\<^sub>E {.. p})"
 | 
| 62390 | 532 | 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 | 533 |                intro!: bexI[of _ "restrict a {..< n}"]) }
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 534 |   then show "{s. ksimplex p n s} \<subseteq> Pow ((\<lambda>f x. if n \<le> x then p else f x) ` ({..< n} \<rightarrow>\<^sub>E {.. p}))"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 535 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 536 | qed (simp add: finite_PiE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 537 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 538 | lemma ksimplex_card: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 539 | 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 | 540 | using assms proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 541 | case (ksimplex u b) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 542 | then interpret kuhn_simplex p n u b s . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 543 | show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 544 | by (simp add: card_image s_eq inj_enum) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 545 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 546 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 547 | lemma simplex_top_face: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 548 | 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 | 549 |   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 | 550 | using assms | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 551 | proof safe | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 552 |   fix s a assume "ksimplex p (Suc n) s" and a: "a \<in> s" and na: "\<forall>x\<in>s - {a}. x n = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 553 |   then show "ksimplex p n (s - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 554 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 555 | case (ksimplex base upd) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 556 | 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 | 557 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 558 | have "a n < p" | 
| 60420 | 559 | 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 | 560 | then have "a = enum 0" | 
| 60420 | 561 | 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 | 562 |     then have s_eq: "s - {a} = enum ` Suc ` {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 563 | 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 | 564 |     then have "enum 1 \<in> s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 565 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 566 | then have "upd 0 = n" | 
| 60420 | 567 | using \<open>a n < p\<close> \<open>a = enum 0\<close> na[rule_format, of "enum 1"] | 
| 62390 | 568 | 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 | 569 |     then have "bij_betw upd (Suc ` {..< n}) {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 570 | using upd | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 571 | by (subst notIn_Un_bij_betw3[where b=0]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 572 | (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 | 573 |     then have "bij_betw (upd\<circ>Suc) {..<n} {..<n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 574 | 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 | 575 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 576 | have "a n = p - 1" | 
| 60420 | 577 |       using enum_Suc[of 0] na[rule_format, OF \<open>enum 1 \<in> s - {a}\<close>] \<open>a = enum 0\<close> by (auto simp: \<open>upd 0 = n\<close>)
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 578 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 579 | show ?thesis | 
| 61169 | 580 | proof (rule ksimplex.intros, standard) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 581 |       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 | 582 |       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 | 583 | using base base_out by (auto simp: Pi_iff) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 584 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 585 |       have "\<And>i. Suc ` {..< i} = {..< Suc i} - {0}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 586 | by (auto simp: image_iff Ball_def) arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 587 |       then have upd_Suc: "\<And>i. i \<le> n \<Longrightarrow> (upd\<circ>Suc) ` {..< i} = upd ` {..< Suc i} - {n}"
 | 
| 60420 | 588 | using \<open>upd 0 = n\<close> upd_inj | 
| 60303 | 589 | by (auto simp add: 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 | 590 |       have n_in_upd: "\<And>i. n \<in> upd ` {..< Suc i}"
 | 
| 60420 | 591 | using \<open>upd 0 = n\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 592 | |
| 63040 | 593 | define f' where "f' i j = | 
| 594 |         (if j \<in> (upd\<circ>Suc)`{..< i} then Suc ((base(n := p)) j) else (base(n := p)) j)" for i j
 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 595 |       { fix x i assume i[arith]: "i \<le> n" then have "enum (Suc i) x = f' i x"
 | 
| 60420 | 596 | unfolding f'_def enum_def using \<open>a n < p\<close> \<open>a = enum 0\<close> \<open>upd 0 = n\<close> \<open>a n = p - 1\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 597 | 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 | 598 |       then show "s - {a} = f' ` {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 599 | 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 | 600 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 601 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 602 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 603 | 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 | 604 |   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 | 605 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 606 | case (ksimplex base upd) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 607 | then interpret kuhn_simplex p n base upd s' . | 
| 63040 | 608 | define b where "b = base (n := p - 1)" | 
| 609 | define u where "u i = (case i of 0 \<Rightarrow> n | Suc i \<Rightarrow> upd i)" for i | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 610 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 611 |     have "ksimplex p (Suc n) (s' \<union> {b})"
 | 
| 61169 | 612 | proof (rule ksimplex.intros, standard) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 613 |       show "b \<in> {..<Suc n} \<rightarrow> {..<p}"
 | 
| 60420 | 614 | 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 | 615 | 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 | 616 | using base_out by (auto simp: b_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 617 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 618 |       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 | 619 | using upd | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 620 | 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 | 621 |       then show "bij_betw u {..<Suc n} {..<Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 622 | 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 | 623 | |
| 63040 | 624 |       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 | 625 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 626 |       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 | 627 | 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 | 628 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 629 |       { 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 | 630 | 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 | 631 | note n_not_upd = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 632 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 633 |       have *: "f' ` {.. Suc n} = f' ` (Suc ` {.. n} \<union> {0})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 634 | unfolding atMost_Suc_eq_insert_0 by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 635 |       also have "\<dots> = (f' \<circ> Suc) ` {.. n} \<union> {b}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 636 | by (auto simp: f'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 637 |       also have "(f' \<circ> Suc) ` {.. n} = s'"
 | 
| 60420 | 638 | 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 | 639 | 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 | 640 | 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 | 641 |       finally show "s' \<union> {b} = f' ` {.. Suc n}" ..
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 642 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 643 | moreover have "b \<notin> s'" | 
| 60420 | 644 | 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 | 645 | ultimately show ?thesis by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 646 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 647 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 648 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 649 | lemma ksimplex_replace_0: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 650 | 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 | 651 |   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 | 652 |   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 | 653 | using s | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 654 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 655 | case (ksimplex b_s u_s) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 656 | |
| 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 | 657 |   { fix t b assume "ksimplex p n t"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 658 | 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 | 659 | by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 660 | 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 | 661 | by intro_locales fact+ | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 662 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 663 |     assume b: "b \<in> t" "t - {b} = s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 664 | 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 | 665 | by (intro ksimplex_eq_top[of a b]) auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 666 |   then have "{s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = {s}"
 | 
| 60420 | 667 | 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 | 668 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 669 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 670 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 671 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 672 | lemma ksimplex_replace_1: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 673 | 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 | 674 |   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 | 675 |   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 | 676 | using s | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 677 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 678 | case (ksimplex b_s u_s) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 679 | |
| 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 | 680 |   { fix t b assume "ksimplex p n t"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 681 | 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 | 682 | by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 683 | 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 | 684 | by intro_locales fact+ | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 685 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 686 |     assume b: "b \<in> t" "t - {b} = s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 687 | 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 | 688 | by (intro ksimplex_eq_bot[of a b]) auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 689 |   then have "{s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = {s}"
 | 
| 60420 | 690 | 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 | 691 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 692 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 693 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 694 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 695 | lemma card_2_exists: "card s = 2 \<longleftrightarrow> (\<exists>x\<in>s. \<exists>y\<in>s. x \<noteq> y \<and> (\<forall>z\<in>s. z = x \<or> z = y))" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 696 | by (auto simp add: card_Suc_eq eval_nat_numeral) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 697 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 698 | lemma ksimplex_replace_2: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 699 | 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 | 700 |     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 | 701 |     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 | 702 |   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 | 703 | using s | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 704 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 705 | case (ksimplex base upd) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 706 | then interpret kuhn_simplex p n base upd s . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 707 | |
| 60420 | 708 | 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 | 709 | unfolding s_eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 710 | |
| 60420 | 711 | 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 | 712 | by linarith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 713 |   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 | 714 | proof (elim disjE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 715 | assume "i = 0" | 
| 63040 | 716 | 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 | 717 | let ?upd = "upd \<circ> rot" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 718 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 719 |     have rot: "bij_betw rot {..< n} {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 720 | 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 | 721 | arith+ | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 722 |     from rot upd have "bij_betw ?upd {..<n} {..<n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 723 | by (rule bij_betw_trans) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 724 | |
| 63040 | 725 | define f' where [abs_def]: "f' i j = | 
| 726 |       (if j \<in> ?upd`{..< i} then Suc (enum (Suc 0) j) else enum (Suc 0) j)" for i j
 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 727 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 728 |     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 | 729 | proof | 
| 60420 | 730 | 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 | 731 | 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 | 732 | 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 | 733 | then have "enum 1 \<le> enum i'" "enum i' (upd 0) < p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 734 | using enum_le_p[of i' "upd 0"] by (auto simp add: enum_inj enum_mono upd_space) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 735 | then have "enum 1 (upd 0) < p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 736 | by (auto simp add: le_fun_def intro: le_less_trans) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 737 |       then show "enum (Suc 0) \<in> {..<n} \<rightarrow> {..<p}"
 | 
| 60420 | 738 | using base \<open>n \<noteq> 0\<close> by (auto simp add: 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 | 739 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 740 |       { fix i assume "n \<le> i" then show "enum (Suc 0) i = p"
 | 
| 60420 | 741 | 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 | 742 |       show "bij_betw ?upd {..<n} {..<n}" by fact
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 743 | qed (simp add: f'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 744 |     have ks_f': "ksimplex p n (f' ` {.. n})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 745 | by rule unfold_locales | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 746 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 747 | 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 | 748 |     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 | 749 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 750 |     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 | 751 | by (auto simp: rot_def image_iff Ball_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 752 | arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 753 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 754 |     { fix j assume j: "j < n"
 | 
| 60420 | 755 | from j \<open>n \<noteq> 0\<close> have "f' j = enum (Suc j)" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 756 | by (auto simp add: f'_def enum_def upd_inj in_upd_image image_comp[symmetric] fun_eq_iff) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 757 | note f'_eq_enum = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 758 |     then have "enum ` Suc ` {..< n} = f' ` {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 759 | by (force simp: enum_inj) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 760 |     also have "Suc ` {..< n} = {.. n} - {0}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 761 | by (auto simp: image_iff Ball_def) arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 762 |     also have "{..< n} = {.. n} - {n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 763 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 764 |     finally have eq: "s - {a} = f' ` {.. n} - {f' n}"
 | 
| 60420 | 765 | unfolding s_eq \<open>a = enum i\<close> \<open>i = 0\<close> | 
| 60303 | 766 | 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 | 767 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 768 | have "enum 0 < f' 0" | 
| 60420 | 769 | 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 | 770 | also have "\<dots> < f' n" | 
| 60420 | 771 | 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 | 772 | finally have "a \<noteq> f' n" | 
| 60420 | 773 | 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 | 774 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 775 |     { 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 | 776 | obtain b u where "kuhn_simplex p n b u t" | 
| 60420 | 777 | 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 | 778 | 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 | 779 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 780 |       { fix x assume "x \<in> s" "x \<noteq> a"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 781 | then have "x (upd 0) = enum (Suc 0) (upd 0)" | 
| 60420 | 782 | 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 | 783 |       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 | 784 | unfolding eq_sma[symmetric] by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 785 | then have "c (upd 0) \<noteq> enum (Suc 0) (upd 0)" | 
| 60420 | 786 | 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 | 787 | 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 | 788 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 789 |       then have "t = s \<or> t = f' ` {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 790 | proof (elim disjE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 791 | assume *: "c (upd 0) < enum (Suc 0) (upd 0)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 792 | interpret st: kuhn_simplex_pair p n base upd s b u t .. | 
| 60420 | 793 |         { 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 | 794 | 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 | 795 | note top = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 796 | have "s = t" | 
| 60420 | 797 | 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 | 798 | by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq_sma]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 799 | (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 | 800 | then show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 801 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 802 | assume *: "c (upd 0) > enum (Suc 0) (upd 0)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 803 |         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 | 804 |         have eq: "f' ` {..n} - {f' n} = t - {c}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 805 | using eq_sma eq by simp | 
| 60420 | 806 |         { 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 | 807 | 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 | 808 | note top = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 809 |         have "f' ` {..n} = t"
 | 
| 60420 | 810 | 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 | 811 | by (intro st.ksimplex_eq_top[OF _ _ _ _ eq]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 812 | (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 | 813 | then show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 814 | qed } | 
| 60420 | 815 | 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 | 816 |       apply (intro ex1I[of _ "f' ` {.. n}"])
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 817 | apply auto [] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 818 | apply metis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 819 | done | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 820 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 821 | assume "i = n" | 
| 60420 | 822 | 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 | 823 | by (cases n) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 824 | |
| 63040 | 825 | 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 | 826 | let ?upd = "upd \<circ> rot" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 827 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 828 |     have rot: "bij_betw rot {..< n} {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 829 | 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 | 830 | arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 831 |     from rot upd have "bij_betw ?upd {..<n} {..<n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 832 | by (rule bij_betw_trans) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 833 | |
| 63040 | 834 | define b where "b = base (upd n' := base (upd n') - 1)" | 
| 835 |     define f' where [abs_def]: "f' i j = (if j \<in> ?upd`{..< i} then Suc (b j) else b j)" for i j
 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 836 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 837 |     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 | 838 | proof | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 839 |       { 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 | 840 | 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 | 841 |       show "b \<in> {..<n} \<rightarrow> {..<p}"
 | 
| 60420 | 842 | 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 | 843 | 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 | 844 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 845 |       show "bij_betw ?upd {..<n} {..<n}" by fact
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 846 | qed (simp add: f'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 847 | 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 | 848 |     have ks_f': "ksimplex p n (b.enum ` {.. n})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 849 | unfolding f' by rule unfold_locales | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 850 | |
| 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 | 851 | have "0 < n" | 
| 60420 | 852 | using \<open>n \<noteq> 0\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 853 | |
| 60420 | 854 |     { 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 | 855 | 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 | 856 | unfolding s_eq by (auto simp: enum_inj n') | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 857 | moreover have "enum i' (upd n') = base (upd n')" | 
| 60420 | 858 | unfolding enum_def using \<open>i' \<le> n\<close> \<open>enum i' \<noteq> enum n\<close> by (auto simp: n' upd_inj enum_inj) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 859 | ultimately have "0 < base (upd n')" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 860 | by auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 861 | then have benum1: "b.enum (Suc 0) = base" | 
| 60420 | 862 | 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 | 863 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 864 |     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 | 865 | 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 | 866 | 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 | 867 | by (simp_all add: rot_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 868 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 869 |     { fix j assume j: "Suc j \<le> n" then have "b.enum (Suc j) = enum j"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 870 | by (induct j) (auto simp add: benum1 enum_0 b.enum_Suc enum_Suc rot_simps) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 871 | note b_enum_eq_enum = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 872 |     then have "enum ` {..< n} = b.enum ` Suc ` {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 873 | by (auto simp add: image_comp intro!: image_cong) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 874 |     also have "Suc ` {..< n} = {.. n} - {0}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 875 | by (auto simp: image_iff Ball_def) arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 876 |     also have "{..< n} = {.. n} - {n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 877 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 878 |     finally have eq: "s - {a} = b.enum ` {.. n} - {b.enum 0}"
 | 
| 60420 | 879 | unfolding s_eq \<open>a = enum i\<close> \<open>i = n\<close> | 
| 60303 | 880 |       using inj_on_image_set_diff[OF inj_enum Diff_subset, of "{n}"]
 | 
| 881 |             inj_on_image_set_diff[OF b.inj_enum Diff_subset, of "{0}"]
 | |
| 882 | by (simp add: comp_def ) | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 883 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 884 | have "b.enum 0 \<le> b.enum n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 885 | by (simp add: b.enum_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 886 | also have "b.enum n < enum n" | 
| 60420 | 887 | 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 | 888 | finally have "a \<noteq> b.enum 0" | 
| 60420 | 889 | 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 | 890 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 891 |     { 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 | 892 | obtain b' u where "kuhn_simplex p n b' u t" | 
| 60420 | 893 | 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 | 894 | 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 | 895 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 896 |       { fix x assume "x \<in> s" "x \<noteq> a"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 897 | then have "x (upd n') = enum n' (upd n')" | 
| 60420 | 898 | 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 | 899 |       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 | 900 | unfolding eq_sma[symmetric] by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 901 | then have "c (upd n') \<noteq> enum n' (upd n')" | 
| 60420 | 902 | using \<open>n \<noteq> 0\<close> by (intro t.one_step[OF \<open>c\<in>t\<close> ]) (auto simp: n' upd_space[unfolded n']) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 903 | 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 | 904 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 905 |       then have "t = s \<or> t = b.enum ` {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 906 | proof (elim disjE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 907 | assume *: "c (upd n') > enum n' (upd n')" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 908 | interpret st: kuhn_simplex_pair p n base upd s b' u t .. | 
| 60420 | 909 |         { 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 | 910 | 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 | 911 | note top = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 912 | have "s = t" | 
| 60420 | 913 | 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 | 914 | by (intro st.ksimplex_eq_top[OF _ _ _ _ eq_sma]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 915 | (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 | 916 | then show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 917 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 918 | assume *: "c (upd n') < enum n' (upd n')" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 919 |         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 | 920 |         have eq: "f' ` {..n} - {b.enum 0} = t - {c}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 921 | using eq_sma eq f' by simp | 
| 60420 | 922 |         { 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 | 923 | 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 | 924 | note bot = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 925 |         have "f' ` {..n} = t"
 | 
| 60420 | 926 | 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 | 927 | by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 928 | (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 | 929 | with f' show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 930 | qed } | 
| 60420 | 931 | 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 | 932 |       apply (intro ex1I[of _ "b.enum ` {.. n}"])
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 933 | apply auto [] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 934 | apply metis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 935 | done | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 936 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 937 | assume i: "0 < i" "i < n" | 
| 63040 | 938 | define i' where "i' = i - 1" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 939 | with i have "Suc i' < n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 940 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 941 | with i have Suc_i': "Suc i' = i" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 942 | by (simp add: i'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 943 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 944 | let ?upd = "Fun.swap i' i upd" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 945 |     from i upd have "bij_betw ?upd {..< n} {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 946 | 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 | 947 | |
| 63040 | 948 |     define f' where [abs_def]: "f' i j = (if j \<in> ?upd`{..< i} then Suc (base j) else base j)"
 | 
| 949 | for i j | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 950 |     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 | 951 | proof | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 952 |       show "base \<in> {..<n} \<rightarrow> {..<p}" by fact
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 953 |       { fix i assume "n \<le> i" then show "base i = p" by fact }
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 954 |       show "bij_betw ?upd {..<n} {..<n}" by fact
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 955 | qed (simp add: f'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 956 | 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 | 957 |     have ks_f': "ksimplex p n (b.enum ` {.. n})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 958 | unfolding f' by rule unfold_locales | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 959 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 960 |     have "{i} \<subseteq> {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 961 | using i by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 962 |     { fix j assume "j \<le> n"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 963 | 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 | 964 | moreover note i | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 965 | 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 | 966 | unfolding enum_def[abs_def] b.enum_def[abs_def] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 967 | by (auto simp add: fun_eq_iff swap_image i'_def | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 968 | 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 | 969 | note enum_eq_benum = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 970 |     then have "enum ` ({.. n} - {i}) = b.enum ` ({.. n} - {i})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 971 | by (intro image_cong) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 972 |     then have eq: "s - {a} = b.enum ` {.. n} - {b.enum i}"
 | 
| 60420 | 973 | unfolding s_eq \<open>a = enum i\<close> | 
| 974 |       using inj_on_image_set_diff[OF inj_enum Diff_subset \<open>{i} \<subseteq> {..n}\<close>]
 | |
| 975 |             inj_on_image_set_diff[OF b.inj_enum Diff_subset \<open>{i} \<subseteq> {..n}\<close>]
 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 976 | by (simp add: comp_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 977 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 978 | have "a \<noteq> b.enum i" | 
| 60420 | 979 | 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 | 980 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 981 |     { 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 | 982 | obtain b' u where "kuhn_simplex p n b' u t" | 
| 60420 | 983 | 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 | 984 | 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 | 985 |       have "enum i' \<in> s - {a}" "enum (i + 1) \<in> s - {a}"
 | 
| 60420 | 986 | 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 | 987 | then obtain l k where | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 988 | 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 | 989 | 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 | 990 | unfolding eq_sma by (auto simp: t.s_eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 991 | with i have "t.enum l < t.enum k" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 992 | by (simp add: enum_strict_mono i'_def) | 
| 60420 | 993 | 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 | 994 | by (simp add: t.enum_strict_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 995 |       { assume "Suc l = k"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 996 | have "enum (Suc (Suc i')) = t.enum (Suc l)" | 
| 60420 | 997 | 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 | 998 | then have False | 
| 60420 | 999 | using \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close> | 
| 62390 | 1000 | 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 | 1001 | (metis Suc_lessD n_not_Suc_n upd_inj) } | 
| 60420 | 1002 | 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 | 1003 | by arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1004 | have c_eq: "c = t.enum (Suc l)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1005 | proof (rule ccontr) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1006 | assume "c \<noteq> t.enum (Suc l)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1007 |         then have "t.enum (Suc l) \<in> s - {a}"
 | 
| 60420 | 1008 | 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 | 1009 | then obtain j where "t.enum (Suc l) = enum j" "j \<le> n" "enum j \<noteq> enum i" | 
| 60420 | 1010 | 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 | 1011 | with i have "t.enum (Suc l) \<le> t.enum l \<or> t.enum k \<le> t.enum (Suc l)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1012 | by (auto simp add: i'_def enum_mono enum_inj l k) | 
| 60420 | 1013 | 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 | 1014 | by (simp add: t.enum_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1015 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1016 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1017 |       { have "t.enum (Suc (Suc l)) \<in> s - {a}"
 | 
| 60420 | 1018 | unfolding eq_sma c_eq t.s_eq using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (auto simp: t.enum_inj) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1019 | then obtain j where eq: "t.enum (Suc (Suc l)) = enum j" and "j \<le> n" "j \<noteq> i" | 
| 60420 | 1020 | 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 | 1021 | moreover have "enum i' < t.enum (Suc (Suc l))" | 
| 60420 | 1022 | 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 | 1023 | ultimately have "i' < j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1024 | using i by (simp add: enum_strict_mono i'_def) | 
| 60420 | 1025 | 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 | 1026 | 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 | 1027 | then have "k \<le> Suc (Suc l)" | 
| 60420 | 1028 | using \<open>k \<le> n\<close> \<open>Suc l < k\<close> by (simp add: t.enum_mono) } | 
| 1029 | with \<open>Suc l < k\<close> have "Suc (Suc l) = k" by simp | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1030 | 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 | 1031 | using i by (simp add: k i'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1032 | also have "\<dots> = (enum i') (u l := Suc (enum i' (u l)), u (Suc l) := Suc (enum i' (u (Suc l))))" | 
| 60420 | 1033 | 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 | 1034 | 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 | 1035 | (u l = upd (Suc i') \<and> u (Suc l) = upd i')" | 
| 62390 | 1036 | 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 | 1037 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1038 |       then have "t = s \<or> t = b.enum ` {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1039 | proof (elim disjE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1040 | assume u: "u l = upd i'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1041 | have "c = t.enum (Suc l)" unfolding c_eq .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1042 | also have "t.enum (Suc l) = enum (Suc i')" | 
| 60420 | 1043 | using u \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close> by (simp add: enum_Suc t.enum_Suc l) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1044 | also have "\<dots> = a" | 
| 60420 | 1045 | 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 | 1046 | finally show ?thesis | 
| 60420 | 1047 | 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 | 1048 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1049 | assume u: "u l = upd (Suc i')" | 
| 63040 | 1050 |         define B where "B = b.enum ` {..n}"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1051 | have "b.enum i' = enum i'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1052 | using enum_eq_benum[of i'] i by (auto simp add: i'_def gr0_conv_Suc) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1053 | have "c = t.enum (Suc l)" unfolding c_eq .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1054 | also have "t.enum (Suc l) = b.enum (Suc i')" | 
| 60420 | 1055 | using u \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close> | 
| 1056 | by (simp_all add: enum_Suc t.enum_Suc l b.enum_Suc \<open>b.enum i' = enum i'\<close> swap_apply1) | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1057 | (simp add: Suc_i') | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1058 | also have "\<dots> = b.enum i" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1059 | using i by (simp add: i'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1060 | finally have "c = b.enum i" . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1061 |         then have "t - {c} = B - {c}" "c \<in> B"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1062 | unfolding eq_sma[symmetric] eq B_def using i by auto | 
| 60420 | 1063 | 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 | 1064 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1065 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1066 | by (simp add: B_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1067 | qed } | 
| 60420 | 1068 | 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 | 1069 |       apply (intro ex1I[of _ "b.enum ` {.. n}"])
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1070 | apply auto [] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1071 | apply metis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1072 | done | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1073 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1074 | then show ?thesis | 
| 60420 | 1075 | 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 | 1076 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1077 | |
| 60420 | 1078 | 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 | 1079 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1080 | lemma kuhn_simplex_lemma: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1081 |   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 | 1082 |     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 | 1083 |       rl ` f = {..n} \<and> ((\<exists>j\<le>n. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>f. x j = p))})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1084 |   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 | 1085 | 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 | 1086 |     where bnd="\<lambda>f. (\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = p)"],
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1087 | safe del: notI) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1088 | |
| 53186 | 1089 | have *: "\<And>x y. x = y \<Longrightarrow> odd (card x) \<Longrightarrow> odd (card y)" | 
| 1090 | by auto | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1091 |   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 | 1092 |     rl ` f = {..n} \<and> ((\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = p))})"
 | 
| 53186 | 1093 | apply (rule *[OF _ assms(2)]) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1094 | apply (auto simp: atLeast0AtMost) | 
| 53186 | 1095 | done | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1096 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1097 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1098 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1099 | fix s assume s: "ksimplex p (Suc n) s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1100 | then show "card s = n + 2" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1101 | by (simp add: ksimplex_card) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1102 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1103 | 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 | 1104 | using assms(1) s by (auto simp: subset_eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1105 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1106 |   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 | 1107 |   { 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 | 1108 | with s a show "card ?S = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1109 | 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 | 1110 | by (subst eq_commute) simp } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1111 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1112 |   { 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 | 1113 | with s a show "card ?S = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1114 | 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 | 1115 | by (subst eq_commute) simp } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1116 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1117 |   { 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 | 1118 |     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 | 1119 | 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 | 1120 | by (subst (asm) eq_commute) auto } | 
| 53186 | 1121 | qed | 
| 1122 | ||
| 60420 | 1123 | subsection \<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 | 1124 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1125 | 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 | 1126 | |
| 53186 | 1127 | lemma reduced_labelling: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1128 | shows "reduced n x \<le> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1129 | and "\<forall>i<reduced n x. x i = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1130 | and "reduced n x = n \<or> x (reduced n x) \<noteq> 0" | 
| 53186 | 1131 | proof - | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1132 | show "reduced n x \<le> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1133 | 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 | 1134 | show "\<forall>i<reduced n x. x i = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1135 | 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 | 1136 | 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 | 1137 | unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) fastforce+ | 
| 53186 | 1138 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1139 | |
| 53186 | 1140 | lemma reduced_labelling_unique: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1141 | "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 | 1142 | 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 | 1143 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1144 | 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 | 1145 | 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 | 1146 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1147 | lemma reduce_labelling_zero[simp]: "reduced 0 x = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1148 | 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 | 1149 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1150 | 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 | 1151 | 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 | 1152 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1153 | 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 | 1154 | using reduced_labelling[of "Suc n" x] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1155 | 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 | 1156 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1157 | lemma complete_face_top: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1158 | 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 | 1159 | 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 | 1160 |     and eq: "(reduced (Suc n) \<circ> lab) ` f = {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1161 | shows "((\<exists>j\<le>n. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>f. x j = p)) \<longleftrightarrow> (\<forall>x\<in>f. x n = p)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1162 | proof (safe del: disjCI) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1163 | 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 | 1164 |   { 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 | 1165 | by (intro reduced_labelling_zero) auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1166 | 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 | 1167 | using j eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1168 | ultimately show "x n = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1169 | by force | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1170 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1171 | 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 | 1172 | have "j = n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1173 | proof (rule ccontr) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1174 | assume "\<not> ?thesis" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1175 |     { fix x assume "x \<in> f"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1176 | 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 | 1177 | by (intro reduced_labelling_nonzero) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1178 | then have "reduced (Suc n) (lab x) \<noteq> n" | 
| 60420 | 1179 | 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 | 1180 | 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 | 1181 | 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 | 1182 | using eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1183 | ultimately show False | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1184 | by force | 
| 53186 | 1185 | qed | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1186 | 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 | 1187 | using j eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1188 | ultimately show "x n = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1189 | using j x by auto | 
| 53688 | 1190 | 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 | 1191 | |
| 60420 | 1192 | 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 | 1193 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1194 | lemma kuhn_induction: | 
| 53688 | 1195 | assumes "0 < p" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1196 | and lab_0: "\<forall>x. \<forall>j\<le>n. (\<forall>j. x j \<le> p) \<and> x j = 0 \<longrightarrow> lab x j = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1197 | and lab_1: "\<forall>x. \<forall>j\<le>n. (\<forall>j. x j \<le> p) \<and> x j = p \<longrightarrow> lab x j = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1198 |     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 | 1199 |   shows "odd (card {s. ksimplex p (Suc n) s \<and> (reduced (Suc n)\<circ>lab) ` s = {..Suc n}})"
 | 
| 53248 | 1200 | proof - | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1201 | 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 | 1202 | let ?ext = "\<lambda>s. (\<exists>j\<le>n. \<forall>x\<in>s. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>s. x j = p)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1203 |   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 | 1204 | by (simp add: reduced_labelling subset_eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1205 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1206 |   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 | 1207 |         {f. \<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> f = s - {a} \<and> ?rl ` f = {..n} \<and> ?ext f}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1208 | proof (intro set_eqI, safe del: disjCI equalityI disjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1209 |     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 | 1210 | 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 | 1211 | then interpret kuhn_simplex p n u b s . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1212 | 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 | 1213 | by (auto simp: out_eq_p) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1214 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1215 |     { fix x assume "x \<in> s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1216 | 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 | 1217 | have "?rl x \<le> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1218 | by (auto intro!: reduced_labelling_nonzero) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1219 | then have "?rl x = reduced n (lab x)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1220 | by (auto intro!: reduced_labelling_Suc) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1221 |     then have "?rl ` s = {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1222 | using rl by (simp cong: image_cong) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1223 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1224 |     obtain t a where "ksimplex p (Suc n) t" "a \<in> t" "s = t - {a}"
 | 
| 60420 | 1225 | 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 | 1226 | ultimately | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1227 |     show "\<exists>t a. ksimplex p (Suc n) t \<and> a \<in> t \<and> s = t - {a} \<and> ?rl ` s = {..n} \<and> ?ext s"
 | 
| 53688 | 1228 | by auto | 
| 53248 | 1229 | next | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1230 |     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 | 1231 |       and a: "a \<in> s" and "?ext (s - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1232 | 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 | 1233 | 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 | 1234 | 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 | 1235 | by (auto simp: out_eq_p) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1236 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1237 |     { fix x assume "x \<in> s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1238 |       then have "?rl x \<in> ?rl ` (s - {a})"
 | 
| 53248 | 1239 | by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1240 | then have "?rl x \<le> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1241 | unfolding rl by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1242 | then have "?rl x = reduced n (lab x)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1243 | by (auto intro!: reduced_labelling_Suc) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1244 |     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 | 1245 | unfolding rl[symmetric] by (intro image_cong) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1246 | |
| 60420 | 1247 |     from \<open>?ext (s - {a})\<close>
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1248 |     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 | 1249 | proof (elim disjE exE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1250 |       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 | 1251 | 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 | 1252 |       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 | 1253 | by (intro reduced_labelling_zero) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1254 |       moreover have "j \<in> ?rl ` (s - {a})"
 | 
| 60420 | 1255 | 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 | 1256 | ultimately show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1257 | by force | 
| 53248 | 1258 | next | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1259 |       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 | 1260 | show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1261 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1262 | 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 | 1263 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1264 | assume "j \<noteq> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1265 |         { fix x assume x: "x \<in> s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1266 | have "reduced n (lab x) \<le> j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1267 | proof (rule reduced_labelling_nonzero) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1268 | show "lab x j \<noteq> 0" | 
| 60420 | 1269 | 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 | 1270 | show "j < n" | 
| 60420 | 1271 | 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 | 1272 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1273 | then have "reduced n (lab x) \<noteq> n" | 
| 60420 | 1274 | 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 | 1275 |         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 | 1276 | unfolding rl' by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1277 | ultimately show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1278 | by force | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1279 | qed | 
| 53248 | 1280 | qed | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1281 |     show "ksimplex p n (s - {a})"
 | 
| 60420 | 1282 | unfolding simplex_top_face[OF \<open>0 < p\<close> all_eq_p] using s a by auto | 
| 53248 | 1283 | qed | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1284 | ultimately show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1285 | using assms by (intro kuhn_simplex_lemma) auto | 
| 53248 | 1286 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1287 | |
| 60420 | 1288 | 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 | 1289 | |
| 53688 | 1290 | lemma ksimplex_0: "ksimplex p 0 s \<longleftrightarrow> s = {(\<lambda>x. p)}"
 | 
| 53248 | 1291 | proof | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1292 |   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 | 1293 | by (blast dest: kuhn_simplex.ksimplex_0 elim: ksimplex.cases) | 
| 53248 | 1294 | next | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1295 |   assume s: "s = {(\<lambda>x. p)}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1296 | show "ksimplex p 0 s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1297 | proof (intro ksimplex, unfold_locales) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1298 |     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 | 1299 |     show "bij_betw id {..<0} {..<0}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1300 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1301 | qed (auto simp: s) | 
| 53248 | 1302 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1303 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1304 | lemma kuhn_combinatorial: | 
| 53688 | 1305 | assumes "0 < p" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1306 | 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 | 1307 | 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 | 1308 |   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 | 1309 | (is "odd (card (?M n))") | 
| 53248 | 1310 | using assms | 
| 1311 | proof (induct n) | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1312 | case 0 then show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1313 | by (simp add: ksimplex_0 cong: conj_cong) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1314 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1315 | case (Suc n) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1316 | then have "odd (card (?M n))" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1317 | by force | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1318 | with Suc show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1319 | using kuhn_induction[of p n] by (auto simp: comp_def) | 
| 53248 | 1320 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1321 | |
| 53248 | 1322 | lemma kuhn_lemma: | 
| 53688 | 1323 | fixes n p :: nat | 
| 1324 | assumes "0 < p" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1325 | and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. label x i = (0::nat) \<or> label x i = 1)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1326 | 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 | 1327 | 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 | 1328 | obtains q where "\<forall>i<n. q i < p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1329 | and "\<forall>i<n. \<exists>r s. (\<forall>j<n. q j \<le> r j \<and> r j \<le> q j + 1) \<and> (\<forall>j<n. q j \<le> s j \<and> s j \<le> q j + 1) \<and> label r i \<noteq> label s i" | 
| 53248 | 1330 | proof - | 
| 60580 | 1331 | let ?rl = "reduced n \<circ> label" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1332 |   let ?A = "{s. ksimplex p n s \<and> ?rl ` s = {..n}}"
 | 
| 53248 | 1333 | have "odd (card ?A)" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1334 | using assms by (intro kuhn_combinatorial[of p n label]) auto | 
| 53688 | 1335 |   then have "?A \<noteq> {}"
 | 
| 60580 | 1336 | by fastforce | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1337 |   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 | 1338 | by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1339 | 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 | 1340 | |
| 53248 | 1341 | show ?thesis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1342 | proof (intro that[of b] allI impI) | 
| 60580 | 1343 | fix i | 
| 1344 | assume "i < n" | |
| 1345 | then show "b i < p" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1346 | using base by auto | 
| 53248 | 1347 | next | 
| 60580 | 1348 | fix i | 
| 1349 | assume "i < n" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1350 |     then have "i \<in> {.. n}" "Suc i \<in> {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1351 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1352 | 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 | 1353 | unfolding rl[symmetric] by blast | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1354 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1355 | have "label u i \<noteq> label v i" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1356 | using reduced_labelling [of n "label u"] reduced_labelling [of n "label v"] | 
| 60420 | 1357 | 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 | 1358 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1359 | moreover | 
| 60580 | 1360 | have "b j \<le> u j" "u j \<le> b j + 1" "b j \<le> v j" "v j \<le> b j + 1" if "j < n" for j | 
| 1361 | using that base_le[OF \<open>u\<in>s\<close>] le_Suc_base[OF \<open>u\<in>s\<close>] base_le[OF \<open>v\<in>s\<close>] le_Suc_base[OF \<open>v\<in>s\<close>] | |
| 1362 | by auto | |
| 1363 | ultimately show "\<exists>r s. (\<forall>j<n. b j \<le> r j \<and> r j \<le> b j + 1) \<and> | |
| 1364 | (\<forall>j<n. b j \<le> s j \<and> s j \<le> b j + 1) \<and> label r i \<noteq> label s i" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1365 | by blast | 
| 53248 | 1366 | qed | 
| 1367 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1368 | |
| 60420 | 1369 | subsection \<open>The 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 | 1370 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1371 | lemma kuhn_labelling_lemma': | 
| 53688 | 1372 | assumes "(\<forall>x::nat\<Rightarrow>real. P x \<longrightarrow> P (f x))" | 
| 1373 | and "\<forall>x. P x \<longrightarrow> (\<forall>i::nat. Q i \<longrightarrow> 0 \<le> x i \<and> x i \<le> 1)" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1374 | shows "\<exists>l. (\<forall>x i. l x i \<le> (1::nat)) \<and> | 
| 53688 | 1375 | (\<forall>x i. P x \<and> Q i \<and> x i = 0 \<longrightarrow> l x i = 0) \<and> | 
| 1376 | (\<forall>x i. P x \<and> Q i \<and> x i = 1 \<longrightarrow> l x i = 1) \<and> | |
| 1377 | (\<forall>x i. P x \<and> Q i \<and> l x i = 0 \<longrightarrow> x i \<le> f x i) \<and> | |
| 1378 | (\<forall>x i. P x \<and> Q i \<and> l x i = 1 \<longrightarrow> f x i \<le> x i)" | |
| 53185 | 1379 | proof - | 
| 53688 | 1380 | have and_forall_thm: "\<And>P Q. (\<forall>x. P x) \<and> (\<forall>x. Q x) \<longleftrightarrow> (\<forall>x. P x \<and> Q x)" | 
| 1381 | by auto | |
| 1382 | have *: "\<forall>x y::real. 0 \<le> x \<and> x \<le> 1 \<and> 0 \<le> y \<and> y \<le> 1 \<longrightarrow> x \<noteq> 1 \<and> x \<le> y \<or> x \<noteq> 0 \<and> y \<le> x" | |
| 53185 | 1383 | by auto | 
| 1384 | show ?thesis | |
| 1385 | unfolding and_forall_thm | |
| 1386 | apply (subst choice_iff[symmetric])+ | |
| 53688 | 1387 | apply rule | 
| 1388 | apply rule | |
| 1389 | proof - | |
| 60580 | 1390 | fix x x' | 
| 53688 | 1391 | let ?R = "\<lambda>y::nat. | 
| 60580 | 1392 | (P x \<and> Q x' \<and> x x' = 0 \<longrightarrow> y = 0) \<and> | 
| 1393 | (P x \<and> Q x' \<and> x x' = 1 \<longrightarrow> y = 1) \<and> | |
| 1394 | (P x \<and> Q x' \<and> y = 0 \<longrightarrow> x x' \<le> (f x) x') \<and> | |
| 1395 | (P x \<and> Q x' \<and> y = 1 \<longrightarrow> (f x) x' \<le> x x')" | |
| 1396 | have "0 \<le> f x x' \<and> f x x' \<le> 1" if "P x" "Q x'" | |
| 1397 | using assms(2)[rule_format,of "f x" x'] that | |
| 1398 | apply (drule_tac assms(1)[rule_format]) | |
| 1399 | apply auto | |
| 1400 | done | |
| 53688 | 1401 | then have "?R 0 \<or> ?R 1" | 
| 1402 | by auto | |
| 60580 | 1403 | then show "\<exists>y\<le>1. ?R y" | 
| 53688 | 1404 | by auto | 
| 53185 | 1405 | qed | 
| 1406 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1407 | |
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1408 | definition unit_cube :: "'a::euclidean_space set" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1409 |   where "unit_cube = {x. \<forall>i\<in>Basis. 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1}"
 | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1410 | |
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1411 | lemma mem_unit_cube: "x \<in> unit_cube \<longleftrightarrow> (\<forall>i\<in>Basis. 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1412 | unfolding unit_cube_def by simp | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1413 | |
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1414 | lemma bounded_unit_cube: "bounded unit_cube" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1415 | unfolding bounded_def | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1416 | proof (intro exI ballI) | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1417 | fix y :: 'a assume y: "y \<in> unit_cube" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1418 | have "dist 0 y = norm y" by (rule dist_0_norm) | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1419 | also have "\<dots> = norm (\<Sum>i\<in>Basis. (y \<bullet> i) *\<^sub>R i)" unfolding euclidean_representation .. | 
| 64267 | 1420 | also have "\<dots> \<le> (\<Sum>i\<in>Basis. norm ((y \<bullet> i) *\<^sub>R i))" by (rule norm_sum) | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1421 | also have "\<dots> \<le> (\<Sum>i::'a\<in>Basis. 1)" | 
| 64267 | 1422 | by (rule sum_mono, simp add: y [unfolded mem_unit_cube]) | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1423 | finally show "dist 0 y \<le> (\<Sum>i::'a\<in>Basis. 1)" . | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1424 | qed | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1425 | |
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1426 | lemma closed_unit_cube: "closed unit_cube" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1427 | unfolding unit_cube_def Collect_ball_eq Collect_conj_eq | 
| 63332 | 1428 | by (rule closed_INT, auto intro!: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1429 | |
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1430 | lemma compact_unit_cube: "compact unit_cube" (is "compact ?C") | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1431 | unfolding compact_eq_seq_compact_metric | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1432 | using bounded_unit_cube closed_unit_cube | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1433 | by (rule bounded_closed_imp_seq_compact) | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1434 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1435 | lemma brouwer_cube: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1436 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1437 | assumes "continuous_on unit_cube f" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1438 | and "f ` unit_cube \<subseteq> unit_cube" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1439 | shows "\<exists>x\<in>unit_cube. f x = x" | 
| 53185 | 1440 | proof (rule ccontr) | 
| 63040 | 1441 |   define n where "n = DIM('a)"
 | 
| 53185 | 1442 | have n: "1 \<le> n" "0 < n" "n \<noteq> 0" | 
| 1443 | unfolding n_def by (auto simp add: Suc_le_eq DIM_positive) | |
| 53674 | 1444 | assume "\<not> ?thesis" | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1445 | then have *: "\<not> (\<exists>x\<in>unit_cube. f x - x = 0)" | 
| 53674 | 1446 | by auto | 
| 55522 | 1447 | obtain d where | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1448 | d: "d > 0" "\<And>x. x \<in> unit_cube \<Longrightarrow> d \<le> norm (f x - x)" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1449 | apply (rule brouwer_compactness_lemma[OF compact_unit_cube _ *]) | 
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56273diff
changeset | 1450 | apply (rule continuous_intros assms)+ | 
| 55522 | 1451 | apply blast | 
| 53185 | 1452 | done | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1453 | have *: "\<forall>x. x \<in> unit_cube \<longrightarrow> f x \<in> unit_cube" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1454 | "\<forall>x. x \<in> (unit_cube::'a set) \<longrightarrow> (\<forall>i\<in>Basis. True \<longrightarrow> 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)" | 
| 53185 | 1455 | using assms(2)[unfolded image_subset_iff Ball_def] | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1456 | unfolding mem_unit_cube | 
| 55522 | 1457 | by auto | 
| 1458 | obtain label :: "'a \<Rightarrow> 'a \<Rightarrow> nat" where | |
| 1459 | "\<forall>x. \<forall>i\<in>Basis. label x i \<le> 1" | |
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1460 | "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> x \<bullet> i = 0 \<longrightarrow> label x i = 0" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1461 | "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> x \<bullet> i = 1 \<longrightarrow> label x i = 1" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1462 | "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> label x i = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1463 | "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> label x i = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i" | 
| 55522 | 1464 | using kuhn_labelling_lemma[OF *] by blast | 
| 53185 | 1465 | note label = this [rule_format] | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1466 | have lem1: "\<forall>x\<in>unit_cube. \<forall>y\<in>unit_cube. \<forall>i\<in>Basis. label x i \<noteq> label y i \<longrightarrow> | 
| 61945 | 1467 | \<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)" | 
| 53185 | 1468 | proof safe | 
| 1469 | fix x y :: 'a | |
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1470 | assume x: "x \<in> unit_cube" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1471 | assume y: "y \<in> unit_cube" | 
| 53185 | 1472 | fix i | 
| 1473 | assume i: "label x i \<noteq> label y i" "i \<in> Basis" | |
| 1474 | have *: "\<And>x y fx fy :: real. x \<le> fx \<and> fy \<le> y \<or> fx \<le> x \<and> y \<le> fy \<Longrightarrow> | |
| 61945 | 1475 | \<bar>fx - x\<bar> \<le> \<bar>fy - fx\<bar> + \<bar>y - x\<bar>" by auto | 
| 1476 | have "\<bar>(f x - x) \<bullet> i\<bar> \<le> \<bar>(f y - f x)\<bullet>i\<bar> + \<bar>(y - x)\<bullet>i\<bar>" | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1477 | unfolding inner_simps | 
| 53185 | 1478 | apply (rule *) | 
| 1479 | apply (cases "label x i = 0") | |
| 53688 | 1480 | apply (rule disjI1) | 
| 1481 | apply rule | |
| 53185 | 1482 | prefer 3 | 
| 53688 | 1483 | apply (rule disjI2) | 
| 1484 | apply rule | |
| 1485 | proof - | |
| 53185 | 1486 | assume lx: "label x i = 0" | 
| 53674 | 1487 | then have ly: "label y i = 1" | 
| 53688 | 1488 | using i label(1)[of i y] | 
| 1489 | by auto | |
| 53185 | 1490 | show "x \<bullet> i \<le> f x \<bullet> i" | 
| 1491 | apply (rule label(4)[rule_format]) | |
| 53674 | 1492 | using x y lx i(2) | 
| 53252 | 1493 | apply auto | 
| 53185 | 1494 | done | 
| 1495 | show "f y \<bullet> i \<le> y \<bullet> i" | |
| 1496 | apply (rule label(5)[rule_format]) | |
| 53674 | 1497 | using x y ly i(2) | 
| 53252 | 1498 | apply auto | 
| 53185 | 1499 | done | 
| 1500 | next | |
| 1501 | assume "label x i \<noteq> 0" | |
| 53688 | 1502 | then have l: "label x i = 1" "label y i = 0" | 
| 1503 | using i label(1)[of i x] label(1)[of i y] | |
| 1504 | by auto | |
| 53185 | 1505 | show "f x \<bullet> i \<le> x \<bullet> i" | 
| 1506 | apply (rule label(5)[rule_format]) | |
| 53674 | 1507 | using x y l i(2) | 
| 53252 | 1508 | apply auto | 
| 53185 | 1509 | done | 
| 1510 | show "y \<bullet> i \<le> f y \<bullet> i" | |
| 1511 | apply (rule label(4)[rule_format]) | |
| 53674 | 1512 | using x y l i(2) | 
| 53252 | 1513 | apply auto | 
| 53185 | 1514 | done | 
| 1515 | qed | |
| 1516 | also have "\<dots> \<le> norm (f y - f x) + norm (y - x)" | |
| 1517 | apply (rule add_mono) | |
| 1518 | apply (rule Basis_le_norm[OF i(2)])+ | |
| 1519 | done | |
| 1520 | finally show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)" | |
| 1521 | unfolding inner_simps . | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1522 | qed | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1523 | have "\<exists>e>0. \<forall>x\<in>unit_cube. \<forall>y\<in>unit_cube. \<forall>z\<in>unit_cube. \<forall>i\<in>Basis. | 
| 53688 | 1524 | norm (x - z) < e \<and> norm (y - z) < e \<and> label x i \<noteq> label y i \<longrightarrow> | 
| 61945 | 1525 | \<bar>(f(z) - z)\<bullet>i\<bar> < d / (real n)" | 
| 53185 | 1526 | proof - | 
| 53688 | 1527 | have d': "d / real n / 8 > 0" | 
| 56541 | 1528 | using d(1) by (simp add: n_def DIM_positive) | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1529 | have *: "uniformly_continuous_on unit_cube f" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1530 | by (rule compact_uniformly_continuous[OF assms(1) compact_unit_cube]) | 
| 55522 | 1531 | obtain e where e: | 
| 1532 | "e > 0" | |
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1533 | "\<And>x x'. x \<in> unit_cube \<Longrightarrow> | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1534 | x' \<in> unit_cube \<Longrightarrow> | 
| 55522 | 1535 | norm (x' - x) < e \<Longrightarrow> | 
| 1536 | norm (f x' - f x) < d / real n / 8" | |
| 1537 | using *[unfolded uniformly_continuous_on_def,rule_format,OF d'] | |
| 1538 | unfolding dist_norm | |
| 1539 | by blast | |
| 53185 | 1540 | show ?thesis | 
| 1541 | apply (rule_tac x="min (e/2) (d/real n/8)" in exI) | |
| 53248 | 1542 | apply safe | 
| 1543 | proof - | |
| 53185 | 1544 | show "0 < min (e / 2) (d / real n / 8)" | 
| 1545 | using d' e by auto | |
| 1546 | fix x y z i | |
| 53688 | 1547 | assume as: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1548 | "x \<in> unit_cube" "y \<in> unit_cube" "z \<in> unit_cube" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36587diff
changeset | 1549 | "norm (x - z) < min (e / 2) (d / real n / 8)" | 
| 53688 | 1550 | "norm (y - z) < min (e / 2) (d / real n / 8)" | 
| 1551 | "label x i \<noteq> label y i" | |
| 1552 | assume i: "i \<in> Basis" | |
| 61945 | 1553 | have *: "\<And>z fz x fx n1 n2 n3 n4 d4 d :: real. \<bar>fx - x\<bar> \<le> n1 + n2 \<Longrightarrow> | 
| 1554 | \<bar>fx - fz\<bar> \<le> n3 \<Longrightarrow> \<bar>x - z\<bar> \<le> n4 \<Longrightarrow> | |
| 53185 | 1555 | n1 < d4 \<Longrightarrow> n2 < 2 * d4 \<Longrightarrow> n3 < d4 \<Longrightarrow> n4 < d4 \<Longrightarrow> | 
| 61945 | 1556 | (8 * d4 = d) \<Longrightarrow> \<bar>fz - z\<bar> < d" | 
| 53688 | 1557 | by auto | 
| 1558 | show "\<bar>(f z - z) \<bullet> i\<bar> < d / real n" | |
| 1559 | unfolding inner_simps | |
| 53185 | 1560 | proof (rule *) | 
| 1561 | show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y -f x) + norm (y - x)" | |
| 1562 | apply (rule lem1[rule_format]) | |
| 53688 | 1563 | using as i | 
| 1564 | apply auto | |
| 53185 | 1565 | done | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1566 | show "\<bar>f x \<bullet> i - f z \<bullet> i\<bar> \<le> norm (f x - f z)" "\<bar>x \<bullet> i - z \<bullet> i\<bar> \<le> norm (x - z)" | 
| 55522 | 1567 | unfolding inner_diff_left[symmetric] | 
| 53688 | 1568 | by (rule Basis_le_norm[OF i])+ | 
| 1569 | have tria: "norm (y - x) \<le> norm (y - z) + norm (x - z)" | |
| 53185 | 1570 | using dist_triangle[of y x z, unfolded dist_norm] | 
| 53688 | 1571 | unfolding norm_minus_commute | 
| 1572 | by auto | |
| 53185 | 1573 | also have "\<dots> < e / 2 + e / 2" | 
| 1574 | apply (rule add_strict_mono) | |
| 53252 | 1575 | using as(4,5) | 
| 1576 | apply auto | |
| 53185 | 1577 | done | 
| 1578 | finally show "norm (f y - f x) < d / real n / 8" | |
| 1579 | apply - | |
| 1580 | apply (rule e(2)) | |
| 53252 | 1581 | using as | 
| 1582 | apply auto | |
| 53185 | 1583 | done | 
| 1584 | have "norm (y - z) + norm (x - z) < d / real n / 8 + d / real n / 8" | |
| 1585 | apply (rule add_strict_mono) | |
| 53252 | 1586 | using as | 
| 1587 | apply auto | |
| 53185 | 1588 | done | 
| 53688 | 1589 | then show "norm (y - x) < 2 * (d / real n / 8)" | 
| 1590 | using tria | |
| 1591 | by auto | |
| 53185 | 1592 | show "norm (f x - f z) < d / real n / 8" | 
| 1593 | apply (rule e(2)) | |
| 53252 | 1594 | using as e(1) | 
| 1595 | apply auto | |
| 53185 | 1596 | done | 
| 1597 | qed (insert as, auto) | |
| 1598 | qed | |
| 1599 | qed | |
| 55522 | 1600 | then | 
| 1601 | obtain e where e: | |
| 1602 | "e > 0" | |
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1603 | "\<And>x y z i. x \<in> unit_cube \<Longrightarrow> | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1604 | y \<in> unit_cube \<Longrightarrow> | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1605 | z \<in> unit_cube \<Longrightarrow> | 
| 55522 | 1606 | i \<in> Basis \<Longrightarrow> | 
| 1607 | norm (x - z) < e \<and> norm (y - z) < e \<and> label x i \<noteq> label y i \<Longrightarrow> | |
| 1608 | \<bar>(f z - z) \<bullet> i\<bar> < d / real n" | |
| 1609 | by blast | |
| 1610 | obtain p :: nat where p: "1 + real n / e \<le> real p" | |
| 1611 | using real_arch_simple .. | |
| 53185 | 1612 | have "1 + real n / e > 0" | 
| 56541 | 1613 | using e(1) n by (simp add: add_pos_pos) | 
| 53688 | 1614 | then have "p > 0" | 
| 1615 | 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 | 1616 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1617 |   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 | 1618 | by atomize_elim (auto simp: n_def intro!: finite_same_card_bij) | 
| 63040 | 1619 |   define b' where "b' = inv_into {..< n} b"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1620 |   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 | 1621 | 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 | 1622 |   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 | 1623 | 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 | 1624 | have bb'[simp]:"\<And>i. i \<in> Basis \<Longrightarrow> b (b' i) = i" | 
| 53688 | 1625 | unfolding b'_def | 
| 1626 | using b | |
| 1627 | 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 | 1628 | have b'b[simp]:"\<And>i. i < n \<Longrightarrow> b' (b i) = i" | 
| 53688 | 1629 | unfolding b'_def | 
| 1630 | using b | |
| 1631 | by (auto simp: inv_into_f_eq bij_betw_def) | |
| 1632 | have *: "\<And>x :: nat. x = 0 \<or> x = 1 \<longleftrightarrow> x \<le> 1" | |
| 1633 | by auto | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1634 | have b'': "\<And>j. j < n \<Longrightarrow> b j \<in> Basis" | 
| 53185 | 1635 | using b unfolding bij_betw_def by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1636 | 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 | 1637 | (\<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 | 1638 | (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)" | 
| 53688 | 1639 | unfolding * | 
| 60420 | 1640 | using \<open>p > 0\<close> \<open>n > 0\<close> | 
| 53688 | 1641 | using label(1)[OF b''] | 
| 1642 | by auto | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1643 |   { fix x :: "nat \<Rightarrow> nat" and i assume "\<forall>i<n. x i \<le> p" "i < n" "x i = p \<or> x i = 0"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1644 | then have "(\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<in> (unit_cube::'a set)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1645 | using b'_Basis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1646 | by (auto simp add: mem_unit_cube inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1647 | note cube = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1648 | 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 | 1649 | (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0)" | 
| 60420 | 1650 | unfolding o_def using cube \<open>p > 0\<close> by (intro allI impI label(2)) (auto simp add: 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 | 1651 | 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 | 1652 | (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)" | 
| 60420 | 1653 | using cube \<open>p > 0\<close> unfolding o_def by (intro allI impI label(3)) (auto simp add: b'') | 
| 55522 | 1654 | obtain q where q: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1655 | "\<forall>i<n. q i < p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1656 | "\<forall>i<n. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1657 | \<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 | 1658 | (\<forall>j<n. q j \<le> s j \<and> s j \<le> q j + 1) \<and> | 
| 55522 | 1659 | (label (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i) \<circ> b) i \<noteq> | 
| 1660 | (label (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i) \<circ> b) i" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1661 | by (rule kuhn_lemma[OF q1 q2 q3]) | 
| 63040 | 1662 | define z :: 'a where "z = (\<Sum>i\<in>Basis. (real (q (b' i)) / real p) *\<^sub>R i)" | 
| 61945 | 1663 | have "\<exists>i\<in>Basis. d / real n \<le> \<bar>(f z - z)\<bullet>i\<bar>" | 
| 53185 | 1664 | 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 | 1665 |     have "\<forall>i\<in>Basis. q (b' i) \<in> {0..p}"
 | 
| 53688 | 1666 | using q(1) b' | 
| 1667 | by (auto intro: less_imp_le simp: bij_betw_def) | |
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1668 | then have "z \<in> unit_cube" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1669 | unfolding z_def mem_unit_cube | 
| 53688 | 1670 | using b'_Basis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1671 | by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1) | 
| 53688 | 1672 | then have d_fz_z: "d \<le> norm (f z - z)" | 
| 1673 | by (rule d) | |
| 1674 | assume "\<not> ?thesis" | |
| 53674 | 1675 | then have as: "\<forall>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar> < d / real n" | 
| 60420 | 1676 | using \<open>n > 0\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1677 | by (auto simp add: 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 | 1678 | have "norm (f z - z) \<le> (\<Sum>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar>)" | 
| 53688 | 1679 | unfolding inner_diff_left[symmetric] | 
| 1680 | by (rule norm_le_l1) | |
| 53185 | 1681 | also have "\<dots> < (\<Sum>(i::'a) \<in> Basis. d / real n)" | 
| 64267 | 1682 | apply (rule sum_strict_mono) | 
| 53688 | 1683 | using as | 
| 1684 | apply auto | |
| 53185 | 1685 | done | 
| 1686 | also have "\<dots> = d" | |
| 53688 | 1687 | using DIM_positive[where 'a='a] | 
| 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 | 1688 | by (auto simp: n_def) | 
| 53688 | 1689 | finally show False | 
| 1690 | using d_fz_z by auto | |
| 53185 | 1691 | qed | 
| 55522 | 1692 | 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 | 1693 | have *: "b' i < n" | 
| 55522 | 1694 | using i and b'[unfolded bij_betw_def] | 
| 53688 | 1695 | by auto | 
| 55522 | 1696 | obtain r s where rs: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1697 | "\<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 | 1698 | "\<And>j. j < n \<Longrightarrow> q j \<le> s j \<and> s j \<le> q j + 1" | 
| 55522 | 1699 | "(label (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i) \<circ> b) (b' i) \<noteq> | 
| 1700 | (label (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i) \<circ> b) (b' i)" | |
| 1701 | using q(2)[rule_format,OF *] by blast | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1702 | have b'_im: "\<And>i. i \<in> Basis \<Longrightarrow> b' i < n" | 
| 53185 | 1703 | using b' unfolding bij_betw_def by auto | 
| 63040 | 1704 | define r' ::'a where "r' = (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i)" | 
| 53185 | 1705 | have "\<And>i. i \<in> Basis \<Longrightarrow> r (b' i) \<le> p" | 
| 1706 | apply (rule order_trans) | |
| 1707 | apply (rule rs(1)[OF b'_im,THEN conjunct2]) | |
| 53252 | 1708 | using q(1)[rule_format,OF b'_im] | 
| 1709 | apply (auto simp add: Suc_le_eq) | |
| 53185 | 1710 | done | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1711 | then have "r' \<in> unit_cube" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1712 | unfolding r'_def mem_unit_cube | 
| 53688 | 1713 | using b'_Basis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1714 | by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1) | 
| 63040 | 1715 | define s' :: 'a where "s' = (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i)" | 
| 53688 | 1716 | have "\<And>i. i \<in> Basis \<Longrightarrow> s (b' i) \<le> p" | 
| 53185 | 1717 | apply (rule order_trans) | 
| 1718 | apply (rule rs(2)[OF b'_im, THEN conjunct2]) | |
| 53252 | 1719 | using q(1)[rule_format,OF b'_im] | 
| 1720 | apply (auto simp add: Suc_le_eq) | |
| 53185 | 1721 | done | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1722 | then have "s' \<in> unit_cube" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1723 | unfolding s'_def mem_unit_cube | 
| 53688 | 1724 | using b'_Basis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1725 | by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1) | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1726 | have "z \<in> unit_cube" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1727 | unfolding z_def mem_unit_cube | 
| 60420 | 1728 | using b'_Basis q(1)[rule_format,OF b'_im] \<open>p > 0\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1729 | by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le) | 
| 53688 | 1730 | have *: "\<And>x. 1 + real x = real (Suc x)" | 
| 1731 | by auto | |
| 1732 |   {
 | |
| 1733 | have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)" | |
| 64267 | 1734 | apply (rule sum_mono) | 
| 53252 | 1735 | using rs(1)[OF b'_im] | 
| 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 | 1736 | apply (auto simp add:* field_simps simp del: of_nat_Suc) | 
| 53185 | 1737 | done | 
| 53688 | 1738 | also have "\<dots> < e * real p" | 
| 60420 | 1739 | using p \<open>e > 0\<close> \<open>p > 0\<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 | 1740 | by (auto simp add: field_simps n_def) | 
| 53185 | 1741 | finally have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) < e * real p" . | 
| 1742 | } | |
| 1743 | moreover | |
| 53688 | 1744 |   {
 | 
| 1745 | have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)" | |
| 64267 | 1746 | apply (rule sum_mono) | 
| 53252 | 1747 | using rs(2)[OF b'_im] | 
| 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 | 1748 | apply (auto simp add:* field_simps simp del: of_nat_Suc) | 
| 53185 | 1749 | done | 
| 53688 | 1750 | also have "\<dots> < e * real p" | 
| 60420 | 1751 | using p \<open>e > 0\<close> \<open>p > 0\<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 | 1752 | by (auto simp add: field_simps n_def) | 
| 53185 | 1753 | finally have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) < e * real p" . | 
| 1754 | } | |
| 1755 | ultimately | |
| 53688 | 1756 | have "norm (r' - z) < e" and "norm (s' - z) < e" | 
| 53185 | 1757 | unfolding r'_def s'_def z_def | 
| 60420 | 1758 | using \<open>p > 0\<close> | 
| 53185 | 1759 | apply (rule_tac[!] le_less_trans[OF norm_le_l1]) | 
| 64267 | 1760 | apply (auto simp add: field_simps sum_divide_distrib[symmetric] inner_diff_left) | 
| 53185 | 1761 | done | 
| 53674 | 1762 | then have "\<bar>(f z - z) \<bullet> i\<bar> < d / real n" | 
| 53688 | 1763 | using rs(3) i | 
| 1764 | unfolding r'_def[symmetric] s'_def[symmetric] o_def bb' | |
| 60420 | 1765 | by (intro e(2)[OF \<open>r'\<in>unit_cube\<close> \<open>s'\<in>unit_cube\<close> \<open>z\<in>unit_cube\<close>]) auto | 
| 53688 | 1766 | then show False | 
| 1767 | 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 | 1768 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1769 | |
| 53185 | 1770 | |
| 60420 | 1771 | subsection \<open>Retractions\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1772 | |
| 53688 | 1773 | definition "retraction s t r \<longleftrightarrow> t \<subseteq> s \<and> continuous_on s r \<and> r ` s \<subseteq> t \<and> (\<forall>x\<in>t. r x = x)" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1774 | |
| 62843 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 1775 | definition retract_of (infixl "retract'_of" 50) | 
| 53185 | 1776 | where "(t retract_of s) \<longleftrightarrow> (\<exists>r. retraction s t r)" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1777 | |
| 53674 | 1778 | lemma retraction_idempotent: "retraction s t r \<Longrightarrow> x \<in> s \<Longrightarrow> r (r x) = r x" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1779 | unfolding retraction_def by auto | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1780 | |
| 60420 | 1781 | subsection \<open>Preservation of fixpoints under (more general notion of) retraction\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1782 | |
| 53185 | 1783 | lemma invertible_fixpoint_property: | 
| 53674 | 1784 | fixes s :: "'a::euclidean_space set" | 
| 1785 | and t :: "'b::euclidean_space set" | |
| 1786 | assumes "continuous_on t i" | |
| 1787 | and "i ` t \<subseteq> s" | |
| 53688 | 1788 | and "continuous_on s r" | 
| 1789 | and "r ` s \<subseteq> t" | |
| 53674 | 1790 | and "\<forall>y\<in>t. r (i y) = y" | 
| 1791 | and "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)" | |
| 1792 | and "continuous_on t g" | |
| 1793 | and "g ` t \<subseteq> t" | |
| 1794 | obtains y where "y \<in> t" and "g y = y" | |
| 53185 | 1795 | proof - | 
| 1796 | have "\<exists>x\<in>s. (i \<circ> g \<circ> r) x = x" | |
| 53688 | 1797 | apply (rule assms(6)[rule_format]) | 
| 1798 | apply rule | |
| 53185 | 1799 | apply (rule continuous_on_compose assms)+ | 
| 53688 | 1800 | apply ((rule continuous_on_subset)?, rule assms)+ | 
| 1801 | using assms(2,4,8) | |
| 53185 | 1802 | apply auto | 
| 1803 | apply blast | |
| 1804 | done | |
| 55522 | 1805 | then obtain x where x: "x \<in> s" "(i \<circ> g \<circ> r) x = x" .. | 
| 53674 | 1806 | then have *: "g (r x) \<in> t" | 
| 1807 | using assms(4,8) by auto | |
| 1808 | have "r ((i \<circ> g \<circ> r) x) = r x" | |
| 1809 | using x by auto | |
| 1810 | then show ?thesis | |
| 53185 | 1811 | apply (rule_tac that[of "r x"]) | 
| 53674 | 1812 | using x | 
| 1813 | unfolding o_def | |
| 1814 | unfolding assms(5)[rule_format,OF *] | |
| 1815 | using assms(4) | |
| 53185 | 1816 | apply auto | 
| 1817 | done | |
| 1818 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1819 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1820 | lemma homeomorphic_fixpoint_property: | 
| 53674 | 1821 | fixes s :: "'a::euclidean_space set" | 
| 1822 | and t :: "'b::euclidean_space set" | |
| 53185 | 1823 | assumes "s homeomorphic t" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1824 | shows "(\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)) \<longleftrightarrow> | 
| 53248 | 1825 | (\<forall>g. continuous_on t g \<and> g ` t \<subseteq> t \<longrightarrow> (\<exists>y\<in>t. g y = y))" | 
| 53185 | 1826 | proof - | 
| 55522 | 1827 | obtain r i where | 
| 1828 | "\<forall>x\<in>s. i (r x) = x" | |
| 1829 | "r ` s = t" | |
| 1830 | "continuous_on s r" | |
| 1831 | "\<forall>y\<in>t. r (i y) = y" | |
| 1832 | "i ` t = s" | |
| 1833 | "continuous_on t i" | |
| 1834 | using assms | |
| 1835 | unfolding homeomorphic_def homeomorphism_def | |
| 1836 | by blast | |
| 53674 | 1837 | then show ?thesis | 
| 53185 | 1838 | apply - | 
| 1839 | apply rule | |
| 1840 | apply (rule_tac[!] allI impI)+ | |
| 1841 | apply (rule_tac g=g in invertible_fixpoint_property[of t i s r]) | |
| 1842 | prefer 10 | |
| 1843 | apply (rule_tac g=f in invertible_fixpoint_property[of s r t i]) | |
| 1844 | apply auto | |
| 1845 | done | |
| 1846 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1847 | |
| 53185 | 1848 | lemma retract_fixpoint_property: | 
| 53688 | 1849 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 53674 | 1850 | and s :: "'a set" | 
| 53185 | 1851 | assumes "t retract_of s" | 
| 53674 | 1852 | and "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)" | 
| 1853 | and "continuous_on t g" | |
| 1854 | and "g ` t \<subseteq> t" | |
| 1855 | obtains y where "y \<in> t" and "g y = y" | |
| 53185 | 1856 | proof - | 
| 55522 | 1857 | obtain h where "retraction s t h" | 
| 1858 | using assms(1) unfolding retract_of_def .. | |
| 53674 | 1859 | then show ?thesis | 
| 53185 | 1860 | unfolding retraction_def | 
| 1861 | apply - | |
| 1862 | apply (rule invertible_fixpoint_property[OF continuous_on_id _ _ _ _ assms(2), of t h g]) | |
| 1863 | prefer 7 | |
| 53248 | 1864 | apply (rule_tac y = y in that) | 
| 1865 | using assms | |
| 1866 | apply auto | |
| 53185 | 1867 | done | 
| 1868 | qed | |
| 1869 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1870 | |
| 60420 | 1871 | subsection \<open>The Brouwer theorem for any set with nonempty interior\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1872 | |
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1873 | lemma convex_unit_cube: "convex unit_cube" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1874 | apply (rule is_interval_convex) | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1875 | apply (clarsimp simp add: is_interval_def mem_unit_cube) | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1876 | apply (drule (1) bspec)+ | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1877 | apply auto | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1878 | done | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1879 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1880 | lemma brouwer_weak: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1881 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | 
| 53674 | 1882 | assumes "compact s" | 
| 1883 | and "convex s" | |
| 1884 |     and "interior s \<noteq> {}"
 | |
| 1885 | and "continuous_on s f" | |
| 1886 | and "f ` s \<subseteq> s" | |
| 1887 | obtains x where "x \<in> s" and "f x = x" | |
| 53185 | 1888 | proof - | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1889 | let ?U = "unit_cube :: 'a set" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1890 | have "\<Sum>Basis /\<^sub>R 2 \<in> interior ?U" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1891 | proof (rule interiorI) | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1892 |     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 | 1893 | show "open ?I" | 
| 63332 | 1894 | by (intro open_INT finite_Basis ballI open_Int, auto intro: open_Collect_less simp: continuous_on_inner continuous_on_const continuous_on_id) | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1895 | show "\<Sum>Basis /\<^sub>R 2 \<in> ?I" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1896 | by simp | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1897 | show "?I \<subseteq> unit_cube" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1898 | unfolding unit_cube_def by force | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1899 | qed | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1900 |   then have *: "interior ?U \<noteq> {}" by fast
 | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1901 | have *: "?U homeomorphic s" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1902 | using homeomorphic_convex_compact[OF convex_unit_cube compact_unit_cube * assms(2,1,3)] . | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1903 | 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 | 1904 | (\<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 | 1905 | using brouwer_cube by auto | 
| 53674 | 1906 | then show ?thesis | 
| 53185 | 1907 | unfolding homeomorphic_fixpoint_property[OF *] | 
| 53252 | 1908 | using assms | 
| 59765 
26d1c71784f1
tweaked a few slow or very ugly proofs
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 1909 | by (auto simp: intro: that) | 
| 53185 | 1910 | qed | 
| 1911 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1912 | |
| 60420 | 1913 | text \<open>And in particular for a closed ball.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1914 | |
| 53185 | 1915 | lemma brouwer_ball: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1916 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | 
| 53674 | 1917 | assumes "e > 0" | 
| 1918 | and "continuous_on (cball a e) f" | |
| 53688 | 1919 | and "f ` cball a e \<subseteq> cball a e" | 
| 53674 | 1920 | obtains x where "x \<in> cball a e" and "f x = x" | 
| 53185 | 1921 | using brouwer_weak[OF compact_cball convex_cball, of a e f] | 
| 1922 | 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 | 1923 | 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 | 1924 | |
| 60420 | 1925 | text \<open>Still more general form; could derive this directly without using the | 
| 61808 | 1926 | rather involved \<open>HOMEOMORPHIC_CONVEX_COMPACT\<close> theorem, just using | 
| 60420 | 1927 | a scaling and translation to put the set inside 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 | 1928 | |
| 53248 | 1929 | lemma brouwer: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1930 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | 
| 53674 | 1931 | assumes "compact s" | 
| 1932 | and "convex s" | |
| 1933 |     and "s \<noteq> {}"
 | |
| 1934 | and "continuous_on s f" | |
| 1935 | and "f ` s \<subseteq> s" | |
| 1936 | obtains x where "x \<in> s" and "f x = x" | |
| 53185 | 1937 | proof - | 
| 1938 | have "\<exists>e>0. s \<subseteq> cball 0 e" | |
| 53688 | 1939 | using compact_imp_bounded[OF assms(1)] | 
| 1940 | unfolding bounded_pos | |
| 53674 | 1941 | apply (erule_tac exE) | 
| 1942 | apply (rule_tac x=b in exI) | |
| 53185 | 1943 | apply (auto simp add: dist_norm) | 
| 1944 | done | |
| 55522 | 1945 | then obtain e where e: "e > 0" "s \<subseteq> cball 0 e" | 
| 1946 | by blast | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1947 | have "\<exists>x\<in> cball 0 e. (f \<circ> closest_point s) x = x" | 
| 53185 | 1948 | apply (rule_tac brouwer_ball[OF e(1), of 0 "f \<circ> closest_point s"]) | 
| 1949 | apply (rule continuous_on_compose ) | |
| 1950 | apply (rule continuous_on_closest_point[OF assms(2) compact_imp_closed[OF assms(1)] assms(3)]) | |
| 1951 | apply (rule continuous_on_subset[OF assms(4)]) | |
| 1952 | apply (insert closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)]) | |
| 1953 | using assms(5)[unfolded subset_eq] | |
| 1954 | using e(2)[unfolded subset_eq mem_cball] | |
| 1955 | apply (auto simp add: dist_norm) | |
| 1956 | done | |
| 55522 | 1957 | then obtain x where x: "x \<in> cball 0 e" "(f \<circ> closest_point s) x = x" .. | 
| 53185 | 1958 | have *: "closest_point s x = x" | 
| 1959 | apply (rule closest_point_self) | |
| 1960 | apply (rule assms(5)[unfolded subset_eq,THEN bspec[where x="x"], unfolded image_iff]) | |
| 1961 | apply (rule_tac x="closest_point s x" in bexI) | |
| 1962 | using x | |
| 1963 | unfolding o_def | |
| 1964 | using closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3), of x] | |
| 1965 | apply auto | |
| 1966 | done | |
| 1967 | show thesis | |
| 1968 | apply (rule_tac x="closest_point s x" in that) | |
| 1969 | unfolding x(2)[unfolded o_def] | |
| 1970 | apply (rule closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)]) | |
| 53674 | 1971 | using * | 
| 1972 | apply auto | |
| 1973 | done | |
| 53185 | 1974 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1975 | |
| 60420 | 1976 | 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 | 1977 | |
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1978 | theorem no_retraction_cball: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1979 | fixes a :: "'a::euclidean_space" | 
| 53674 | 1980 | assumes "e > 0" | 
| 1981 | shows "\<not> (frontier (cball a e) retract_of (cball a e))" | |
| 53185 | 1982 | proof | 
| 60580 | 1983 | assume *: "frontier (cball a e) retract_of (cball a e)" | 
| 1984 | have **: "\<And>xa. a - (2 *\<^sub>R a - xa) = - (a - xa)" | |
| 53185 | 1985 | using scaleR_left_distrib[of 1 1 a] by auto | 
| 55522 | 1986 | obtain x where x: | 
| 1987 |       "x \<in> {x. norm (a - x) = e}"
 | |
| 1988 | "2 *\<^sub>R a - x = x" | |
| 60580 | 1989 | apply (rule retract_fixpoint_property[OF *, of "\<lambda>x. scaleR 2 a - x"]) | 
| 59765 
26d1c71784f1
tweaked a few slow or very ugly proofs
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 1990 | apply (blast intro: brouwer_ball[OF assms]) | 
| 
26d1c71784f1
tweaked a few slow or very ugly proofs
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 1991 | apply (intro continuous_intros) | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62061diff
changeset | 1992 | unfolding frontier_cball subset_eq Ball_def image_iff dist_norm sphere_def | 
| 60580 | 1993 | apply (auto simp add: ** norm_minus_commute) | 
| 53185 | 1994 | done | 
| 53674 | 1995 | then have "scaleR 2 a = scaleR 1 x + scaleR 1 x" | 
| 53248 | 1996 | by (auto simp add: algebra_simps) | 
| 53674 | 1997 | then have "a = x" | 
| 53688 | 1998 | unfolding scaleR_left_distrib[symmetric] | 
| 1999 | by auto | |
| 53674 | 2000 | then show False | 
| 2001 | using x assms by auto | |
| 53185 | 2002 | qed | 
| 2003 | ||
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2004 | corollary contractible_sphere: | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2005 | fixes a :: "'a::euclidean_space" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2006 | shows "contractible(sphere a r) \<longleftrightarrow> r \<le> 0" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2007 | proof (cases "0 < r") | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2008 | case True | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2009 | then show ?thesis | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2010 | unfolding contractible_def nullhomotopic_from_sphere_extension | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2011 | using no_retraction_cball [OF True, of a] | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2012 | by (auto simp: retract_of_def retraction_def) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2013 | next | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2014 | case False | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2015 | then show ?thesis | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2016 | unfolding contractible_def nullhomotopic_from_sphere_extension | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2017 | apply (simp add: not_less) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2018 | apply (rule_tac x=id in exI) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2019 | apply (auto simp: continuous_on_def) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2020 | apply (meson dist_not_less_zero le_less less_le_trans) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2021 | done | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2022 | qed | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 2023 | |
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2024 | lemma connected_sphere_eq: | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2025 | fixes a :: "'a :: euclidean_space" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2026 |   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 | 2027 | (is "?lhs = ?rhs") | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2028 | proof (cases r "0::real" rule: linorder_cases) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2029 | case less | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2030 | then show ?thesis by auto | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2031 | next | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2032 | case equal | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2033 | then show ?thesis by auto | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2034 | next | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2035 | case greater | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2036 | show ?thesis | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2037 | proof | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2038 | assume L: ?lhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2039 |     have "False" if 1: "DIM('a) = 1"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2040 | proof - | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2041 |       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 | 2042 | using sphere_1D_doubleton [OF 1 greater] | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2043 | 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 | 2044 | then have "finite (sphere a r)" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2045 | by auto | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2046 | with L \<open>r > 0\<close> show "False" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2047 | apply (auto simp: connected_finite_iff_sing) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2048 | using xy by auto | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2049 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2050 | with greater show ?rhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2051 | 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 | 2052 | next | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2053 | assume ?rhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2054 | then show ?lhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2055 | using connected_sphere greater by auto | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2056 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2057 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2058 | |
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2059 | lemma path_connected_sphere_eq: | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2060 | fixes a :: "'a :: euclidean_space" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2061 |   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 | 2062 | (is "?lhs = ?rhs") | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2063 | proof | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2064 | assume ?lhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2065 | then show ?rhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2066 | 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 | 2067 | next | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2068 | assume R: ?rhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2069 | then show ?lhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2070 | 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 | 2071 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2072 | |
| 64122 | 2073 | proposition frontier_subset_retraction: | 
| 2074 | fixes S :: "'a::euclidean_space set" | |
| 2075 | assumes "bounded S" and fros: "frontier S \<subseteq> T" | |
| 2076 | and contf: "continuous_on (closure S) f" | |
| 2077 | and fim: "f ` S \<subseteq> T" | |
| 2078 | and fid: "\<And>x. x \<in> T \<Longrightarrow> f x = x" | |
| 2079 | shows "S \<subseteq> T" | |
| 2080 | proof (rule ccontr) | |
| 2081 | assume "\<not> S \<subseteq> T" | |
| 2082 | then obtain a where "a \<in> S" "a \<notin> T" by blast | |
| 2083 | define g where "g \<equiv> \<lambda>z. if z \<in> closure S then f z else z" | |
| 2084 | have "continuous_on (closure S \<union> closure(-S)) g" | |
| 2085 | unfolding g_def | |
| 2086 | apply (rule continuous_on_cases) | |
| 2087 | using fros fid frontier_closures | |
| 2088 | apply (auto simp: contf continuous_on_id) | |
| 2089 | done | |
| 2090 | moreover have "closure S \<union> closure(- S) = UNIV" | |
| 2091 | using closure_Un by fastforce | |
| 2092 | ultimately have contg: "continuous_on UNIV g" by metis | |
| 2093 | obtain B where "0 < B" and B: "closure S \<subseteq> ball a B" | |
| 2094 | using \<open>bounded S\<close> bounded_subset_ballD by blast | |
| 2095 | have notga: "g x \<noteq> a" for x | |
| 2096 | unfolding g_def using fros fim \<open>a \<notin> T\<close> | |
| 2097 | apply (auto simp: frontier_def) | |
| 2098 | using fid interior_subset apply fastforce | |
| 2099 | by (simp add: \<open>a \<in> S\<close> closure_def) | |
| 2100 | define h where "h \<equiv> (\<lambda>y. a + (B / norm(y - a)) *\<^sub>R (y - a)) \<circ> g" | |
| 2101 | have "\<not> (frontier (cball a B) retract_of (cball a B))" | |
| 2102 | by (metis no_retraction_cball \<open>0 < B\<close>) | |
| 2103 | then have "\<And>k. \<not> retraction (cball a B) (frontier (cball a B)) k" | |
| 2104 | by (simp add: retract_of_def) | |
| 2105 | moreover have "retraction (cball a B) (frontier (cball a B)) h" | |
| 2106 | unfolding retraction_def | |
| 2107 | proof (intro conjI ballI) | |
| 2108 | show "frontier (cball a B) \<subseteq> cball a B" | |
| 2109 | by (force simp:) | |
| 2110 | show "continuous_on (cball a B) h" | |
| 2111 | unfolding h_def | |
| 2112 | apply (intro continuous_intros) | |
| 2113 | using contg continuous_on_subset notga apply auto | |
| 2114 | done | |
| 2115 | show "h ` cball a B \<subseteq> frontier (cball a B)" | |
| 2116 | using \<open>0 < B\<close> by (auto simp: h_def notga dist_norm) | |
| 2117 | show "\<And>x. x \<in> frontier (cball a B) \<Longrightarrow> h x = x" | |
| 2118 | apply (auto simp: h_def algebra_simps) | |
| 2119 | apply (simp add: vector_add_divide_simps notga) | |
| 2120 | 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) | |
| 2121 | qed | |
| 2122 | ultimately show False by simp | |
| 2123 | qed | |
| 2124 | ||
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2125 | subsection\<open>More Properties of Retractions\<close> | 
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2126 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2127 | lemma retraction: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2128 | "retraction s t r \<longleftrightarrow> | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2129 | t \<subseteq> s \<and> continuous_on s r \<and> r ` s = t \<and> (\<forall>x \<in> t. r x = x)" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2130 | by (force simp: retraction_def) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2131 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2132 | lemma retract_of_imp_extensible: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2133 | assumes "s retract_of t" and "continuous_on s f" and "f ` s \<subseteq> u" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2134 | obtains g where "continuous_on t g" "g ` t \<subseteq> u" "\<And>x. x \<in> s \<Longrightarrow> g x = f x" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2135 | using assms | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2136 | apply (clarsimp simp add: retract_of_def retraction) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2137 | apply (rule_tac g = "f o r" in that) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2138 | apply (auto simp: continuous_on_compose2) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2139 | done | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2140 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2141 | lemma idempotent_imp_retraction: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2142 | assumes "continuous_on s f" and "f ` s \<subseteq> s" and "\<And>x. x \<in> s \<Longrightarrow> f(f x) = f x" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2143 | shows "retraction s (f ` s) f" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2144 | by (simp add: assms retraction) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2145 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2146 | lemma retraction_subset: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2147 | assumes "retraction s t r" and "t \<subseteq> s'" and "s' \<subseteq> s" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2148 | shows "retraction s' t r" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2149 | apply (simp add: retraction_def) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2150 | by (metis assms continuous_on_subset image_mono retraction) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2151 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2152 | lemma retract_of_subset: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2153 | assumes "t retract_of s" and "t \<subseteq> s'" and "s' \<subseteq> s" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2154 | shows "t retract_of s'" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2155 | by (meson assms retract_of_def retraction_subset) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2156 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2157 | lemma retraction_refl [simp]: "retraction s s (\<lambda>x. x)" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2158 | by (simp add: continuous_on_id retraction) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2159 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2160 | lemma retract_of_refl [iff]: "s retract_of s" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2161 | using continuous_on_id retract_of_def retraction_def by fastforce | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2162 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2163 | lemma retract_of_imp_subset: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2164 | "s retract_of t \<Longrightarrow> s \<subseteq> t" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2165 | by (simp add: retract_of_def retraction_def) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2166 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2167 | lemma retract_of_empty [simp]: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2168 |      "({} retract_of s) \<longleftrightarrow> s = {}"  "(s retract_of {}) \<longleftrightarrow> s = {}"
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2169 | by (auto simp: retract_of_def retraction_def) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2170 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2171 | lemma retract_of_singleton [iff]: "({x} retract_of s) \<longleftrightarrow> x \<in> s"
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2172 | using continuous_on_const | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2173 | by (auto simp: retract_of_def retraction_def) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2174 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2175 | lemma retraction_comp: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2176 | "\<lbrakk>retraction s t f; retraction t u g\<rbrakk> | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2177 | \<Longrightarrow> retraction s u (g o f)" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2178 | apply (auto simp: retraction_def intro: continuous_on_compose2) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2179 | by blast | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2180 | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2181 | lemma retract_of_trans [trans]: | 
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2182 | assumes "s retract_of t" and "t retract_of u" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2183 | shows "s retract_of u" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2184 | using assms by (auto simp: retract_of_def intro: retraction_comp) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2185 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2186 | lemma closedin_retract: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2187 | fixes s :: "'a :: real_normed_vector set" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2188 | assumes "s retract_of t" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2189 | shows "closedin (subtopology euclidean t) s" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2190 | proof - | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2191 | obtain r where "s \<subseteq> t" "continuous_on t r" "r ` t \<subseteq> s" "\<And>x. x \<in> s \<Longrightarrow> r x = x" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2192 | using assms by (auto simp: retract_of_def retraction_def) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2193 |   then have s: "s = {x \<in> t. (norm(r x - x)) = 0}" by auto
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2194 | show ?thesis | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2195 | apply (subst s) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2196 | apply (rule continuous_closedin_preimage_constant) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2197 | by (simp add: \<open>continuous_on t r\<close> continuous_on_diff continuous_on_id continuous_on_norm) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2198 | qed | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2199 | |
| 63301 | 2200 | lemma closedin_self [simp]: | 
| 2201 | fixes S :: "'a :: real_normed_vector set" | |
| 2202 | shows "closedin (subtopology euclidean S) S" | |
| 2203 | by (simp add: closedin_retract) | |
| 2204 | ||
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2205 | lemma retract_of_contractible: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2206 | assumes "contractible t" "s retract_of t" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2207 | shows "contractible s" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2208 | using assms | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2209 | apply (clarsimp simp add: retract_of_def contractible_def retraction_def homotopic_with) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2210 | apply (rule_tac x="r a" in exI) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2211 | apply (rule_tac x="r o h" in exI) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2212 | apply (intro conjI continuous_intros continuous_on_compose) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2213 | apply (erule continuous_on_subset | force)+ | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2214 | done | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2215 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2216 | lemma retract_of_compact: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2217 | "\<lbrakk>compact t; s retract_of t\<rbrakk> \<Longrightarrow> compact s" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2218 | by (metis compact_continuous_image retract_of_def retraction) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2219 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2220 | lemma retract_of_closed: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2221 | fixes s :: "'a :: real_normed_vector set" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2222 | shows "\<lbrakk>closed t; s retract_of t\<rbrakk> \<Longrightarrow> closed s" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2223 | by (metis closedin_retract closedin_closed_eq) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2224 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2225 | lemma retract_of_connected: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2226 | "\<lbrakk>connected t; s retract_of t\<rbrakk> \<Longrightarrow> connected s" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2227 | by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2228 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2229 | lemma retract_of_path_connected: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2230 | "\<lbrakk>path_connected t; s retract_of t\<rbrakk> \<Longrightarrow> path_connected s" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2231 | by (metis path_connected_continuous_image retract_of_def retraction) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2232 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2233 | lemma retract_of_simply_connected: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2234 | "\<lbrakk>simply_connected t; s retract_of t\<rbrakk> \<Longrightarrow> simply_connected s" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2235 | apply (simp add: retract_of_def retraction_def, clarify) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2236 | apply (rule simply_connected_retraction_gen) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2237 | apply (force simp: continuous_on_id elim!: continuous_on_subset)+ | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2238 | done | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2239 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2240 | lemma retract_of_homotopically_trivial: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2241 | assumes ts: "t retract_of s" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2242 | and hom: "\<And>f g. \<lbrakk>continuous_on u f; f ` u \<subseteq> s; | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2243 | continuous_on u g; g ` u \<subseteq> s\<rbrakk> | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2244 | \<Longrightarrow> homotopic_with (\<lambda>x. True) u s f g" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2245 | and "continuous_on u f" "f ` u \<subseteq> t" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2246 | and "continuous_on u g" "g ` u \<subseteq> t" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2247 | shows "homotopic_with (\<lambda>x. True) u t f g" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2248 | proof - | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2249 | obtain r where "r ` s \<subseteq> s" "continuous_on s r" "\<forall>x\<in>s. r (r x) = r x" "t = r ` s" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2250 | using ts by (auto simp: retract_of_def retraction) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2251 | then obtain k where "Retracts s r t k" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2252 | unfolding Retracts_def | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2253 | by (metis continuous_on_subset dual_order.trans image_iff image_mono) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2254 | then show ?thesis | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2255 | apply (rule Retracts.homotopically_trivial_retraction_gen) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2256 | using assms | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2257 | apply (force simp: hom)+ | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2258 | done | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2259 | qed | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2260 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2261 | lemma retract_of_homotopically_trivial_null: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2262 | assumes ts: "t retract_of s" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2263 | and hom: "\<And>f. \<lbrakk>continuous_on u f; f ` u \<subseteq> s\<rbrakk> | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2264 | \<Longrightarrow> \<exists>c. homotopic_with (\<lambda>x. True) u s f (\<lambda>x. c)" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2265 | and "continuous_on u f" "f ` u \<subseteq> t" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2266 | obtains c where "homotopic_with (\<lambda>x. True) u t f (\<lambda>x. c)" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2267 | proof - | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2268 | obtain r where "r ` s \<subseteq> s" "continuous_on s r" "\<forall>x\<in>s. r (r x) = r x" "t = r ` s" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2269 | using ts by (auto simp: retract_of_def retraction) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2270 | then obtain k where "Retracts s r t k" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2271 | unfolding Retracts_def | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2272 | by (metis continuous_on_subset dual_order.trans image_iff image_mono) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2273 | then show ?thesis | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2274 | apply (rule Retracts.homotopically_trivial_retraction_null_gen) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2275 | apply (rule TrueI refl assms that | assumption)+ | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2276 | done | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2277 | qed | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2278 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2279 | lemma retraction_imp_quotient_map: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2280 | "retraction s t r | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2281 | \<Longrightarrow> u \<subseteq> t | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 2282 | \<Longrightarrow> (openin (subtopology euclidean s) (s \<inter> r -` u) \<longleftrightarrow> | 
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2283 | openin (subtopology euclidean t) u)" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2284 | apply (clarsimp simp add: retraction) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2285 | apply (rule continuous_right_inverse_imp_quotient_map [where g=r]) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2286 | apply (auto simp: elim: continuous_on_subset) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2287 | done | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2288 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2289 | lemma retract_of_locally_compact: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2290 |     fixes s :: "'a :: {heine_borel,real_normed_vector} set"
 | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2291 | shows "\<lbrakk> locally compact s; t retract_of s\<rbrakk> \<Longrightarrow> locally compact t" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2292 | by (metis locally_compact_closedin closedin_retract) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2293 | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 2294 | lemma retract_of_Times: | 
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2295 | "\<lbrakk>s retract_of s'; t retract_of t'\<rbrakk> \<Longrightarrow> (s \<times> t) retract_of (s' \<times> t')" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2296 | apply (simp add: retract_of_def retraction_def Sigma_mono, clarify) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2297 | apply (rename_tac f g) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2298 | apply (rule_tac x="\<lambda>z. ((f o fst) z, (g o snd) z)" in exI) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2299 | apply (rule conjI continuous_intros | erule continuous_on_subset | force)+ | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2300 | done | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2301 | |
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2302 | lemma homotopic_into_retract: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2303 | "\<lbrakk>f ` s \<subseteq> t; g ` s \<subseteq> t; t retract_of u; | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2304 | homotopic_with (\<lambda>x. True) s u f g\<rbrakk> | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2305 | \<Longrightarrow> homotopic_with (\<lambda>x. True) s t f g" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2306 | apply (subst (asm) homotopic_with_def) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2307 | apply (simp add: homotopic_with retract_of_def retraction_def, clarify) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2308 | apply (rule_tac x="r o h" in exI) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2309 | apply (rule conjI continuous_intros | erule continuous_on_subset | force simp: image_subset_iff)+ | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2310 | done | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62843diff
changeset | 2311 | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 2312 | lemma retract_of_locally_connected: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 2313 | assumes "locally connected T" "S retract_of T" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 2314 | shows "locally connected S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 2315 | using assms | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 2316 | by (auto simp: retract_of_def retraction intro!: retraction_imp_quotient_map elim!: locally_connected_quotient_image) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 2317 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 2318 | lemma retract_of_locally_path_connected: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 2319 | assumes "locally path_connected T" "S retract_of T" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 2320 | shows "locally path_connected S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 2321 | using assms | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 2322 | by (auto simp: retract_of_def retraction intro!: retraction_imp_quotient_map elim!: locally_path_connected_quotient_image) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 2323 | |
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2324 | subsubsection\<open>A few simple lemmas about deformation retracts\<close> | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2325 | |
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2326 | lemma deformation_retract_imp_homotopy_eqv: | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2327 | fixes S :: "'a::euclidean_space set" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2328 | assumes "homotopic_with (\<lambda>x. True) S S id r" "retraction S T r" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2329 | shows "S homotopy_eqv T" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2330 | apply (simp add: homotopy_eqv_def) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2331 | apply (rule_tac x=r in exI) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2332 | using assms apply (simp add: retraction_def) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2333 | apply (rule_tac x=id in exI) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2334 | apply (auto simp: continuous_on_id) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2335 | apply (metis homotopic_with_symD) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2336 | by (metis continuous_on_id' homotopic_with_equal homotopic_with_symD id_apply image_id subset_refl) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2337 | |
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2338 | lemma deformation_retract: | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2339 | fixes S :: "'a::euclidean_space set" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2340 | shows "(\<exists>r. homotopic_with (\<lambda>x. True) S S id r \<and> retraction S T r) \<longleftrightarrow> | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2341 | T retract_of S \<and> (\<exists>f. homotopic_with (\<lambda>x. True) S S id f \<and> f ` S \<subseteq> T)" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2342 | (is "?lhs = ?rhs") | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2343 | proof | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2344 | assume ?lhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2345 | then show ?rhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2346 | by (auto simp: retract_of_def retraction_def) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2347 | next | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2348 | assume ?rhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2349 | then show ?lhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2350 | apply (clarsimp simp add: retract_of_def retraction_def) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2351 | apply (rule_tac x=r in exI, simp) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2352 | apply (rule homotopic_with_trans, assumption) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2353 | apply (rule_tac f = "r \<circ> f" and g="r \<circ> id" in homotopic_with_eq) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2354 | apply (rule_tac Y=S in homotopic_compose_continuous_left) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2355 | apply (auto simp: homotopic_with_sym) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2356 | done | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2357 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2358 | |
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2359 | lemma deformation_retract_of_contractible_sing: | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2360 | fixes S :: "'a::euclidean_space set" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2361 | assumes "contractible S" "a \<in> S" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2362 |   obtains r where "homotopic_with (\<lambda>x. True) S S id r" "retraction S {a} r"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2363 | proof - | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2364 |   have "{a} retract_of S"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2365 | by (simp add: \<open>a \<in> S\<close>) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2366 | moreover have "homotopic_with (\<lambda>x. True) S S id (\<lambda>x. a)" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2367 | using assms | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2368 | apply (clarsimp simp add: contractible_def) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2369 | apply (rule homotopic_with_trans, assumption) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2370 | by (metis assms(1) contractible_imp_path_connected homotopic_constant_maps homotopic_with_sym homotopic_with_trans insert_absorb insert_not_empty path_component_mem(1) path_connected_component) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2371 |   moreover have "(\<lambda>x. a) ` S \<subseteq> {a}"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2372 | by (simp add: image_subsetI) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2373 | ultimately show ?thesis | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2374 | using that deformation_retract by metis | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2375 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2376 | |
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2377 | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2378 | subsection\<open>Punctured affine hulls, etc.\<close> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2379 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2380 | lemma continuous_on_compact_surface_projection_aux: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2381 | fixes S :: "'a::t2_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2382 | assumes "compact S" "S \<subseteq> T" "image q T \<subseteq> S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2383 | and contp: "continuous_on T p" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2384 | and "\<And>x. x \<in> S \<Longrightarrow> q x = x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2385 | and [simp]: "\<And>x. x \<in> T \<Longrightarrow> q(p x) = q x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2386 | and "\<And>x. x \<in> T \<Longrightarrow> p(q x) = p x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2387 | shows "continuous_on T q" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2388 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2389 | have *: "image p T = image p S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2390 | using assms by auto (metis imageI subset_iff) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2391 | have contp': "continuous_on S p" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2392 | by (rule continuous_on_subset [OF contp \<open>S \<subseteq> T\<close>]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2393 | have "continuous_on T (q \<circ> p)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2394 | apply (rule continuous_on_compose [OF contp]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2395 | apply (simp add: *) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2396 | apply (rule continuous_on_inv [OF contp' \<open>compact S\<close>]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2397 | using assms by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2398 | then show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2399 | apply (rule continuous_on_eq [of _ "q o p"]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2400 | apply (simp add: o_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2401 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2402 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2403 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2404 | lemma continuous_on_compact_surface_projection: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2405 | fixes S :: "'a::real_normed_vector set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2406 | assumes "compact S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2407 |       and S: "S \<subseteq> V - {0}" and "cone V"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2408 |       and iff: "\<And>x k. x \<in> V - {0} \<Longrightarrow> 0 < k \<and> (k *\<^sub>R x) \<in> S \<longleftrightarrow> d x = k"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2409 |   shows "continuous_on (V - {0}) (\<lambda>x. d x *\<^sub>R x)"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2410 | proof (rule continuous_on_compact_surface_projection_aux [OF \<open>compact S\<close> S]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2411 |   show "(\<lambda>x. d x *\<^sub>R x) ` (V - {0}) \<subseteq> S"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2412 | using iff by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2413 |   show "continuous_on (V - {0}) (\<lambda>x. inverse(norm x) *\<^sub>R x)"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2414 | by (intro continuous_intros) force | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2415 | show "\<And>x. x \<in> S \<Longrightarrow> d x *\<^sub>R x = x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2416 | by (metis S zero_less_one local.iff scaleR_one subset_eq) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2417 |   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
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2418 | using iff [of "inverse(norm x) *\<^sub>R x" "norm x * d x", symmetric] iff that \<open>cone V\<close> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2419 | by (simp add: field_simps cone_def zero_less_mult_iff) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2420 |   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
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2421 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2422 | have "0 < d x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2423 | using local.iff that by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2424 | then show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2425 | by simp | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2426 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2427 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2428 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2429 | proposition rel_frontier_deformation_retract_of_punctured_convex: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2430 | fixes S :: "'a::euclidean_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2431 | assumes "convex S" "convex T" "bounded S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2432 | and arelS: "a \<in> rel_interior S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2433 | and relS: "rel_frontier S \<subseteq> T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2434 | and affS: "T \<subseteq> affine hull S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2435 |   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 | 2436 |                   "retraction (T - {a}) (rel_frontier S) r"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2437 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2438 | 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 | 2439 | (\<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 | 2440 | 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 | 2441 | 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 | 2442 | apply (rule that)+ | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2443 | by metis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2444 | then obtain dd | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2445 | 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 | 2446 | 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 | 2447 | \<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 | 2448 | by metis+ | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2449 | have aaffS: "a \<in> affine hull S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2450 | 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 | 2451 |   have "((\<lambda>z. z - a) ` (affine hull S - {a})) = ((\<lambda>z. z - a) ` (affine hull S)) - {0}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2452 | by (auto simp: ) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2453 |   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 | 2454 | proof (rule continuous_on_compact_surface_projection) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2455 | 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 | 2456 | 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 | 2457 | 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 | 2458 | using rel_frontier_translation [of "-a"] add.commute by simp | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2459 |     also have "... \<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 | 2460 | 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 | 2461 |     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 | 2462 | show "cone ((\<lambda>z. z - a) ` (affine hull S))" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2463 | apply (rule subspace_imp_cone) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2464 | using aaffS | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2465 | apply (simp add: subspace_affine image_comp o_def affine_translation_aux [of a]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2466 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2467 | 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 | 2468 |          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 | 2469 | proof | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2470 | 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 | 2471 | 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 | 2472 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2473 | 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 | 2474 | have False if "dd x < k" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2475 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2476 | 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 | 2477 | using k closure_translation [of "-a"] | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2478 | by (auto simp: rel_frontier_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2479 | 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 | 2480 | 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 | 2481 | have "x \<noteq> 0" and xaffS: "a + x \<in> affine hull S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2482 | using x by (auto simp: ) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2483 | 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 | 2484 | using dd1 by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2485 | 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 | 2486 | 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 | 2487 | apply (simp add: in_segment) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2488 | 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 | 2489 | apply (simp add: field_simps that) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2490 | 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 | 2491 | 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 | 2492 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2493 | ultimately show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2494 | using segsub by (auto simp add: rel_frontier_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2495 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2496 | moreover have False if "k < dd x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2497 | using x k that rel_frontier_def | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2498 | by (fastforce simp: algebra_simps releq dest!: dd2) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2499 | ultimately show "dd x = k" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2500 | by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2501 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2502 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2503 |   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 | 2504 | by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2505 |   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 | 2506 | by (intro * continuous_intros continuous_on_compose) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2507 |   with affS have contdd: "continuous_on (T - {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 | 2508 | by (blast intro: continuous_on_subset elim: ) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2509 | show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2510 | proof | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2511 |     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 | 2512 | proof (rule homotopic_with_linear) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2513 |       show "continuous_on (T - {a}) id"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2514 | by (intro continuous_intros continuous_on_compose) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2515 |       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 | 2516 | using contdd by (simp add: o_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2517 |       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 | 2518 |            if "x \<in> T - {a}" for x
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2519 | proof (clarsimp simp: in_segment, intro conjI) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2520 | fix u::real assume u: "0 \<le> u" "u \<le> 1" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2521 | show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \<in> T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2522 | apply (rule convexD [OF \<open>convex T\<close>]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2523 | using that u apply (auto simp add: ) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2524 | apply (metis add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 relS subsetD) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2525 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2526 | 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 | 2527 | (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 | 2528 | by (auto simp: algebra_simps) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2529 | 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 | 2530 | 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 | 2531 | 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 | 2532 | 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 | 2533 | 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 | 2534 | 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 | 2535 | apply (auto simp: iff) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2536 | 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 | 2537 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2538 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2539 |     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 | 2540 | proof (simp add: retraction_def, intro conjI ballI) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2541 |       show "rel_frontier S \<subseteq> T - {a}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2542 | using arelS relS rel_frontier_def by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2543 |       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 | 2544 | using contdd by (simp add: o_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2545 |       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 | 2546 | apply (auto simp: rel_frontier_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2547 | 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 | 2548 | 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 | 2549 | 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 | 2550 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2551 | have "x \<noteq> a" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2552 | using that arelS by (auto simp add: rel_frontier_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2553 | have False if "dd (x - a) < 1" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2554 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2555 | have "x \<in> closure S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2556 | using x by (auto simp: rel_frontier_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2557 | 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 | 2558 | 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 | 2559 | have xaffS: "x \<in> affine hull S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2560 | using affS relS x by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2561 | then have "0 < dd (x - a)" and inS: "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 | 2562 | using dd1 by (auto simp add: \<open>x \<noteq> a\<close>) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2563 | 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 | 2564 | 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 | 2565 | apply (simp add: in_segment) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2566 | 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 | 2567 | apply (simp add: algebra_simps that) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2568 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2569 | ultimately show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2570 | using segsub by (auto simp add: rel_frontier_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2571 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2572 | moreover have False if "1 < dd (x - a)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2573 | 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 | 2574 | by (auto simp: rel_frontier_def) | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67399diff
changeset | 2575 | 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 | 2576 | by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2577 | with that show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2578 | by (simp add: rel_frontier_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2579 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2580 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2581 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2582 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2583 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2584 | corollary rel_frontier_retract_of_punctured_affine_hull: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2585 | fixes S :: "'a::euclidean_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2586 | 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 | 2587 |     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 | 2588 | apply (rule rel_frontier_deformation_retract_of_punctured_convex [of S "affine hull S" a]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2589 | apply (auto simp add: affine_imp_convex rel_frontier_affine_hull retract_of_def assms) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2590 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2591 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2592 | corollary rel_boundary_retract_of_punctured_affine_hull: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2593 | fixes S :: "'a::euclidean_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2594 | 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 | 2595 |     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 | 2596 | 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 | 2597 | rel_frontier_retract_of_punctured_affine_hull) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2598 | |
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2599 | lemma homotopy_eqv_rel_frontier_punctured_convex: | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2600 | fixes S :: "'a::euclidean_space set" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2601 | 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 | 2602 |   shows "(rel_frontier S) homotopy_eqv (T - {a})"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2603 | 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 | 2604 | using assms | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2605 | apply auto | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2606 | apply (subst homotopy_eqv_sym) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2607 | using deformation_retract_imp_homotopy_eqv by blast | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2608 | |
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2609 | lemma homotopy_eqv_rel_frontier_punctured_affine_hull: | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2610 | fixes S :: "'a::euclidean_space set" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2611 | 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 | 2612 |     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 | 2613 | apply (rule homotopy_eqv_rel_frontier_punctured_convex) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2614 | using assms rel_frontier_affine_hull by force+ | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2615 | |
| 64394 | 2616 | lemma path_connected_sphere_gen: | 
| 2617 | assumes "convex S" "bounded S" "aff_dim S \<noteq> 1" | |
| 2618 | shows "path_connected(rel_frontier S)" | |
| 2619 | proof (cases "rel_interior S = {}")
 | |
| 2620 | case True | |
| 2621 | then show ?thesis | |
| 2622 | by (simp add: \<open>convex S\<close> convex_imp_path_connected rel_frontier_def) | |
| 2623 | next | |
| 2624 | case False | |
| 2625 | then show ?thesis | |
| 2626 | 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) | |
| 2627 | qed | |
| 2628 | ||
| 2629 | lemma connected_sphere_gen: | |
| 2630 | assumes "convex S" "bounded S" "aff_dim S \<noteq> 1" | |
| 2631 | shows "connected(rel_frontier S)" | |
| 2632 | by (simp add: assms path_connected_imp_connected path_connected_sphere_gen) | |
| 2633 | ||
| 63301 | 2634 | subsection\<open>Borsuk-style characterization of separation\<close> | 
| 2635 | ||
| 2636 | lemma continuous_on_Borsuk_map: | |
| 2637 | "a \<notin> s \<Longrightarrow> continuous_on s (\<lambda>x. inverse(norm (x - a)) *\<^sub>R (x - a))" | |
| 2638 | by (rule continuous_intros | force)+ | |
| 2639 | ||
| 2640 | lemma Borsuk_map_into_sphere: | |
| 2641 | "(\<lambda>x. inverse(norm (x - a)) *\<^sub>R (x - a)) ` s \<subseteq> sphere 0 1 \<longleftrightarrow> (a \<notin> s)" | |
| 2642 | by auto (metis eq_iff_diff_eq_0 left_inverse norm_eq_zero) | |
| 2643 | ||
| 2644 | lemma Borsuk_maps_homotopic_in_path_component: | |
| 2645 | assumes "path_component (- s) a b" | |
| 2646 | shows "homotopic_with (\<lambda>x. True) s (sphere 0 1) | |
| 2647 | (\<lambda>x. inverse(norm(x - a)) *\<^sub>R (x - a)) | |
| 2648 | (\<lambda>x. inverse(norm(x - b)) *\<^sub>R (x - b))" | |
| 2649 | proof - | |
| 2650 | obtain g where "path g" "path_image g \<subseteq> -s" "pathstart g = a" "pathfinish g = b" | |
| 2651 | using assms by (auto simp: path_component_def) | |
| 2652 | then show ?thesis | |
| 2653 | apply (simp add: path_def path_image_def pathstart_def pathfinish_def homotopic_with_def) | |
| 2654 | apply (rule_tac x = "\<lambda>z. inverse(norm(snd z - (g o fst)z)) *\<^sub>R (snd z - (g o fst)z)" in exI) | |
| 2655 | apply (intro conjI continuous_intros) | |
| 2656 | apply (rule continuous_intros | erule continuous_on_subset | fastforce simp: divide_simps sphere_def)+ | |
| 2657 | done | |
| 2658 | qed | |
| 2659 | ||
| 2660 | lemma non_extensible_Borsuk_map: | |
| 2661 | fixes a :: "'a :: euclidean_space" | |
| 2662 | assumes "compact s" and cin: "c \<in> components(- s)" and boc: "bounded c" and "a \<in> c" | |
| 2663 | shows "~ (\<exists>g. continuous_on (s \<union> c) g \<and> | |
| 2664 | g ` (s \<union> c) \<subseteq> sphere 0 1 \<and> | |
| 2665 | (\<forall>x \<in> s. g x = inverse(norm(x - a)) *\<^sub>R (x - a)))" | |
| 2666 | proof - | |
| 2667 | have "closed s" using assms by (simp add: compact_imp_closed) | |
| 2668 | have "c \<subseteq> -s" | |
| 2669 | using assms by (simp add: in_components_subset) | |
| 2670 | with \<open>a \<in> c\<close> have "a \<notin> s" by blast | |
| 2671 | then have ceq: "c = connected_component_set (- s) a" | |
| 2672 | by (metis \<open>a \<in> c\<close> cin components_iff connected_component_eq) | |
| 2673 | then have "bounded (s \<union> connected_component_set (- s) a)" | |
| 2674 | using \<open>compact s\<close> boc compact_imp_bounded by auto | |
| 2675 | with bounded_subset_ballD obtain r where "0 < r" and r: "(s \<union> connected_component_set (- s) a) \<subseteq> ball a r" | |
| 2676 | by blast | |
| 2677 |   { fix g
 | |
| 2678 | assume "continuous_on (s \<union> c) g" | |
| 2679 | "g ` (s \<union> c) \<subseteq> sphere 0 1" | |
| 2680 | and [simp]: "\<And>x. x \<in> s \<Longrightarrow> g x = (x - a) /\<^sub>R norm (x - a)" | |
| 2681 | then have [simp]: "\<And>x. x \<in> s \<union> c \<Longrightarrow> norm (g x) = 1" | |
| 2682 | by force | |
| 2683 | have cb_eq: "cball a r = (s \<union> connected_component_set (- s) a) \<union> | |
| 2684 | (cball a r - connected_component_set (- s) a)" | |
| 2685 | using ball_subset_cball [of a r] r by auto | |
| 2686 | have cont1: "continuous_on (s \<union> connected_component_set (- s) a) | |
| 2687 | (\<lambda>x. a + r *\<^sub>R g x)" | |
| 2688 | apply (rule continuous_intros)+ | |
| 2689 | using \<open>continuous_on (s \<union> c) g\<close> ceq by blast | |
| 2690 | have cont2: "continuous_on (cball a r - connected_component_set (- s) a) | |
| 2691 | (\<lambda>x. a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" | |
| 2692 | by (rule continuous_intros | force simp: \<open>a \<notin> s\<close>)+ | |
| 2693 | have 1: "continuous_on (cball a r) | |
| 2694 | (\<lambda>x. if connected_component (- s) a x | |
| 2695 | then a + r *\<^sub>R g x | |
| 2696 | else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" | |
| 2697 | apply (subst cb_eq) | |
| 2698 | apply (rule continuous_on_cases [OF _ _ cont1 cont2]) | |
| 2699 | using ceq cin | |
| 2700 | apply (auto intro: closed_Un_complement_component | |
| 2701 | simp: \<open>closed s\<close> open_Compl open_connected_component) | |
| 2702 | done | |
| 2703 | have 2: "(\<lambda>x. a + r *\<^sub>R g x) ` (cball a r \<inter> connected_component_set (- s) a) | |
| 2704 | \<subseteq> sphere a r " | |
| 2705 | using \<open>0 < r\<close> by (force simp: dist_norm ceq) | |
| 2706 | have "retraction (cball a r) (sphere a r) | |
| 2707 | (\<lambda>x. if x \<in> connected_component_set (- s) a | |
| 2708 | then a + r *\<^sub>R g x | |
| 2709 | else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" | |
| 2710 | using \<open>0 < r\<close> | |
| 2711 | apply (simp add: retraction_def dist_norm 1 2, safe) | |
| 2712 | apply (force simp: dist_norm abs_if mult_less_0_iff divide_simps \<open>a \<notin> s\<close>) | |
| 2713 | using r | |
| 2714 | by (auto simp: dist_norm norm_minus_commute) | |
| 2715 | then have False | |
| 2716 | using no_retraction_cball | |
| 2717 | [OF \<open>0 < r\<close>, of a, unfolded retract_of_def, simplified, rule_format, | |
| 2718 | of "\<lambda>x. if x \<in> connected_component_set (- s) a | |
| 2719 | then a + r *\<^sub>R g x | |
| 2720 | else a + r *\<^sub>R inverse(norm(x - a)) *\<^sub>R (x - a)"] | |
| 2721 | by blast | |
| 2722 | } | |
| 2723 | then show ?thesis | |
| 2724 | by blast | |
| 2725 | qed | |
| 2726 | ||
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2727 | subsection\<open>Absolute retracts, Etc.\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2728 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2729 | text\<open>Absolute retracts (AR), absolute neighbourhood retracts (ANR) and also | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2730 | Euclidean neighbourhood retracts (ENR). We define AR and ANR by | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2731 | specializing the standard definitions for a set to embedding in | 
| 63306 
00090a0cd17f
Removed instances of ^ from theory markup
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 2732 | spaces of higher dimension. \<close> | 
| 
00090a0cd17f
Removed instances of ^ from theory markup
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 2733 | |
| 
00090a0cd17f
Removed instances of ^ from theory markup
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 2734 | (*This turns out to be sufficient (since any set in | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2735 | R^n can be embedded as a closed subset of a convex subset of R^{n+1}) to
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2736 | derive the usual definitions, but we need to split them into two | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2737 | implications because of the lack of type quantifiers. Then ENR turns out | 
| 63306 
00090a0cd17f
Removed instances of ^ from theory markup
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 2738 | to be equivalent to ANR plus local compactness. -- JRH*) | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2739 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2740 | definition AR :: "'a::topological_space set => bool" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2741 | where | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2742 |    "AR S \<equiv> \<forall>U. \<forall>S'::('a * real) set. S homeomorphic S' \<and> closedin (subtopology euclidean U) S'
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2743 | \<longrightarrow> S' retract_of U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2744 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2745 | definition ANR :: "'a::topological_space set => bool" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2746 | where | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2747 |    "ANR S \<equiv> \<forall>U. \<forall>S'::('a * real) set. S homeomorphic S' \<and> closedin (subtopology euclidean U) S'
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2748 | \<longrightarrow> (\<exists>T. openin (subtopology euclidean U) T \<and> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2749 | S' retract_of T)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2750 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2751 | definition ENR :: "'a::topological_space set => bool" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2752 | where "ENR S \<equiv> \<exists>U. open U \<and> S retract_of U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2753 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2754 | text\<open> First, show that we do indeed get the "usual" properties of ARs and ANRs.\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2755 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2756 | proposition AR_imp_absolute_extensor: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2757 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2758 | assumes "AR S" and contf: "continuous_on T f" and "f ` T \<subseteq> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2759 | and cloUT: "closedin (subtopology euclidean U) T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2760 | obtains g where "continuous_on U g" "g ` U \<subseteq> S" "\<And>x. x \<in> T \<Longrightarrow> g x = f x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2761 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2762 |   have "aff_dim S < int (DIM('b \<times> real))"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2763 | using aff_dim_le_DIM [of S] by simp | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2764 |   then obtain C and S' :: "('b * real) set"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2765 |           where C: "convex C" "C \<noteq> {}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2766 | and cloCS: "closedin (subtopology euclidean C) S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2767 | and hom: "S homeomorphic S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2768 | by (metis that homeomorphic_closedin_convex) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2769 | then have "S' retract_of C" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2770 | using \<open>AR S\<close> by (simp add: AR_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2771 | then obtain r where "S' \<subseteq> C" and contr: "continuous_on C r" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2772 | and "r ` C \<subseteq> S'" and rid: "\<And>x. x\<in>S' \<Longrightarrow> r x = x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2773 | by (auto simp: retraction_def retract_of_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2774 | obtain g h where "homeomorphism S S' g h" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2775 | using hom by (force simp: homeomorphic_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2776 | then have "continuous_on (f ` T) g" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2777 | by (meson \<open>f ` T \<subseteq> S\<close> continuous_on_subset homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2778 | then have contgf: "continuous_on T (g o f)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2779 | by (metis continuous_on_compose contf) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2780 | have gfTC: "(g \<circ> f) ` T \<subseteq> C" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2781 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2782 | have "g ` S = S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2783 | by (metis (no_types) \<open>homeomorphism S S' g h\<close> homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2784 | with \<open>S' \<subseteq> C\<close> \<open>f ` T \<subseteq> S\<close> show ?thesis by force | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2785 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2786 | obtain f' where f': "continuous_on U f'" "f' ` U \<subseteq> C" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2787 | "\<And>x. x \<in> T \<Longrightarrow> f' x = (g \<circ> f) x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2788 | by (metis Dugundji [OF C cloUT contgf gfTC]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2789 | show ?thesis | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2790 | proof (rule_tac g = "h o r o f'" in that) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2791 | show "continuous_on U (h \<circ> r \<circ> f')" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2792 | apply (intro continuous_on_compose f') | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2793 | using continuous_on_subset contr f' apply blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2794 | 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) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2795 | show "(h \<circ> r \<circ> f') ` U \<subseteq> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2796 | using \<open>homeomorphism S S' g h\<close> \<open>r ` C \<subseteq> S'\<close> \<open>f' ` U \<subseteq> C\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2797 | by (fastforce simp: homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2798 | show "\<And>x. x \<in> T \<Longrightarrow> (h \<circ> r \<circ> f') x = f x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2799 | using \<open>homeomorphism S S' g h\<close> \<open>f ` T \<subseteq> S\<close> f' | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2800 | by (auto simp: rid homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2801 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2802 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2803 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2804 | lemma AR_imp_absolute_retract: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2805 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2806 | assumes "AR S" "S homeomorphic S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2807 | and clo: "closedin (subtopology euclidean U) S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2808 | shows "S' retract_of U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2809 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2810 | obtain g h where hom: "homeomorphism S S' g h" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2811 | using assms by (force simp: homeomorphic_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2812 | have h: "continuous_on S' h" " h ` S' \<subseteq> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2813 | using hom homeomorphism_def apply blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2814 | apply (metis hom equalityE homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2815 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2816 | obtain h' where h': "continuous_on U h'" "h' ` U \<subseteq> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2817 | and h'h: "\<And>x. x \<in> S' \<Longrightarrow> h' x = h x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2818 | by (blast intro: AR_imp_absolute_extensor [OF \<open>AR S\<close> h clo]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2819 | have [simp]: "S' \<subseteq> U" using clo closedin_limpt by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2820 | show ?thesis | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2821 | proof (simp add: retraction_def retract_of_def, intro exI conjI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2822 | show "continuous_on U (g o h')" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2823 | apply (intro continuous_on_compose h') | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2824 | apply (meson hom continuous_on_subset h' homeomorphism_cont1) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2825 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2826 | show "(g \<circ> h') ` U \<subseteq> S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2827 | using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2828 | show "\<forall>x\<in>S'. (g \<circ> h') x = x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2829 | by clarsimp (metis h'h hom homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2830 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2831 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2832 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2833 | lemma AR_imp_absolute_retract_UNIV: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2834 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2835 | assumes "AR S" and hom: "S homeomorphic S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2836 | and clo: "closed S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2837 | shows "S' retract_of UNIV" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2838 | apply (rule AR_imp_absolute_retract [OF \<open>AR S\<close> hom]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2839 | using clo closed_closedin by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2840 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2841 | lemma absolute_extensor_imp_AR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2842 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2843 | assumes "\<And>f :: 'a * real \<Rightarrow> 'a. | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2844 | \<And>U T. \<lbrakk>continuous_on T f; f ` T \<subseteq> S; | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2845 | closedin (subtopology euclidean U) T\<rbrakk> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2846 | \<Longrightarrow> \<exists>g. continuous_on U g \<and> g ` U \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2847 | shows "AR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2848 | proof (clarsimp simp: AR_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2849 |   fix U and T :: "('a * real) set"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2850 | assume "S homeomorphic T" and clo: "closedin (subtopology euclidean U) T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2851 | then obtain g h where hom: "homeomorphism S T g h" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2852 | by (force simp: homeomorphic_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2853 | have h: "continuous_on T h" " h ` T \<subseteq> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2854 | using hom homeomorphism_def apply blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2855 | apply (metis hom equalityE homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2856 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2857 | obtain h' where h': "continuous_on U h'" "h' ` U \<subseteq> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2858 | and h'h: "\<forall>x\<in>T. h' x = h x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2859 | using assms [OF h clo] by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2860 | have [simp]: "T \<subseteq> U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2861 | using clo closedin_imp_subset by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2862 | show "T retract_of U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2863 | proof (simp add: retraction_def retract_of_def, intro exI conjI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2864 | show "continuous_on U (g o h')" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2865 | apply (intro continuous_on_compose h') | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2866 | apply (meson hom continuous_on_subset h' homeomorphism_cont1) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2867 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2868 | show "(g \<circ> h') ` U \<subseteq> T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2869 | using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2870 | show "\<forall>x\<in>T. (g \<circ> h') x = x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2871 | by clarsimp (metis h'h hom homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2872 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2873 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2874 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2875 | lemma AR_eq_absolute_extensor: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2876 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2877 | shows "AR S \<longleftrightarrow> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2878 | (\<forall>f :: 'a * real \<Rightarrow> 'a. | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2879 | \<forall>U T. continuous_on T f \<longrightarrow> f ` T \<subseteq> S \<longrightarrow> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2880 | closedin (subtopology euclidean U) T \<longrightarrow> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2881 | (\<exists>g. continuous_on U g \<and> g ` U \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)))" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2882 | apply (rule iffI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2883 | apply (metis AR_imp_absolute_extensor) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2884 | apply (simp add: absolute_extensor_imp_AR) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2885 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2886 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2887 | lemma AR_imp_retract: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2888 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2889 | assumes "AR S \<and> closedin (subtopology euclidean U) S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2890 | shows "S retract_of U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2891 | using AR_imp_absolute_retract assms homeomorphic_refl by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2892 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2893 | lemma AR_homeomorphic_AR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2894 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2895 | assumes "AR T" "S homeomorphic T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2896 | shows "AR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2897 | unfolding AR_def | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2898 | by (metis assms AR_imp_absolute_retract homeomorphic_trans [of _ S] homeomorphic_sym) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2899 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2900 | lemma homeomorphic_AR_iff_AR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2901 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2902 | shows "S homeomorphic T \<Longrightarrow> AR S \<longleftrightarrow> AR T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2903 | by (metis AR_homeomorphic_AR homeomorphic_sym) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2904 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2905 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2906 | proposition ANR_imp_absolute_neighbourhood_extensor: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2907 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2908 | assumes "ANR S" and contf: "continuous_on T f" and "f ` T \<subseteq> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2909 | and cloUT: "closedin (subtopology euclidean U) T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2910 | obtains V g where "T \<subseteq> V" "openin (subtopology euclidean U) V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2911 | "continuous_on V g" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2912 | "g ` V \<subseteq> S" "\<And>x. x \<in> T \<Longrightarrow> g x = f x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2913 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2914 |   have "aff_dim S < int (DIM('b \<times> real))"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2915 | using aff_dim_le_DIM [of S] by simp | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2916 |   then obtain C and S' :: "('b * real) set"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2917 |           where C: "convex C" "C \<noteq> {}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2918 | and cloCS: "closedin (subtopology euclidean C) S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2919 | and hom: "S homeomorphic S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2920 | by (metis that homeomorphic_closedin_convex) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2921 | then obtain D where opD: "openin (subtopology euclidean C) D" and "S' retract_of D" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2922 | using \<open>ANR S\<close> by (auto simp: ANR_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2923 | then obtain r where "S' \<subseteq> D" and contr: "continuous_on D r" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2924 | and "r ` D \<subseteq> S'" and rid: "\<And>x. x \<in> S' \<Longrightarrow> r x = x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2925 | by (auto simp: retraction_def retract_of_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2926 | obtain g h where homgh: "homeomorphism S S' g h" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2927 | using hom by (force simp: homeomorphic_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2928 | have "continuous_on (f ` T) g" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2929 | by (meson \<open>f ` T \<subseteq> S\<close> continuous_on_subset homeomorphism_def homgh) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2930 | then have contgf: "continuous_on T (g o f)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2931 | by (intro continuous_on_compose contf) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2932 | have gfTC: "(g \<circ> f) ` T \<subseteq> C" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2933 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2934 | have "g ` S = S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2935 | by (metis (no_types) homeomorphism_def homgh) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2936 | then show ?thesis | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2937 | by (metis (no_types) assms(3) cloCS closedin_def image_comp image_mono order.trans topspace_euclidean_subtopology) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2938 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2939 | obtain f' where contf': "continuous_on U f'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2940 | and "f' ` U \<subseteq> C" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2941 | and eq: "\<And>x. x \<in> T \<Longrightarrow> f' x = (g \<circ> f) x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2942 | by (metis Dugundji [OF C cloUT contgf gfTC]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2943 | show ?thesis | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 2944 | proof (rule_tac V = "U \<inter> f' -` D" and g = "h o r o f'" in that) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 2945 | show "T \<subseteq> U \<inter> f' -` D" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2946 | using cloUT closedin_imp_subset \<open>S' \<subseteq> D\<close> \<open>f ` T \<subseteq> S\<close> eq homeomorphism_image1 homgh | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2947 | by fastforce | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 2948 | show ope: "openin (subtopology euclidean U) (U \<inter> f' -` D)" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2949 | using \<open>f' ` U \<subseteq> C\<close> by (auto simp: opD contf' continuous_openin_preimage) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 2950 | have conth: "continuous_on (r ` f' ` (U \<inter> f' -` D)) h" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2951 | apply (rule continuous_on_subset [of S']) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2952 | using homeomorphism_def homgh apply blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2953 | using \<open>r ` D \<subseteq> S'\<close> by blast | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 2954 | show "continuous_on (U \<inter> f' -` D) (h \<circ> r \<circ> f')" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2955 | apply (intro continuous_on_compose conth | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2956 | continuous_on_subset [OF contr] continuous_on_subset [OF contf'], auto) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2957 | done | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 2958 | show "(h \<circ> r \<circ> f') ` (U \<inter> f' -` D) \<subseteq> S" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2959 | using \<open>homeomorphism S S' g h\<close> \<open>f' ` U \<subseteq> C\<close> \<open>r ` D \<subseteq> S'\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2960 | by (auto simp: homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2961 | show "\<And>x. x \<in> T \<Longrightarrow> (h \<circ> r \<circ> f') x = f x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2962 | using \<open>homeomorphism S S' g h\<close> \<open>f ` T \<subseteq> S\<close> eq | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2963 | by (auto simp: rid homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2964 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2965 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2966 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2967 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2968 | corollary ANR_imp_absolute_neighbourhood_retract: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2969 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2970 | assumes "ANR S" "S homeomorphic S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2971 | and clo: "closedin (subtopology euclidean U) S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2972 | obtains V where "openin (subtopology euclidean U) V" "S' retract_of V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2973 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2974 | obtain g h where hom: "homeomorphism S S' g h" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2975 | using assms by (force simp: homeomorphic_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2976 | have h: "continuous_on S' h" " h ` S' \<subseteq> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2977 | using hom homeomorphism_def apply blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2978 | apply (metis hom equalityE homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2979 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2980 | from ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> h clo] | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2981 | obtain V h' where "S' \<subseteq> V" and opUV: "openin (subtopology euclidean U) V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2982 | and h': "continuous_on V h'" "h' ` V \<subseteq> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2983 | and h'h:"\<And>x. x \<in> S' \<Longrightarrow> h' x = h x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2984 | by (blast intro: ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> h clo]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2985 | have "S' retract_of V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2986 | proof (simp add: retraction_def retract_of_def, intro exI conjI \<open>S' \<subseteq> V\<close>) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2987 | show "continuous_on V (g o h')" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2988 | apply (intro continuous_on_compose h') | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2989 | apply (meson hom continuous_on_subset h' homeomorphism_cont1) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2990 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2991 | show "(g \<circ> h') ` V \<subseteq> S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2992 | using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2993 | show "\<forall>x\<in>S'. (g \<circ> h') x = x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2994 | by clarsimp (metis h'h hom homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2995 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2996 | then show ?thesis | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2997 | by (rule that [OF opUV]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2998 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2999 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3000 | corollary ANR_imp_absolute_neighbourhood_retract_UNIV: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3001 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3002 | assumes "ANR S" and hom: "S homeomorphic S'" and clo: "closed S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3003 | obtains V where "open V" "S' retract_of V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3004 | using ANR_imp_absolute_neighbourhood_retract [OF \<open>ANR S\<close> hom] | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3005 | by (metis clo closed_closedin open_openin subtopology_UNIV) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3006 | |
| 63928 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3007 | corollary neighbourhood_extension_into_ANR: | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3008 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3009 | assumes contf: "continuous_on S f" and fim: "f ` S \<subseteq> T" and "ANR T" "closed S" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3010 | obtains V g where "S \<subseteq> V" "open V" "continuous_on V g" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3011 | "g ` V \<subseteq> T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3012 | using ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> contf fim] | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3013 | by (metis \<open>closed S\<close> closed_closedin open_openin subtopology_UNIV) | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3014 | |
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3015 | lemma absolute_neighbourhood_extensor_imp_ANR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3016 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3017 | assumes "\<And>f :: 'a * real \<Rightarrow> 'a. | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3018 | \<And>U T. \<lbrakk>continuous_on T f; f ` T \<subseteq> S; | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3019 | closedin (subtopology euclidean U) T\<rbrakk> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3020 | \<Longrightarrow> \<exists>V g. T \<subseteq> V \<and> openin (subtopology euclidean U) V \<and> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3021 | continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3022 | shows "ANR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3023 | proof (clarsimp simp: ANR_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3024 |   fix U and T :: "('a * real) set"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3025 | assume "S homeomorphic T" and clo: "closedin (subtopology euclidean U) T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3026 | then obtain g h where hom: "homeomorphism S T g h" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3027 | by (force simp: homeomorphic_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3028 | have h: "continuous_on T h" " h ` T \<subseteq> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3029 | using hom homeomorphism_def apply blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3030 | apply (metis hom equalityE homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3031 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3032 | obtain V h' where "T \<subseteq> V" and opV: "openin (subtopology euclidean U) V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3033 | and h': "continuous_on V h'" "h' ` V \<subseteq> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3034 | and h'h: "\<forall>x\<in>T. h' x = h x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3035 | using assms [OF h clo] by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3036 | have [simp]: "T \<subseteq> U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3037 | using clo closedin_imp_subset by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3038 | have "T retract_of V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3039 | proof (simp add: retraction_def retract_of_def, intro exI conjI \<open>T \<subseteq> V\<close>) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3040 | show "continuous_on V (g o h')" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3041 | apply (intro continuous_on_compose h') | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3042 | apply (meson hom continuous_on_subset h' homeomorphism_cont1) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3043 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3044 | show "(g \<circ> h') ` V \<subseteq> T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3045 | using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3046 | show "\<forall>x\<in>T. (g \<circ> h') x = x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3047 | by clarsimp (metis h'h hom homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3048 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3049 | then show "\<exists>V. openin (subtopology euclidean U) V \<and> T retract_of V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3050 | using opV by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3051 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3052 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3053 | lemma ANR_eq_absolute_neighbourhood_extensor: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3054 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3055 | shows "ANR S \<longleftrightarrow> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3056 | (\<forall>f :: 'a * real \<Rightarrow> 'a. | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3057 | \<forall>U T. continuous_on T f \<longrightarrow> f ` T \<subseteq> S \<longrightarrow> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3058 | closedin (subtopology euclidean U) T \<longrightarrow> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3059 | (\<exists>V g. T \<subseteq> V \<and> openin (subtopology euclidean U) V \<and> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3060 | continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x \<in> T. g x = f x)))" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3061 | apply (rule iffI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3062 | apply (metis ANR_imp_absolute_neighbourhood_extensor) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3063 | apply (simp add: absolute_neighbourhood_extensor_imp_ANR) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3064 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3065 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3066 | lemma ANR_imp_neighbourhood_retract: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3067 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3068 | assumes "ANR S" "closedin (subtopology euclidean U) S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3069 | obtains V where "openin (subtopology euclidean U) V" "S retract_of V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3070 | using ANR_imp_absolute_neighbourhood_retract assms homeomorphic_refl by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3071 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3072 | lemma ANR_imp_absolute_closed_neighbourhood_retract: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3073 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3074 | assumes "ANR S" "S homeomorphic S'" and US': "closedin (subtopology euclidean U) S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3075 | obtains V W | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3076 | where "openin (subtopology euclidean U) V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3077 | "closedin (subtopology euclidean U) W" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3078 | "S' \<subseteq> V" "V \<subseteq> W" "S' retract_of W" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3079 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3080 | obtain Z where "openin (subtopology euclidean U) Z" and S'Z: "S' retract_of Z" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3081 | by (blast intro: assms ANR_imp_absolute_neighbourhood_retract) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3082 | then have UUZ: "closedin (subtopology euclidean U) (U - Z)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3083 | by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3084 |   have "S' \<inter> (U - Z) = {}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3085 | using \<open>S' retract_of Z\<close> closedin_retract closedin_subtopology by fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3086 | then obtain V W | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3087 | where "openin (subtopology euclidean U) V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3088 | and "openin (subtopology euclidean U) W" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3089 |         and "S' \<subseteq> V" "U - Z \<subseteq> W" "V \<inter> W = {}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3090 | using separation_normal_local [OF US' UUZ] by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3091 | moreover have "S' retract_of U - W" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3092 | apply (rule retract_of_subset [OF S'Z]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3093 |     using US' \<open>S' \<subseteq> V\<close> \<open>V \<inter> W = {}\<close> closedin_subset apply fastforce
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3094 | using Diff_subset_conv \<open>U - Z \<subseteq> W\<close> by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3095 | ultimately show ?thesis | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3096 | apply (rule_tac V=V and W = "U-W" in that) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3097 | using openin_imp_subset apply (force simp:)+ | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3098 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3099 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3100 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3101 | lemma ANR_imp_closed_neighbourhood_retract: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3102 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3103 | assumes "ANR S" "closedin (subtopology euclidean U) S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3104 | obtains V W where "openin (subtopology euclidean U) V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3105 | "closedin (subtopology euclidean U) W" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3106 | "S \<subseteq> V" "V \<subseteq> W" "S retract_of W" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3107 | by (meson ANR_imp_absolute_closed_neighbourhood_retract assms homeomorphic_refl) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3108 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3109 | lemma ANR_homeomorphic_ANR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3110 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3111 | assumes "ANR T" "S homeomorphic T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3112 | shows "ANR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3113 | unfolding ANR_def | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3114 | by (metis assms ANR_imp_absolute_neighbourhood_retract homeomorphic_trans [of _ S] homeomorphic_sym) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3115 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3116 | lemma homeomorphic_ANR_iff_ANR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3117 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3118 | shows "S homeomorphic T \<Longrightarrow> ANR S \<longleftrightarrow> ANR T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3119 | by (metis ANR_homeomorphic_ANR homeomorphic_sym) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3120 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3121 | subsection\<open> Analogous properties of ENRs.\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3122 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3123 | proposition ENR_imp_absolute_neighbourhood_retract: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3124 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3125 | assumes "ENR S" and hom: "S homeomorphic S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3126 | and "S' \<subseteq> U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3127 | obtains V where "openin (subtopology euclidean U) V" "S' retract_of V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3128 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3129 | obtain X where "open X" "S retract_of X" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3130 | using \<open>ENR S\<close> by (auto simp: ENR_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3131 | then obtain r where "retraction X S r" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3132 | by (auto simp: retract_of_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3133 | have "locally compact S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3134 | using retract_of_locally_compact open_imp_locally_compact | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3135 | homeomorphic_local_compactness \<open>S retract_of X\<close> \<open>open X\<close> hom by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3136 | then obtain W where UW: "openin (subtopology euclidean U) W" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3137 | and WS': "closedin (subtopology euclidean W) S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3138 | apply (rule locally_compact_closedin_open) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3139 | apply (rename_tac W) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3140 | apply (rule_tac W = "U \<inter> W" in that, blast) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3141 | by (simp add: \<open>S' \<subseteq> U\<close> closedin_limpt) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3142 | obtain f g where hom: "homeomorphism S S' f g" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3143 | using assms by (force simp: homeomorphic_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3144 | have contg: "continuous_on S' g" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3145 | using hom homeomorphism_def by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3146 | moreover have "g ` S' \<subseteq> S" by (metis hom equalityE homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3147 | ultimately obtain h where conth: "continuous_on W h" and hg: "\<And>x. x \<in> S' \<Longrightarrow> h x = g x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3148 | using Tietze_unbounded [of S' g W] WS' by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3149 | have "W \<subseteq> U" using UW openin_open by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3150 | have "S' \<subseteq> W" using WS' closedin_closed by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3151 | have him: "\<And>x. x \<in> S' \<Longrightarrow> h x \<in> X" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3152 | 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) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3153 | have "S' retract_of (W \<inter> h -` X)" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3154 | proof (simp add: retraction_def retract_of_def, intro exI conjI) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3155 | show "S' \<subseteq> W" "S' \<subseteq> h -` X" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3156 | using him WS' closedin_imp_subset by blast+ | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3157 | show "continuous_on (W \<inter> h -` X) (f o r o h)" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3158 | proof (intro continuous_on_compose) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3159 | show "continuous_on (W \<inter> h -` X) h" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3160 | by (meson conth continuous_on_subset inf_le1) | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3161 | show "continuous_on (h ` (W \<inter> h -` X)) r" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3162 | proof - | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3163 | have "h ` (W \<inter> h -` X) \<subseteq> X" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3164 | by blast | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3165 | then show "continuous_on (h ` (W \<inter> h -` X)) r" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3166 | by (meson \<open>retraction X S r\<close> continuous_on_subset retraction) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3167 | qed | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3168 | show "continuous_on (r ` h ` (W \<inter> h -` X)) f" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3169 | apply (rule continuous_on_subset [of S]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3170 | using hom homeomorphism_def apply blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3171 | apply clarify | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3172 | apply (meson \<open>retraction X S r\<close> subsetD imageI retraction_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3173 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3174 | qed | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3175 | show "(f \<circ> r \<circ> h) ` (W \<inter> h -` X) \<subseteq> S'" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3176 | using \<open>retraction X S r\<close> hom | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3177 | by (auto simp: retraction_def homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3178 | show "\<forall>x\<in>S'. (f \<circ> r \<circ> h) x = x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3179 | using \<open>retraction X S r\<close> hom by (auto simp: retraction_def homeomorphism_def hg) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3180 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3181 | then show ?thesis | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3182 | apply (rule_tac V = "W \<inter> h -` X" in that) | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3183 | apply (rule openin_trans [OF _ UW]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3184 | using \<open>continuous_on W h\<close> \<open>open X\<close> continuous_openin_preimage_eq apply blast+ | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3185 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3186 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3187 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3188 | corollary ENR_imp_absolute_neighbourhood_retract_UNIV: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3189 | fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3190 | assumes "ENR S" "S homeomorphic S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3191 | obtains T' where "open T'" "S' retract_of T'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3192 | by (metis ENR_imp_absolute_neighbourhood_retract UNIV_I assms(1) assms(2) open_openin subsetI subtopology_UNIV) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3193 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3194 | lemma ENR_homeomorphic_ENR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3195 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3196 | assumes "ENR T" "S homeomorphic T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3197 | shows "ENR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3198 | unfolding ENR_def | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3199 | by (meson ENR_imp_absolute_neighbourhood_retract_UNIV assms homeomorphic_sym) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3200 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3201 | lemma homeomorphic_ENR_iff_ENR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3202 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3203 | assumes "S homeomorphic T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3204 | shows "ENR S \<longleftrightarrow> ENR T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3205 | by (meson ENR_homeomorphic_ENR assms homeomorphic_sym) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3206 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3207 | lemma ENR_translation: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3208 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3209 | shows "ENR(image (\<lambda>x. a + x) S) \<longleftrightarrow> ENR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3210 | by (meson homeomorphic_sym homeomorphic_translation homeomorphic_ENR_iff_ENR) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3211 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3212 | lemma ENR_linear_image_eq: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3213 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3214 | assumes "linear f" "inj f" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3215 | shows "ENR (image f S) \<longleftrightarrow> ENR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3216 | apply (rule homeomorphic_ENR_iff_ENR) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3217 | using assms homeomorphic_sym linear_homeomorphic_image by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3218 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3219 | subsection\<open>Some relations among the concepts\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3220 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3221 | text\<open>We also relate AR to being a retract of UNIV, which is often a more convenient proxy in the closed case.\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3222 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3223 | lemma AR_imp_ANR: "AR S \<Longrightarrow> ANR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3224 | using ANR_def AR_def by fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3225 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3226 | lemma ENR_imp_ANR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3227 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3228 | shows "ENR S \<Longrightarrow> ANR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3229 | apply (simp add: ANR_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3230 | by (metis ENR_imp_absolute_neighbourhood_retract closedin_imp_subset) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3231 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3232 | lemma ENR_ANR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3233 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3234 | shows "ENR S \<longleftrightarrow> ANR S \<and> locally compact S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3235 | proof | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3236 | assume "ENR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3237 | then have "locally compact S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3238 | using ENR_def open_imp_locally_compact retract_of_locally_compact by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3239 | then show "ANR S \<and> locally compact S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3240 | using ENR_imp_ANR \<open>ENR S\<close> by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3241 | next | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3242 | assume "ANR S \<and> locally compact S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3243 | then have "ANR S" "locally compact S" by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3244 |   then obtain T :: "('a * real) set" where "closed T" "S homeomorphic T"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3245 | using locally_compact_homeomorphic_closed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3246 | by (metis DIM_prod DIM_real Suc_eq_plus1 lessI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3247 | then show "ENR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3248 | using \<open>ANR S\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3249 | apply (simp add: ANR_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3250 | apply (drule_tac x=UNIV in spec) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3251 | apply (drule_tac x=T in spec, clarsimp) | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3252 | apply (meson ENR_def ENR_homeomorphic_ENR open_openin) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3253 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3254 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3255 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3256 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3257 | proposition AR_ANR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3258 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3259 |   shows "AR S \<longleftrightarrow> ANR S \<and> contractible S \<and> S \<noteq> {}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3260 | (is "?lhs = ?rhs") | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3261 | proof | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3262 | assume ?lhs | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3263 |   obtain C and S' :: "('a * real) set"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3264 |     where "convex C" "C \<noteq> {}" "closedin (subtopology euclidean C) S'" "S homeomorphic S'"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3265 | apply (rule homeomorphic_closedin_convex [of S, where 'n = "'a * real"]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3266 | using aff_dim_le_DIM [of S] by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3267 | with \<open>AR S\<close> have "contractible S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3268 | apply (simp add: AR_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3269 | apply (drule_tac x=C in spec) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3270 | apply (drule_tac x="S'" in spec, simp) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3271 | using convex_imp_contractible homeomorphic_contractible_eq retract_of_contractible by fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3272 | with \<open>AR S\<close> show ?rhs | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3273 | apply (auto simp: AR_imp_ANR) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3274 | apply (force simp: AR_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3275 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3276 | next | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3277 | assume ?rhs | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3278 | then obtain a and h:: "real \<times> 'a \<Rightarrow> 'a" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3279 |       where conth: "continuous_on ({0..1} \<times> S) h"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3280 |         and hS: "h ` ({0..1} \<times> S) \<subseteq> S"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3281 | and [simp]: "\<And>x. h(0, x) = x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3282 | and [simp]: "\<And>x. h(1, x) = a" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3283 |         and "ANR S" "S \<noteq> {}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3284 | by (auto simp: contractible_def homotopic_with_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3285 | then have "a \<in> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3286 | by (metis all_not_in_conv atLeastAtMost_iff image_subset_iff mem_Sigma_iff order_refl zero_le_one) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3287 | have "\<exists>g. continuous_on W g \<and> g ` W \<subseteq> S \<and> (\<forall>x\<in>T. g x = f x)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3288 | if f: "continuous_on T f" "f ` T \<subseteq> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3289 | and WT: "closedin (subtopology euclidean W) T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3290 | for W T and f :: "'a \<times> real \<Rightarrow> 'a" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3291 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3292 | obtain U g | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3293 | where "T \<subseteq> U" and WU: "openin (subtopology euclidean W) U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3294 | and contg: "continuous_on U g" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3295 | and "g ` U \<subseteq> S" and gf: "\<And>x. x \<in> T \<Longrightarrow> g x = f x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3296 | using iffD1 [OF ANR_eq_absolute_neighbourhood_extensor \<open>ANR S\<close>, rule_format, OF f WT] | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3297 | by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3298 | have WWU: "closedin (subtopology euclidean W) (W - U)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3299 | using WU closedin_diff by fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3300 |     moreover have "(W - U) \<inter> T = {}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3301 | using \<open>T \<subseteq> U\<close> by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3302 | ultimately obtain V V' | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3303 | where WV': "openin (subtopology euclidean W) V'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3304 | and WV: "openin (subtopology euclidean W) V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3305 |         and "W - U \<subseteq> V'" "T \<subseteq> V" "V' \<inter> V = {}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3306 | using separation_normal_local [of W "W-U" T] WT by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3307 |     then have WVT: "T \<inter> (W - V) = {}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3308 | by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3309 | have WWV: "closedin (subtopology euclidean W) (W - V)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3310 | using WV closedin_diff by fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3311 | obtain j :: " 'a \<times> real \<Rightarrow> real" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3312 | where contj: "continuous_on W j" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3313 |         and j:  "\<And>x. x \<in> W \<Longrightarrow> j x \<in> {0..1}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3314 | and j0: "\<And>x. x \<in> W - V \<Longrightarrow> j x = 1" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3315 | and j1: "\<And>x. x \<in> T \<Longrightarrow> j x = 0" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3316 | by (rule Urysohn_local [OF WT WWV WVT, of 0 "1::real"]) (auto simp: in_segment) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3317 | have Weq: "W = (W - V) \<union> (W - V')" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3318 |       using \<open>V' \<inter> V = {}\<close> by force
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3319 | show ?thesis | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3320 | proof (intro conjI exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3321 | have *: "continuous_on (W - V') (\<lambda>x. h (j x, g x))" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3322 | apply (rule continuous_on_compose2 [OF conth continuous_on_Pair]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3323 | apply (rule continuous_on_subset [OF contj Diff_subset]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3324 | apply (rule continuous_on_subset [OF contg]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3325 | apply (metis Diff_subset_conv Un_commute \<open>W - U \<subseteq> V'\<close>) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3326 | using j \<open>g ` U \<subseteq> S\<close> \<open>W - U \<subseteq> V'\<close> apply fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3327 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3328 | show "continuous_on W (\<lambda>x. if x \<in> W - V then a else h (j x, g x))" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3329 | apply (subst Weq) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3330 | apply (rule continuous_on_cases_local) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3331 | apply (simp_all add: Weq [symmetric] WWV continuous_on_const *) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3332 | using WV' closedin_diff apply fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3333 | apply (auto simp: j0 j1) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3334 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3335 | next | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3336 | have "h (j (x, y), g (x, y)) \<in> S" if "(x, y) \<in> W" "(x, y) \<in> V" for x y | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3337 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3338 |         have "j(x, y) \<in> {0..1}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3339 | using j that by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3340 | moreover have "g(x, y) \<in> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3341 |           using \<open>V' \<inter> V = {}\<close> \<open>W - U \<subseteq> V'\<close> \<open>g ` U \<subseteq> S\<close> that by fastforce
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3342 | ultimately show ?thesis | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3343 | using hS by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3344 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3345 | with \<open>a \<in> S\<close> \<open>g ` U \<subseteq> S\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3346 | show "(\<lambda>x. if x \<in> W - V then a else h (j x, g x)) ` W \<subseteq> S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3347 | by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3348 | next | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3349 | show "\<forall>x\<in>T. (if x \<in> W - V then a else h (j x, g x)) = f x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3350 | using \<open>T \<subseteq> V\<close> by (auto simp: j0 j1 gf) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3351 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3352 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3353 | then show ?lhs | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3354 | by (simp add: AR_eq_absolute_extensor) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3355 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3356 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3357 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3358 | lemma ANR_retract_of_ANR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3359 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3360 | assumes "ANR T" "S retract_of T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3361 | shows "ANR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3362 | using assms | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3363 | apply (simp add: ANR_eq_absolute_neighbourhood_extensor retract_of_def retraction_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3364 | apply (clarsimp elim!: all_forward) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3365 | apply (erule impCE, metis subset_trans) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3366 | apply (clarsimp elim!: ex_forward) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3367 | apply (rule_tac x="r o g" in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3368 | by (metis comp_apply continuous_on_compose continuous_on_subset subsetD imageI image_comp image_mono subset_trans) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3369 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3370 | lemma AR_retract_of_AR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3371 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3372 | shows "\<lbrakk>AR T; S retract_of T\<rbrakk> \<Longrightarrow> AR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3373 | using ANR_retract_of_ANR AR_ANR retract_of_contractible by fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3374 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3375 | lemma ENR_retract_of_ENR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3376 | "\<lbrakk>ENR T; S retract_of T\<rbrakk> \<Longrightarrow> ENR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3377 | by (meson ENR_def retract_of_trans) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3378 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3379 | lemma retract_of_UNIV: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3380 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3381 | shows "S retract_of UNIV \<longleftrightarrow> AR S \<and> closed S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3382 | 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) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3383 | |
| 64122 | 3384 | lemma compact_AR: | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3385 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3386 | shows "compact S \<and> AR S \<longleftrightarrow> compact S \<and> S retract_of UNIV" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3387 | using compact_imp_closed retract_of_UNIV by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3388 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3389 | subsection\<open>More properties of ARs, ANRs and ENRs\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3390 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3391 | lemma not_AR_empty [simp]: "~ AR({})"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3392 | by (auto simp: AR_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3393 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3394 | lemma ENR_empty [simp]: "ENR {}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3395 | by (simp add: ENR_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3396 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3397 | lemma ANR_empty [simp]: "ANR ({} :: 'a::euclidean_space set)"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3398 | by (simp add: ENR_imp_ANR) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3399 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3400 | lemma convex_imp_AR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3401 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3402 |   shows "\<lbrakk>convex S; S \<noteq> {}\<rbrakk> \<Longrightarrow> AR S"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3403 | apply (rule absolute_extensor_imp_AR) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3404 | apply (rule Dugundji, assumption+) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3405 | by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3406 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3407 | lemma convex_imp_ANR: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3408 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3409 | shows "convex S \<Longrightarrow> ANR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3410 | using ANR_empty AR_imp_ANR convex_imp_AR by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3411 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3412 | lemma ENR_convex_closed: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3413 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3414 | shows "\<lbrakk>closed S; convex S\<rbrakk> \<Longrightarrow> ENR S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3415 | using ENR_def ENR_empty convex_imp_AR retract_of_UNIV by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3416 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3417 | lemma AR_UNIV [simp]: "AR (UNIV :: 'a::euclidean_space set)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3418 | using retract_of_UNIV by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3419 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3420 | lemma ANR_UNIV [simp]: "ANR (UNIV :: 'a::euclidean_space set)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3421 | by (simp add: AR_imp_ANR) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3422 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3423 | lemma ENR_UNIV [simp]:"ENR UNIV" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3424 | using ENR_def by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3425 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3426 | lemma AR_singleton: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3427 | fixes a :: "'a::euclidean_space" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3428 |     shows "AR {a}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3429 | using retract_of_UNIV by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3430 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3431 | lemma ANR_singleton: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3432 | fixes a :: "'a::euclidean_space" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3433 |     shows "ANR {a}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3434 | by (simp add: AR_imp_ANR AR_singleton) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3435 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3436 | lemma ENR_singleton: "ENR {a}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3437 | using ENR_def by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3438 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3439 | subsection\<open>ARs closed under union\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3440 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3441 | lemma AR_closed_Un_local_aux: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3442 | fixes U :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3443 | assumes "closedin (subtopology euclidean U) S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3444 | "closedin (subtopology euclidean U) T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3445 | "AR S" "AR T" "AR(S \<inter> T)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3446 | shows "(S \<union> T) retract_of U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3447 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3448 |   have "S \<inter> T \<noteq> {}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3449 | using assms AR_def by fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3450 | have "S \<subseteq> U" "T \<subseteq> U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3451 | using assms by (auto simp: closedin_imp_subset) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3452 |   define S' where "S' \<equiv> {x \<in> U. setdist {x} S \<le> setdist {x} T}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3453 |   define T' where "T' \<equiv> {x \<in> U. setdist {x} T \<le> setdist {x} S}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3454 |   define W  where "W \<equiv> {x \<in> U. setdist {x} S = setdist {x} T}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3455 | have US': "closedin (subtopology euclidean U) S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3456 |     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} S - setdist {x} T" "{..0}"]
 | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3457 | by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3458 | have UT': "closedin (subtopology euclidean U) T'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3459 |     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} T - setdist {x} S" "{..0}"]
 | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3460 | by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3461 | have "S \<subseteq> S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3462 | using S'_def \<open>S \<subseteq> U\<close> setdist_sing_in_set by fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3463 | have "T \<subseteq> T'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3464 | using T'_def \<open>T \<subseteq> U\<close> setdist_sing_in_set by fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3465 | have "S \<inter> T \<subseteq> W" "W \<subseteq> U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3466 | using \<open>S \<subseteq> U\<close> by (auto simp: W_def setdist_sing_in_set) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3467 | have "(S \<inter> T) retract_of W" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3468 | apply (rule AR_imp_absolute_retract [OF \<open>AR(S \<inter> T)\<close>]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3469 | apply (simp add: homeomorphic_refl) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3470 | apply (rule closedin_subset_trans [of U]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3471 | apply (simp_all add: assms closedin_Int \<open>S \<inter> T \<subseteq> W\<close> \<open>W \<subseteq> U\<close>) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3472 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3473 | then obtain r0 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3474 | where "S \<inter> T \<subseteq> W" and contr0: "continuous_on W r0" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3475 | and "r0 ` W \<subseteq> S \<inter> T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3476 | and r0 [simp]: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> r0 x = x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3477 | by (auto simp: retract_of_def retraction_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3478 | have ST: "x \<in> W \<Longrightarrow> x \<in> S \<longleftrightarrow> x \<in> T" for x | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3479 |     using setdist_eq_0_closedin \<open>S \<inter> T \<noteq> {}\<close> assms
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3480 | by (force simp: W_def setdist_sing_in_set) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3481 | have "S' \<inter> T' = W" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3482 | by (auto simp: S'_def T'_def W_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3483 | then have cloUW: "closedin (subtopology euclidean U) W" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3484 | using closedin_Int US' UT' by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3485 | define r where "r \<equiv> \<lambda>x. if x \<in> W then r0 x else x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3486 | have "r ` (W \<union> S) \<subseteq> S" "r ` (W \<union> T) \<subseteq> T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3487 | using \<open>r0 ` W \<subseteq> S \<inter> T\<close> r_def by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3488 | have contr: "continuous_on (W \<union> (S \<union> T)) r" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3489 | unfolding r_def | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3490 | proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3491 | show "closedin (subtopology euclidean (W \<union> (S \<union> T))) W" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3492 | 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 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3493 | show "closedin (subtopology euclidean (W \<union> (S \<union> T))) (S \<union> T)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3494 | 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) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3495 | show "\<And>x. x \<in> W \<and> x \<notin> W \<or> x \<in> S \<union> T \<and> x \<in> W \<Longrightarrow> r0 x = x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3496 | by (auto simp: ST) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3497 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3498 | have cloUWS: "closedin (subtopology euclidean U) (W \<union> S)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3499 | by (simp add: cloUW assms closedin_Un) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3500 | obtain g where contg: "continuous_on U g" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3501 | and "g ` U \<subseteq> S" and geqr: "\<And>x. x \<in> W \<union> S \<Longrightarrow> g x = r x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3502 | apply (rule AR_imp_absolute_extensor [OF \<open>AR S\<close> _ _ cloUWS]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3503 | apply (rule continuous_on_subset [OF contr]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3504 | using \<open>r ` (W \<union> S) \<subseteq> S\<close> apply auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3505 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3506 | have cloUWT: "closedin (subtopology euclidean U) (W \<union> T)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3507 | by (simp add: cloUW assms closedin_Un) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3508 | obtain h where conth: "continuous_on U h" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3509 | and "h ` U \<subseteq> T" and heqr: "\<And>x. x \<in> W \<union> T \<Longrightarrow> h x = r x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3510 | apply (rule AR_imp_absolute_extensor [OF \<open>AR T\<close> _ _ cloUWT]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3511 | apply (rule continuous_on_subset [OF contr]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3512 | using \<open>r ` (W \<union> T) \<subseteq> T\<close> apply auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3513 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3514 | have "U = S' \<union> T'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3515 | by (force simp: S'_def T'_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3516 | then have cont: "continuous_on U (\<lambda>x. if x \<in> S' then g x else h x)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3517 | apply (rule ssubst) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3518 | apply (rule continuous_on_cases_local) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3519 | using US' UT' \<open>S' \<inter> T' = W\<close> \<open>U = S' \<union> T'\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3520 | contg conth continuous_on_subset geqr heqr apply auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3521 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3522 | have UST: "(\<lambda>x. if x \<in> S' then g x else h x) ` U \<subseteq> S \<union> T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3523 | using \<open>g ` U \<subseteq> S\<close> \<open>h ` U \<subseteq> T\<close> by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3524 | show ?thesis | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3525 | apply (simp add: retract_of_def retraction_def \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close>) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3526 | apply (rule_tac x="\<lambda>x. if x \<in> S' then g x else h x" in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3527 | apply (intro conjI cont UST) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3528 | 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) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3529 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3530 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3531 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3532 | proposition AR_closed_Un_local: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3533 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3534 | assumes STS: "closedin (subtopology euclidean (S \<union> T)) S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3535 | and STT: "closedin (subtopology euclidean (S \<union> T)) T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3536 | and "AR S" "AR T" "AR(S \<inter> T)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3537 | shows "AR(S \<union> T)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3538 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3539 | have "C retract_of U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3540 | if hom: "S \<union> T homeomorphic C" and UC: "closedin (subtopology euclidean U) C" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3541 |        for U and C :: "('a * real) set"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3542 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3543 | obtain f g where hom: "homeomorphism (S \<union> T) C f g" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3544 | using hom by (force simp: homeomorphic_def) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3545 | have US: "closedin (subtopology euclidean U) (C \<inter> g -` S)" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3546 | apply (rule closedin_trans [OF _ UC]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3547 | apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3548 | using hom homeomorphism_def apply blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3549 | apply (metis hom homeomorphism_def set_eq_subset) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3550 | done | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3551 | have UT: "closedin (subtopology euclidean U) (C \<inter> g -` T)" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3552 | apply (rule closedin_trans [OF _ UC]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3553 | apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3554 | using hom homeomorphism_def apply blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3555 | apply (metis hom homeomorphism_def set_eq_subset) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3556 | done | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3557 | have ARS: "AR (C \<inter> g -` S)" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3558 | apply (rule AR_homeomorphic_AR [OF \<open>AR S\<close>]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3559 | apply (simp add: homeomorphic_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3560 | apply (rule_tac x=g in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3561 | apply (rule_tac x=f in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3562 | using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3563 | apply (rule_tac x="f x" in image_eqI, auto) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3564 | done | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3565 | have ART: "AR (C \<inter> g -` T)" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3566 | apply (rule AR_homeomorphic_AR [OF \<open>AR T\<close>]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3567 | apply (simp add: homeomorphic_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3568 | apply (rule_tac x=g in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3569 | apply (rule_tac x=f in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3570 | using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3571 | apply (rule_tac x="f x" in image_eqI, auto) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3572 | done | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3573 | have ARI: "AR ((C \<inter> g -` S) \<inter> (C \<inter> g -` T))" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3574 | apply (rule AR_homeomorphic_AR [OF \<open>AR (S \<inter> T)\<close>]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3575 | apply (simp add: homeomorphic_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3576 | apply (rule_tac x=g in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3577 | apply (rule_tac x=f in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3578 | using hom | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3579 | apply (auto simp: homeomorphism_def elim!: continuous_on_subset) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3580 | apply (rule_tac x="f x" in image_eqI, auto) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3581 | done | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3582 | have "C = (C \<inter> g -` S) \<union> (C \<inter> g -` T)" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3583 | using hom by (auto simp: homeomorphism_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3584 | then show ?thesis | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3585 | by (metis AR_closed_Un_local_aux [OF US UT ARS ART ARI]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3586 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3587 | then show ?thesis | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3588 | by (force simp: AR_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3589 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3590 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3591 | corollary AR_closed_Un: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3592 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3593 | shows "\<lbrakk>closed S; closed T; AR S; AR T; AR (S \<inter> T)\<rbrakk> \<Longrightarrow> AR (S \<union> T)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3594 | by (metis AR_closed_Un_local_aux closed_closedin retract_of_UNIV subtopology_UNIV) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3595 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3596 | subsection\<open>ANRs closed under union\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3597 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3598 | lemma ANR_closed_Un_local_aux: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3599 | fixes U :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3600 | assumes US: "closedin (subtopology euclidean U) S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3601 | and UT: "closedin (subtopology euclidean U) T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3602 | and "ANR S" "ANR T" "ANR(S \<inter> T)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3603 | obtains V where "openin (subtopology euclidean U) V" "(S \<union> T) retract_of V" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3604 | proof (cases "S = {} \<or> T = {}")
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3605 | case True with assms that show ?thesis | 
| 65585 
a043de9ad41e
Some fixes related to compactE_image
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 3606 | by (metis ANR_imp_neighbourhood_retract Un_commute inf_bot_right sup_inf_absorb) | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3607 | next | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3608 | case False | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3609 |   then have [simp]: "S \<noteq> {}" "T \<noteq> {}" by auto
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3610 | have "S \<subseteq> U" "T \<subseteq> U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3611 | using assms by (auto simp: closedin_imp_subset) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3612 |   define S' where "S' \<equiv> {x \<in> U. setdist {x} S \<le> setdist {x} T}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3613 |   define T' where "T' \<equiv> {x \<in> U. setdist {x} T \<le> setdist {x} S}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3614 |   define W  where "W \<equiv> {x \<in> U. setdist {x} S = setdist {x} T}"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3615 | have cloUS': "closedin (subtopology euclidean U) S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3616 |     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} S - setdist {x} T" "{..0}"]
 | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3617 | by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3618 | have cloUT': "closedin (subtopology euclidean U) T'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3619 |     using continuous_closedin_preimage [of U "\<lambda>x. setdist {x} T - setdist {x} S" "{..0}"]
 | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3620 | by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3621 | have "S \<subseteq> S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3622 | using S'_def \<open>S \<subseteq> U\<close> setdist_sing_in_set by fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3623 | have "T \<subseteq> T'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3624 | using T'_def \<open>T \<subseteq> U\<close> setdist_sing_in_set by fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3625 | have "S' \<union> T' = U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3626 | by (auto simp: S'_def T'_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3627 | have "W \<subseteq> S'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3628 | by (simp add: Collect_mono S'_def W_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3629 | have "W \<subseteq> T'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3630 | by (simp add: Collect_mono T'_def W_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3631 | have ST_W: "S \<inter> T \<subseteq> W" and "W \<subseteq> U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3632 | using \<open>S \<subseteq> U\<close> by (force simp: W_def setdist_sing_in_set)+ | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3633 | have "S' \<inter> T' = W" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3634 | by (auto simp: S'_def T'_def W_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3635 | then have cloUW: "closedin (subtopology euclidean U) W" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3636 | using closedin_Int cloUS' cloUT' by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3637 | obtain W' W0 where "openin (subtopology euclidean W) W'" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3638 | and cloWW0: "closedin (subtopology euclidean W) W0" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3639 | and "S \<inter> T \<subseteq> W'" "W' \<subseteq> W0" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3640 | and ret: "(S \<inter> T) retract_of W0" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3641 | apply (rule ANR_imp_closed_neighbourhood_retract [OF \<open>ANR(S \<inter> T)\<close>]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3642 | apply (rule closedin_subset_trans [of U, OF _ ST_W \<open>W \<subseteq> U\<close>]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3643 | apply (blast intro: assms)+ | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3644 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3645 | then obtain U0 where opeUU0: "openin (subtopology euclidean U) U0" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3646 | and U0: "S \<inter> T \<subseteq> U0" "U0 \<inter> W \<subseteq> W0" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3647 | unfolding openin_open using \<open>W \<subseteq> U\<close> by blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3648 | have "W0 \<subseteq> U" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3649 | using \<open>W \<subseteq> U\<close> cloWW0 closedin_subset by fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3650 | obtain r0 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3651 | where "S \<inter> T \<subseteq> W0" and contr0: "continuous_on W0 r0" and "r0 ` W0 \<subseteq> S \<inter> T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3652 | and r0 [simp]: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> r0 x = x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3653 | using ret by (force simp add: retract_of_def retraction_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3654 | have ST: "x \<in> W \<Longrightarrow> x \<in> S \<longleftrightarrow> x \<in> T" for x | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3655 | using assms by (auto simp: W_def setdist_sing_in_set dest!: setdist_eq_0_closedin) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3656 | define r where "r \<equiv> \<lambda>x. if x \<in> W0 then r0 x else x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3657 | have "r ` (W0 \<union> S) \<subseteq> S" "r ` (W0 \<union> T) \<subseteq> T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3658 | using \<open>r0 ` W0 \<subseteq> S \<inter> T\<close> r_def by auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3659 | have contr: "continuous_on (W0 \<union> (S \<union> T)) r" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3660 | unfolding r_def | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3661 | proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3662 | show "closedin (subtopology euclidean (W0 \<union> (S \<union> T))) W0" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3663 | apply (rule closedin_subset_trans [of U]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3664 | using cloWW0 cloUW closedin_trans \<open>W0 \<subseteq> U\<close> \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> apply blast+ | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3665 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3666 | show "closedin (subtopology euclidean (W0 \<union> (S \<union> T))) (S \<union> T)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3667 | 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) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3668 | show "\<And>x. x \<in> W0 \<and> x \<notin> W0 \<or> x \<in> S \<union> T \<and> x \<in> W0 \<Longrightarrow> r0 x = x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3669 | using ST cloWW0 closedin_subset by fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3670 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3671 | have cloS'WS: "closedin (subtopology euclidean S') (W0 \<union> S)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3672 | by (meson closedin_subset_trans US cloUS' \<open>S \<subseteq> S'\<close> \<open>W \<subseteq> S'\<close> cloUW cloWW0 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3673 | closedin_Un closedin_imp_subset closedin_trans) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3674 | obtain W1 g where "W0 \<union> S \<subseteq> W1" and contg: "continuous_on W1 g" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3675 | and opeSW1: "openin (subtopology euclidean S') W1" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3676 | and "g ` W1 \<subseteq> S" and geqr: "\<And>x. x \<in> W0 \<union> S \<Longrightarrow> g x = r x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3677 | apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> _ \<open>r ` (W0 \<union> S) \<subseteq> S\<close> cloS'WS]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3678 | apply (rule continuous_on_subset [OF contr]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3679 | apply (blast intro: elim: )+ | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3680 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3681 | have cloT'WT: "closedin (subtopology euclidean T') (W0 \<union> T)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3682 | by (meson closedin_subset_trans UT cloUT' \<open>T \<subseteq> T'\<close> \<open>W \<subseteq> T'\<close> cloUW cloWW0 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3683 | closedin_Un closedin_imp_subset closedin_trans) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3684 | obtain W2 h where "W0 \<union> T \<subseteq> W2" and conth: "continuous_on W2 h" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3685 | and opeSW2: "openin (subtopology euclidean T') W2" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3686 | and "h ` W2 \<subseteq> T" and heqr: "\<And>x. x \<in> W0 \<union> T \<Longrightarrow> h x = r x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3687 | apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> _ \<open>r ` (W0 \<union> T) \<subseteq> T\<close> cloT'WT]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3688 | apply (rule continuous_on_subset [OF contr]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3689 | apply (blast intro: elim: )+ | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3690 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3691 | have "S' \<inter> T' = W" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3692 | by (force simp: S'_def T'_def W_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3693 | obtain O1 O2 where "open O1" "W1 = S' \<inter> O1" "open O2" "W2 = T' \<inter> O2" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3694 | using opeSW1 opeSW2 by (force simp add: openin_open) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3695 | show ?thesis | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3696 | proof | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3697 | have eq: "W1 - (W - U0) \<union> (W2 - (W - U0)) = | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3698 | ((U - T') \<inter> O1 \<union> (U - S') \<inter> O2 \<union> U \<inter> O1 \<inter> O2) - (W - U0)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3699 | using \<open>U0 \<inter> W \<subseteq> W0\<close> \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W0 \<union> T \<subseteq> W2\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3700 | 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>) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3701 | show "openin (subtopology euclidean U) (W1 - (W - U0) \<union> (W2 - (W - U0)))" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3702 | apply (subst eq) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3703 | apply (intro openin_Un openin_Int_open openin_diff closedin_diff cloUW opeUU0 cloUS' cloUT' \<open>open O1\<close> \<open>open O2\<close>) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3704 | apply simp_all | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3705 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3706 | have cloW1: "closedin (subtopology euclidean (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W1 - (W - U0))" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3707 | using cloUS' apply (simp add: closedin_closed) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3708 | apply (erule ex_forward) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3709 | using U0 \<open>W0 \<union> S \<subseteq> W1\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3710 | apply (auto simp add: \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<union> T' = U\<close> [symmetric]\<open>S' \<inter> T' = W\<close> [symmetric]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3711 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3712 | have cloW2: "closedin (subtopology euclidean (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W2 - (W - U0))" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3713 | using cloUT' apply (simp add: closedin_closed) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3714 | apply (erule ex_forward) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3715 | using U0 \<open>W0 \<union> T \<subseteq> W2\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3716 | apply (auto simp add: \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<union> T' = U\<close> [symmetric]\<open>S' \<inter> T' = W\<close> [symmetric]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3717 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3718 | have *: "\<forall>x\<in>S \<union> T. (if x \<in> S' then g x else h x) = x" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3719 | using ST \<open>S' \<inter> T' = W\<close> cloT'WT closedin_subset geqr heqr | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3720 | apply (auto simp: r_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3721 | apply fastforce | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3722 | 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 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3723 | have "\<exists>r. continuous_on (W1 - (W - U0) \<union> (W2 - (W - U0))) r \<and> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3724 | r ` (W1 - (W - U0) \<union> (W2 - (W - U0))) \<subseteq> S \<union> T \<and> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3725 | (\<forall>x\<in>S \<union> T. r x = x)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3726 | apply (rule_tac x = "\<lambda>x. if x \<in> S' then g x else h x" in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3727 | apply (intro conjI *) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3728 | apply (rule continuous_on_cases_local | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3729 | [OF cloW1 cloW2 continuous_on_subset [OF contg] continuous_on_subset [OF conth]]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3730 | using \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<inter> T' = W\<close> | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3731 | \<open>g ` W1 \<subseteq> S\<close> \<open>h ` W2 \<subseteq> T\<close> apply auto | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3732 | using \<open>U0 \<inter> W \<subseteq> W0\<close> \<open>W0 \<union> S \<subseteq> W1\<close> apply (fastforce simp add: geqr heqr)+ | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3733 | done | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3734 | then show "S \<union> T retract_of W1 - (W - U0) \<union> (W2 - (W - U0))" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3735 | using \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W0 \<union> T \<subseteq> W2\<close> ST opeUU0 U0 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3736 | by (auto simp add: retract_of_def retraction_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3737 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3738 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3739 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3740 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3741 | proposition ANR_closed_Un_local: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3742 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3743 | assumes STS: "closedin (subtopology euclidean (S \<union> T)) S" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3744 | and STT: "closedin (subtopology euclidean (S \<union> T)) T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3745 | and "ANR S" "ANR T" "ANR(S \<inter> T)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3746 | shows "ANR(S \<union> T)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3747 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3748 | have "\<exists>T. openin (subtopology euclidean U) T \<and> C retract_of T" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3749 | if hom: "S \<union> T homeomorphic C" and UC: "closedin (subtopology euclidean U) C" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3750 |        for U and C :: "('a * real) set"
 | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3751 | proof - | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3752 | obtain f g where hom: "homeomorphism (S \<union> T) C f g" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3753 | using hom by (force simp: homeomorphic_def) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3754 | have US: "closedin (subtopology euclidean U) (C \<inter> g -` S)" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3755 | apply (rule closedin_trans [OF _ UC]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3756 | apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3757 | using hom [unfolded homeomorphism_def] apply blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3758 | apply (metis hom homeomorphism_def set_eq_subset) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3759 | done | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3760 | have UT: "closedin (subtopology euclidean U) (C \<inter> g -` T)" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3761 | apply (rule closedin_trans [OF _ UC]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3762 | apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3763 | using hom [unfolded homeomorphism_def] apply blast | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3764 | apply (metis hom homeomorphism_def set_eq_subset) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3765 | done | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3766 | have ANRS: "ANR (C \<inter> g -` S)" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3767 | apply (rule ANR_homeomorphic_ANR [OF \<open>ANR S\<close>]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3768 | apply (simp add: homeomorphic_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3769 | apply (rule_tac x=g in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3770 | apply (rule_tac x=f in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3771 | using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3772 | apply (rule_tac x="f x" in image_eqI, auto) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3773 | done | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3774 | have ANRT: "ANR (C \<inter> g -` T)" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3775 | apply (rule ANR_homeomorphic_ANR [OF \<open>ANR T\<close>]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3776 | apply (simp add: homeomorphic_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3777 | apply (rule_tac x=g in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3778 | apply (rule_tac x=f in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3779 | using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3780 | apply (rule_tac x="f x" in image_eqI, auto) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3781 | done | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3782 | have ANRI: "ANR ((C \<inter> g -` S) \<inter> (C \<inter> g -` T))" | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3783 | apply (rule ANR_homeomorphic_ANR [OF \<open>ANR (S \<inter> T)\<close>]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3784 | apply (simp add: homeomorphic_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3785 | apply (rule_tac x=g in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3786 | apply (rule_tac x=f in exI) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3787 | using hom | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3788 | apply (auto simp: homeomorphism_def elim!: continuous_on_subset) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3789 | apply (rule_tac x="f x" in image_eqI, auto) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3790 | done | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3791 | have "C = (C \<inter> g -` S) \<union> (C \<inter> g -` T)" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3792 | using hom by (auto simp: homeomorphism_def) | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3793 | then show ?thesis | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3794 | by (metis ANR_closed_Un_local_aux [OF US UT ANRS ANRT ANRI]) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3795 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3796 | then show ?thesis | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3797 | by (auto simp: ANR_def) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3798 | qed | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3799 | |
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3800 | corollary ANR_closed_Un: | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3801 | fixes S :: "'a::euclidean_space set" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3802 | shows "\<lbrakk>closed S; closed T; ANR S; ANR T; ANR (S \<inter> T)\<rbrakk> \<Longrightarrow> ANR (S \<union> T)" | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3803 | by (simp add: ANR_closed_Un_local closedin_def diff_eq open_Compl openin_open_Int) | 
| 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 3804 | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3805 | lemma ANR_openin: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3806 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3807 | assumes "ANR T" and opeTS: "openin (subtopology euclidean T) S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3808 | shows "ANR S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3809 | proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3810 | fix f :: "'a \<times> real \<Rightarrow> 'a" and U C | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3811 | assume contf: "continuous_on C f" and fim: "f ` C \<subseteq> S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3812 | and cloUC: "closedin (subtopology euclidean U) C" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3813 | have "f ` C \<subseteq> T" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3814 | using fim opeTS openin_imp_subset by blast | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3815 | obtain W g where "C \<subseteq> W" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3816 | and UW: "openin (subtopology euclidean U) W" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3817 | and contg: "continuous_on W g" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3818 | and gim: "g ` W \<subseteq> T" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3819 | and geq: "\<And>x. x \<in> C \<Longrightarrow> g x = f x" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3820 | apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> contf \<open>f ` C \<subseteq> T\<close> cloUC]) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3821 | using fim by auto | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3822 | 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)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3823 | proof (intro exI conjI) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3824 | show "C \<subseteq> W \<inter> g -` S" | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3825 | using \<open>C \<subseteq> W\<close> fim geq by blast | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3826 | show "openin (subtopology euclidean U) (W \<inter> g -` S)" | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3827 | by (metis (mono_tags, lifting) UW contg continuous_openin_preimage gim opeTS openin_trans) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3828 | show "continuous_on (W \<inter> g -` S) g" | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3829 | by (blast intro: continuous_on_subset [OF contg]) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 3830 | show "g ` (W \<inter> g -` S) \<subseteq> S" | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3831 | using gim by blast | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3832 | show "\<forall>x\<in>C. g x = f x" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3833 | using geq by blast | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3834 | qed | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3835 | qed | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3836 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3837 | lemma ENR_openin: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3838 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3839 | assumes "ENR T" and opeTS: "openin (subtopology euclidean T) S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3840 | shows "ENR S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3841 | using assms apply (simp add: ENR_ANR) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3842 | using ANR_openin locally_open_subset by blast | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3843 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3844 | lemma ANR_neighborhood_retract: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3845 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3846 | assumes "ANR U" "S retract_of T" "openin (subtopology euclidean U) T" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3847 | shows "ANR S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3848 | using ANR_openin ANR_retract_of_ANR assms by blast | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3849 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3850 | lemma ENR_neighborhood_retract: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3851 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3852 | assumes "ENR U" "S retract_of T" "openin (subtopology euclidean U) T" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3853 | shows "ENR S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3854 | using ENR_openin ENR_retract_of_ENR assms by blast | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3855 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3856 | lemma ANR_rel_interior: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3857 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3858 | shows "ANR S \<Longrightarrow> ANR(rel_interior S)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3859 | by (blast intro: ANR_openin openin_set_rel_interior) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3860 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3861 | lemma ANR_delete: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3862 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3863 |   shows "ANR S \<Longrightarrow> ANR(S - {a})"
 | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3864 | by (blast intro: ANR_openin openin_delete openin_subtopology_self) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3865 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3866 | lemma ENR_rel_interior: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3867 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3868 | shows "ENR S \<Longrightarrow> ENR(rel_interior S)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3869 | by (blast intro: ENR_openin openin_set_rel_interior) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3870 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3871 | lemma ENR_delete: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3872 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3873 |   shows "ENR S \<Longrightarrow> ENR(S - {a})"
 | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3874 | by (blast intro: ENR_openin openin_delete openin_subtopology_self) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3875 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3876 | lemma open_imp_ENR: "open S \<Longrightarrow> ENR S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3877 | using ENR_def by blast | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3878 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3879 | lemma open_imp_ANR: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3880 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3881 | shows "open S \<Longrightarrow> ANR S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3882 | by (simp add: ENR_imp_ANR open_imp_ENR) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3883 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3884 | lemma ANR_ball [iff]: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3885 | fixes a :: "'a::euclidean_space" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3886 | shows "ANR(ball a r)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3887 | by (simp add: convex_imp_ANR) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3888 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3889 | lemma ENR_ball [iff]: "ENR(ball a r)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3890 | by (simp add: open_imp_ENR) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3891 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3892 | lemma AR_ball [simp]: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3893 | fixes a :: "'a::euclidean_space" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3894 | shows "AR(ball a r) \<longleftrightarrow> 0 < r" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3895 | by (auto simp: AR_ANR convex_imp_contractible) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3896 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3897 | lemma ANR_cball [iff]: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3898 | fixes a :: "'a::euclidean_space" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3899 | shows "ANR(cball a r)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3900 | by (simp add: convex_imp_ANR) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3901 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3902 | lemma ENR_cball: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3903 | fixes a :: "'a::euclidean_space" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3904 | shows "ENR(cball a r)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3905 | using ENR_convex_closed by blast | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3906 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3907 | lemma AR_cball [simp]: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3908 | fixes a :: "'a::euclidean_space" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3909 | shows "AR(cball a r) \<longleftrightarrow> 0 \<le> r" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3910 | by (auto simp: AR_ANR convex_imp_contractible) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3911 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3912 | lemma ANR_box [iff]: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3913 | fixes a :: "'a::euclidean_space" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3914 | shows "ANR(cbox a b)" "ANR(box a b)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3915 | by (auto simp: convex_imp_ANR open_imp_ANR) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3916 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3917 | lemma ENR_box [iff]: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3918 | fixes a :: "'a::euclidean_space" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3919 | shows "ENR(cbox a b)" "ENR(box a b)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3920 | apply (simp add: ENR_convex_closed closed_cbox) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3921 | by (simp add: open_box open_imp_ENR) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3922 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3923 | lemma AR_box [simp]: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3924 |     "AR(cbox a b) \<longleftrightarrow> cbox a b \<noteq> {}" "AR(box a b) \<longleftrightarrow> box a b \<noteq> {}"
 | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3925 | by (auto simp: AR_ANR convex_imp_contractible) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3926 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3927 | lemma ANR_interior: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3928 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3929 | shows "ANR(interior S)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3930 | by (simp add: open_imp_ANR) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3931 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3932 | lemma ENR_interior: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3933 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3934 | shows "ENR(interior S)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3935 | by (simp add: open_imp_ENR) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3936 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3937 | lemma AR_imp_contractible: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3938 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3939 | shows "AR S \<Longrightarrow> contractible S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3940 | by (simp add: AR_ANR) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3941 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3942 | lemma ENR_imp_locally_compact: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3943 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3944 | shows "ENR S \<Longrightarrow> locally compact S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3945 | by (simp add: ENR_ANR) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3946 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3947 | lemma ANR_imp_locally_path_connected: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3948 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3949 | assumes "ANR S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3950 | shows "locally path_connected S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3951 | proof - | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3952 |   obtain U and T :: "('a \<times> real) set"
 | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3953 |      where "convex U" "U \<noteq> {}"
 | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3954 | and UT: "closedin (subtopology euclidean U) T" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3955 | and "S homeomorphic T" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3956 | apply (rule homeomorphic_closedin_convex [of S]) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3957 | using aff_dim_le_DIM [of S] apply auto | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3958 | done | 
| 63928 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3959 | then have "locally path_connected T" | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3960 | by (meson ANR_imp_absolute_neighbourhood_retract | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 3961 | assms convex_imp_locally_path_connected locally_open_subset retract_of_locally_path_connected) | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3962 | then have S: "locally path_connected S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3963 |       if "openin (subtopology euclidean U) V" "T retract_of V" "U \<noteq> {}" for V
 | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3964 | using \<open>S homeomorphic T\<close> homeomorphic_locally homeomorphic_path_connectedness by blast | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3965 | show ?thesis | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3966 | using assms | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3967 | apply (clarsimp simp: ANR_def) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3968 | apply (drule_tac x=U in spec) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3969 | apply (drule_tac x=T in spec) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3970 |     using \<open>S homeomorphic T\<close> \<open>U \<noteq> {}\<close> UT  apply (blast intro: S)
 | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3971 | done | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3972 | qed | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3973 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3974 | lemma ANR_imp_locally_connected: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3975 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3976 | assumes "ANR S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3977 | shows "locally connected S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3978 | using locally_path_connected_imp_locally_connected ANR_imp_locally_path_connected assms by auto | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3979 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3980 | lemma AR_imp_locally_path_connected: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3981 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3982 | assumes "AR S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3983 | shows "locally path_connected S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3984 | by (simp add: ANR_imp_locally_path_connected AR_imp_ANR assms) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3985 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3986 | lemma AR_imp_locally_connected: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3987 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3988 | assumes "AR S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3989 | shows "locally connected S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3990 | using ANR_imp_locally_connected AR_ANR assms by blast | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3991 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3992 | lemma ENR_imp_locally_path_connected: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3993 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3994 | assumes "ENR S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3995 | shows "locally path_connected S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3996 | by (simp add: ANR_imp_locally_path_connected ENR_imp_ANR assms) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3997 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3998 | lemma ENR_imp_locally_connected: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 3999 | fixes S :: "'a::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4000 | assumes "ENR S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4001 | shows "locally connected S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4002 | using ANR_imp_locally_connected ENR_ANR assms by blast | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4003 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4004 | lemma ANR_Times: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4005 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4006 | assumes "ANR S" "ANR T" shows "ANR(S \<times> T)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4007 | proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4008 |   fix f :: " ('a \<times> 'b) \<times> real \<Rightarrow> 'a \<times> 'b" and U C
 | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4009 | assume "continuous_on C f" and fim: "f ` C \<subseteq> S \<times> T" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4010 | and cloUC: "closedin (subtopology euclidean U) C" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4011 | have contf1: "continuous_on C (fst \<circ> f)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4012 | by (simp add: \<open>continuous_on C f\<close> continuous_on_fst) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4013 | obtain W1 g where "C \<subseteq> W1" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4014 | and UW1: "openin (subtopology euclidean U) W1" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4015 | and contg: "continuous_on W1 g" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4016 | and gim: "g ` W1 \<subseteq> S" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4017 | and geq: "\<And>x. x \<in> C \<Longrightarrow> g x = (fst \<circ> f) x" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4018 | apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> contf1 _ cloUC]) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4019 | using fim apply auto | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4020 | done | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4021 | have contf2: "continuous_on C (snd \<circ> f)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4022 | by (simp add: \<open>continuous_on C f\<close> continuous_on_snd) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4023 | obtain W2 h where "C \<subseteq> W2" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4024 | and UW2: "openin (subtopology euclidean U) W2" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4025 | and conth: "continuous_on W2 h" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4026 | and him: "h ` W2 \<subseteq> T" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4027 | and heq: "\<And>x. x \<in> C \<Longrightarrow> h x = (snd \<circ> f) x" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4028 | apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> contf2 _ cloUC]) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4029 | using fim apply auto | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4030 | done | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4031 | show "\<exists>V g. C \<subseteq> V \<and> | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4032 | openin (subtopology euclidean U) V \<and> | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4033 | continuous_on V g \<and> g ` V \<subseteq> S \<times> T \<and> (\<forall>x\<in>C. g x = f x)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4034 | proof (intro exI conjI) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4035 | show "C \<subseteq> W1 \<inter> W2" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4036 | by (simp add: \<open>C \<subseteq> W1\<close> \<open>C \<subseteq> W2\<close>) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4037 | show "openin (subtopology euclidean U) (W1 \<inter> W2)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4038 | by (simp add: UW1 UW2 openin_Int) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4039 | show "continuous_on (W1 \<inter> W2) (\<lambda>x. (g x, h x))" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4040 | by (metis (no_types) contg conth continuous_on_Pair continuous_on_subset inf_commute inf_le1) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4041 | show "(\<lambda>x. (g x, h x)) ` (W1 \<inter> W2) \<subseteq> S \<times> T" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4042 | using gim him by blast | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4043 | show "(\<forall>x\<in>C. (g x, h x) = f x)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4044 | using geq heq by auto | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4045 | qed | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4046 | qed | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4047 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4048 | lemma AR_Times: | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4049 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4050 | assumes "AR S" "AR T" shows "AR(S \<times> T)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4051 | using assms by (simp add: AR_ANR ANR_Times contractible_Times) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63365diff
changeset | 4052 | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4053 | lemma ENR_rel_frontier_convex: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4054 | fixes S :: "'a::euclidean_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4055 | assumes "bounded S" "convex S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4056 | shows "ENR(rel_frontier S)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4057 | proof (cases "S = {}")
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4058 | case True then show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4059 | by simp | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4060 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4061 | case False | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4062 |   with assms have "rel_interior S \<noteq> {}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4063 | by (simp add: rel_interior_eq_empty) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4064 | 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 | 4065 | by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4066 |   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 | 4067 | by (auto simp: closest_point_self) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4068 |   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 | 4069 | by (simp add: assms a rel_frontier_retract_of_punctured_affine_hull) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4070 |   also have "... retract_of {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 | 4071 | 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 | 4072 | apply (rule_tac x="closest_point (affine hull S)" in exI) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4073 | apply (auto simp add: False closest_point_self affine_imp_convex closest_point_in_set continuous_on_closest_point) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4074 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4075 |   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 | 4076 |   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 | 4077 | apply (rule continuous_openin_preimage_gen) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4078 | apply (auto simp add: False affine_imp_convex continuous_on_closest_point) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4079 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4080 | ultimately show ?thesis | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 4081 | unfolding ENR_def | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 4082 |     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 | 4083 | apply (simp add: vimage_def) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4084 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4085 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4086 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4087 | lemma ANR_rel_frontier_convex: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4088 | fixes S :: "'a::euclidean_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4089 | assumes "bounded S" "convex S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4090 | shows "ANR(rel_frontier S)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4091 | 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 | 4092 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4093 | lemma ENR_closedin_Un_local: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4094 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4095 | 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 | 4096 | 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 | 4097 | \<Longrightarrow> ENR(S \<union> T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4098 | 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 | 4099 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4100 | lemma ENR_closed_Un: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4101 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4102 | 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 | 4103 | by (auto simp: closed_subset ENR_closedin_Un_local) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4104 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4105 | lemma absolute_retract_Un: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4106 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4107 | 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 | 4108 | \<Longrightarrow> (S \<union> T) retract_of UNIV" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4109 | 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 | 4110 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4111 | lemma retract_from_Un_Int: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4112 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4113 | assumes clS: "closedin (subtopology euclidean (S \<union> T)) S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4114 | and clT: "closedin (subtopology euclidean (S \<union> T)) T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4115 | 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 | 4116 | shows "S retract_of U" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4117 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4118 | 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 | 4119 | 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 | 4120 | have "S retract_of S \<union> T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4121 | unfolding retraction_def retract_of_def | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4122 | proof (intro exI conjI) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4123 | 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 | 4124 | apply (rule continuous_on_cases_local [OF clS clT]) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4125 | using r by (auto simp: continuous_on_id) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4126 | qed (use r in auto) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4127 | also have "... retract_of U" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4128 | by (rule Un) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4129 | finally show ?thesis . | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4130 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4131 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4132 | lemma AR_from_Un_Int_local: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4133 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4134 | assumes clS: "closedin (subtopology euclidean (S \<union> T)) S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4135 | and clT: "closedin (subtopology euclidean (S \<union> T)) T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4136 | 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 | 4137 | shows "AR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4138 | apply (rule AR_retract_of_AR [OF Un]) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4139 | 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 | 4140 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4141 | lemma AR_from_Un_Int_local': | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4142 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4143 | assumes "closedin (subtopology euclidean (S \<union> T)) S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4144 | and "closedin (subtopology euclidean (S \<union> T)) T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4145 | and "AR(S \<union> T)" "AR(S \<inter> T)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4146 | shows "AR T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4147 | 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 | 4148 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4149 | lemma AR_from_Un_Int: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4150 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4151 | 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 | 4152 | shows "AR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4153 | 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 | 4154 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4155 | lemma ANR_from_Un_Int_local: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4156 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4157 | assumes clS: "closedin (subtopology euclidean (S \<union> T)) S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4158 | and clT: "closedin (subtopology euclidean (S \<union> T)) T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4159 | 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 | 4160 | shows "ANR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4161 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4162 | 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 | 4163 | and ope: "openin (subtopology euclidean (S \<union> T)) V" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4164 | and ret: "S \<inter> T retract_of V" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4165 | 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 | 4166 | 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 | 4167 | by (auto simp: retraction_def retract_of_def) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4168 | have Vsub: "V \<subseteq> S \<union> T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4169 | by (meson ope openin_contains_cball) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4170 | have Vsup: "S \<inter> T \<subseteq> V" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4171 | by (simp add: retract_of_imp_subset ret) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4172 | 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 | 4173 | by auto | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4174 | 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 | 4175 | using Vsub by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4176 | 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 | 4177 | proof (rule continuous_on_cases_local) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4178 | 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 | 4179 | using clS closedin_subset_trans inf.boundedE by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4180 | 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 | 4181 | using clT Vsup by (auto simp: closedin_closed) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4182 | show "continuous_on (V \<inter> T) r" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4183 | by (meson Int_lower1 continuous_on_subset r) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4184 | qed (use req continuous_on_id in auto) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4185 | with rim have "S retract_of S \<union> V" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4186 | unfolding retraction_def retract_of_def | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4187 | 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 | 4188 | apply (auto simp: eq') | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4189 | done | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4190 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4191 | using ANR_neighborhood_retract [OF Un] | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4192 | 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 | 4193 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4194 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4195 | lemma ANR_from_Un_Int: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4196 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4197 | 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 | 4198 | shows "ANR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4199 | 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 | 4200 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4201 | proposition ANR_finite_Union_convex_closed: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4202 | fixes \<T> :: "'a::euclidean_space set set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4203 | 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 | 4204 | shows "ANR(\<Union>\<T>)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4205 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4206 | 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 | 4207 | using assms that | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4208 | proof (induction n arbitrary: \<T>) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4209 | case 0 then show ?case by simp | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4210 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4211 | case (Suc n) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4212 | 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 | 4213 | using that | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4214 | proof (induction \<U>) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4215 | case empty | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4216 | then show ?case by simp | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4217 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4218 | case (insert C \<U>) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4219 | have "ANR (C \<union> \<Union>\<U>)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4220 | proof (rule ANR_closed_Un) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4221 | show "ANR (C \<inter> \<Union>\<U>)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4222 | unfolding Int_Union | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4223 | proof (rule Suc) | 
| 67399 | 4224 | show "finite ((\<inter>) C ` \<U>)" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4225 | by (simp add: insert.hyps(1)) | 
| 67399 | 4226 | 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 | 4227 | by (metis (no_types, hide_lams) Suc.prems(2) closed_Int subsetD imageE insert.prems insertI1 insertI2) | 
| 67399 | 4228 | 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 | 4229 | by (metis (mono_tags, lifting) Suc.prems(3) convex_Int imageE insert.prems insert_subset subsetCE) | 
| 67399 | 4230 | show "card ((\<inter>) C ` \<U>) < n" | 
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4231 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4232 | have "card \<T> \<le> n" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4233 | 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 | 4234 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4235 | 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 | 4236 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4237 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4238 | show "closed (\<Union>\<U>)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4239 | 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 | 4240 | 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 | 4241 | then show ?case | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4242 | by simp | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4243 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4244 | then show ?case | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4245 | using Suc.prems(1) by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4246 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4247 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4248 | by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4249 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4250 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4251 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4252 | lemma finite_imp_ANR: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4253 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4254 | assumes "finite S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4255 | shows "ANR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4256 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4257 |   have "ANR(\<Union>x \<in> S. {x})"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4258 | by (blast intro: ANR_finite_Union_convex_closed assms) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4259 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4260 | by simp | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4261 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4262 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4263 | lemma ANR_insert: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4264 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4265 | assumes "ANR S" "closed S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4266 | shows "ANR(insert a S)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4267 | 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 | 4268 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4269 | lemma ANR_path_component_ANR: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4270 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4271 | 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 | 4272 | 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 | 4273 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4274 | lemma ANR_connected_component_ANR: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4275 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4276 | 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 | 4277 | 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 | 4278 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4279 | lemma ANR_component_ANR: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4280 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4281 | assumes "ANR S" "c \<in> components S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4282 | shows "ANR c" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4283 | by (metis ANR_connected_component_ANR assms componentsE) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4284 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4285 | subsection\<open>Original ANR material, now for ENRs.\<close> | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4286 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4287 | lemma ENR_bounded: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4288 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4289 | assumes "bounded S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4290 | 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 | 4291 | (is "?lhs = ?rhs") | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4292 | proof | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4293 | 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 | 4294 | using bounded_subset_ballD assms by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4295 | assume ?lhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4296 | then show ?rhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4297 | apply (clarsimp simp: ENR_def) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4298 | 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 | 4299 | 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 | 4300 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4301 | assume ?rhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4302 | then show ?lhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4303 | using ENR_def by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4304 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4305 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4306 | lemma absolute_retract_imp_AR_gen: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4307 | 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 | 4308 |   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 | 4309 | shows "S' retract_of U" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4310 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4311 | have "AR T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4312 | by (simp add: assms convex_imp_AR) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4313 | then have "AR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4314 | using AR_retract_of_AR assms by auto | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4315 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4316 | using assms AR_imp_absolute_retract by metis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4317 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4318 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4319 | lemma absolute_retract_imp_AR: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4320 | 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 | 4321 | 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 | 4322 | shows "S' retract_of UNIV" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4323 | 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 | 4324 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4325 | lemma homeomorphic_compact_arness: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4326 | 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 | 4327 | assumes "S homeomorphic S'" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4328 | 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 | 4329 | using assms homeomorphic_compactness | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4330 | apply auto | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4331 | 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 | 4332 | done | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4333 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4334 | lemma absolute_retract_from_Un_Int: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4335 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4336 | 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 | 4337 | shows "S retract_of UNIV" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4338 | 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 | 4339 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4340 | lemma ENR_from_Un_Int_gen: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4341 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4342 | 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 | 4343 | shows "ENR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4344 | apply (simp add: ENR_ANR) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4345 | 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 | 4346 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4347 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4348 | lemma ENR_from_Un_Int: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4349 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4350 | 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 | 4351 | shows "ENR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4352 | 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 | 4353 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4354 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4355 | lemma ENR_finite_Union_convex_closed: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4356 | fixes \<T> :: "'a::euclidean_space set set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4357 | 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 | 4358 | shows "ENR(\<Union> \<T>)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4359 | 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 | 4360 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4361 | lemma finite_imp_ENR: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4362 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4363 | shows "finite S \<Longrightarrow> ENR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4364 | 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 | 4365 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4366 | lemma ENR_insert: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4367 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4368 | assumes "closed S" "ENR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4369 | shows "ENR(insert a S)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4370 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4371 |   have "ENR ({a} \<union> S)"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4372 | 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 | 4373 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4374 | by auto | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4375 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4376 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4377 | lemma ENR_path_component_ENR: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4378 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4379 | assumes "ENR S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4380 | shows "ENR(path_component_set S x)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4381 | 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 | 4382 | 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 | 4383 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4384 | (*UNUSED | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4385 | lemma ENR_Times: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4386 | 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 | 4387 | 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 | 4388 | 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 | 4389 | thm locally_compact_Times | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4390 | oops | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4391 | 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 | 4392 | *) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4393 | |
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4394 | subsection\<open>Finally, spheres are ANRs and ENRs\<close> | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4395 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4396 | lemma absolute_retract_homeomorphic_convex_compact: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4397 | 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 | 4398 |   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 | 4399 | shows "S retract_of T" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4400 | 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 | 4401 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4402 | lemma frontier_retract_of_punctured_universe: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4403 | fixes S :: "'a::euclidean_space set" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4404 | assumes "convex S" "bounded S" "a \<in> interior S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4405 |   shows "(frontier S) retract_of (- {a})"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4406 | using rel_frontier_retract_of_punctured_affine_hull | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4407 | 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 | 4408 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4409 | lemma sphere_retract_of_punctured_universe_gen: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4410 | fixes a :: "'a::euclidean_space" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4411 | assumes "b \<in> ball a r" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4412 |   shows  "sphere a r retract_of (- {b})"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4413 | proof - | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4414 |   have "frontier (cball a r) retract_of (- {b})"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4415 | apply (rule frontier_retract_of_punctured_universe) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4416 | using assms by auto | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4417 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4418 | by simp | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4419 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4420 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4421 | lemma sphere_retract_of_punctured_universe: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4422 | fixes a :: "'a::euclidean_space" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4423 | assumes "0 < r" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4424 |   shows "sphere a r retract_of (- {a})"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4425 | 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 | 4426 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4427 | proposition ENR_sphere: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4428 | fixes a :: "'a::euclidean_space" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4429 | shows "ENR(sphere a r)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4430 | proof (cases "0 < r") | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4431 | case True | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4432 |   then have "sphere a r retract_of -{a}"
 | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4433 | by (simp add: sphere_retract_of_punctured_universe) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4434 | with open_delete show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4435 | by (auto simp: ENR_def) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4436 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4437 | case False | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4438 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4439 | using finite_imp_ENR | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4440 | 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 | 4441 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4442 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4443 | corollary ANR_sphere: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4444 | fixes a :: "'a::euclidean_space" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4445 | shows "ANR(sphere a r)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4446 | by (simp add: ENR_imp_ANR ENR_sphere) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4447 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4448 | |
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4449 | subsection\<open>Spheres are connected, etc.\<close> | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4450 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4451 | 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 | 4452 | 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 | 4453 | 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 | 4454 | 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 | 4455 | 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 | 4456 | case True | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4457 | with assms show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4458 | 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 | 4459 | next | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4460 | case False | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4461 | 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 | 4462 | by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4463 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4464 | 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 | 4465 |     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 | 4466 | 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 | 4467 |     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 | 4468 | 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 | 4469 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4470 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4471 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4472 | 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 | 4473 | 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 | 4474 | 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 | 4475 | 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 | 4476 | 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 | 4477 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4478 | 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 | 4479 | 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 | 4480 | 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 | 4481 | 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 | 4482 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4483 | lemma locally_connected_sphere: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4484 | 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 | 4485 | 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 | 4486 | 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 | 4487 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4488 | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4489 | subsection\<open>Borsuk homotopy extension theorem\<close> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4490 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4491 | 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 | 4492 | 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 | 4493 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4494 | theorem Borsuk_homotopy_extension_homotopic: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4495 | 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 | 4496 | assumes cloTS: "closedin (subtopology euclidean T) S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4497 | 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 | 4498 | and contf: "continuous_on T f" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4499 | and "f ` T \<subseteq> U" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4500 | 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 | 4501 | 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 | 4502 | "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 | 4503 | "\<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 | 4504 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4505 | 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 | 4506 |   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 | 4507 |              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 | 4508 | 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 | 4509 | using assms by (auto simp: homotopic_with_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4510 | define h' where "h' \<equiv> \<lambda>z. if snd z \<in> S then h z else (f o snd) z" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4511 |   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 | 4512 |   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 | 4513 | by (simp add: closedin_subtopology_refl closedin_Times) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4514 |   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 | 4515 | 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 | 4516 |   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 | 4517 | by (auto simp: B_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4518 |   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 | 4519 | 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 | 4520 |   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 | 4521 | 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 | 4522 | 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 | 4523 |   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 | 4524 | apply (rule continuous_intros)+ | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4525 | apply (simp add: contf) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4526 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4527 | ultimately have conth': "continuous_on B h'" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4528 |     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 | 4529 | apply (auto intro!: continuous_on_cases_local conth) | 
| 
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 | have "image h' B \<subseteq> U" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4532 | 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 | 4533 |   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 | 4534 | 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 | 4535 | 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 | 4536 | using anr | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4537 | proof | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4538 | assume ST: "ANR S \<and> ANR T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4539 |     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 | 4540 | using \<open>S \<subseteq> T\<close> by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4541 | have "ANR B" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4542 | apply (simp add: B_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4543 | apply (rule ANR_closed_Un_local) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4544 | apply (metis cloBT B_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4545 | apply (metis Un_commute cloBS B_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4546 | 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 | 4547 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4548 | note Vk = that | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4549 |     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 | 4550 | "retraction V B r" for V r | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4551 | using that | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4552 | apply (clarsimp simp add: retraction_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4553 | apply (rule Vk [of V "h' o r"], assumption+) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4554 | 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 | 4555 | 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 | 4556 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4557 | show thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4558 | 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 | 4559 | 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 | 4560 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4561 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4562 | assume "ANR U" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4563 | 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 | 4564 | show ?thesis by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4565 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4566 |   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 | 4567 | have "closedin (subtopology euclidean T) S'" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4568 | unfolding S'_def | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4569 | apply (rule closedin_compact_projection, blast) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4570 | using closedin_self opeTV by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4571 |   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 | 4572 | by (auto simp: S'_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4573 | have cloTS': "closedin (subtopology euclidean T) S'" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4574 | 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 | 4575 |   have "S \<inter> S' = {}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4576 | 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 | 4577 | 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 | 4578 | 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 | 4579 | 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 | 4580 | 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 | 4581 |     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 | 4582 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4583 |   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 | 4584 | using closed_segment_eq_real_ivl by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4585 | 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 | 4586 | proof (rule ccontr) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4587 | assume "(u * a t, t) \<notin> V" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4588 | 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 | 4589 | apply simp | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4590 | apply (rule a0) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4591 | 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 | 4592 | show False | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4593 | 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 | 4594 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4595 | show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4596 | proof | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4597 | 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 | 4598 | proof (simp add: homotopic_with, intro exI conjI) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4599 |       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 | 4600 | apply (intro continuous_on_compose continuous_intros) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4601 | apply (rule continuous_on_subset [OF conta], force) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4602 | apply (rule continuous_on_subset [OF contk]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4603 | apply (force intro: inV) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4604 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4605 |       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 | 4606 | using inV kim by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4607 | 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 | 4608 | by (simp add: B_def h'_def keq) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4609 | 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 | 4610 | by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4611 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4612 | 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 | 4613 | using hom homotopic_with_imp_continuous by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4614 | 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 | 4615 | proof clarify | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4616 | fix t | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4617 | assume "t \<in> T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4618 | show "k (a t, t) \<in> U" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4619 | 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 | 4620 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4621 | 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 | 4622 | 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 | 4623 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4624 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4625 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4626 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4627 | corollary nullhomotopic_into_ANR_extension: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4628 | 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 | 4629 | assumes "closed S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4630 | and contf: "continuous_on S f" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4631 | and "ANR T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4632 | and fim: "f ` S \<subseteq> T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4633 |       and "S \<noteq> {}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4634 | 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 | 4635 | (\<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 | 4636 | (is "?lhs = ?rhs") | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4637 | proof | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4638 | assume ?lhs | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4639 | then obtain c where c: "homotopic_with (\<lambda>x. True) S T (\<lambda>x. c) f" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4640 | by (blast intro: homotopic_with_symD elim: ) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4641 | have "closedin (subtopology euclidean UNIV) S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4642 | 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 | 4643 | 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 | 4644 | "\<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 | 4645 | 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 | 4646 |     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 | 4647 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4648 | then show ?rhs by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4649 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4650 | assume ?rhs | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4651 | 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 | 4652 | by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4653 | 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 | 4654 | 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 | 4655 | then show ?lhs | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4656 | apply (rule_tac x="c" in exI) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4657 | 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 | 4658 | apply (rule homotopic_with_subset_left) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4659 | apply (auto simp add: \<open>\<And>x. x \<in> S \<Longrightarrow> g x = f x\<close>) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4660 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4661 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4662 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4663 | corollary nullhomotopic_into_rel_frontier_extension: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4664 | 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 | 4665 | assumes "closed S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4666 | and contf: "continuous_on S f" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4667 | and "convex T" "bounded T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4668 | and fim: "f ` S \<subseteq> rel_frontier T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4669 |       and "S \<noteq> {}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4670 | 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 | 4671 | (\<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 | 4672 | 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 | 4673 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4674 | corollary nullhomotopic_into_sphere_extension: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4675 | 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 | 4676 | 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 | 4677 |       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 | 4678 | 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 | 4679 | (\<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 | 4680 | (is "?lhs = ?rhs") | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4681 | proof (cases "r = 0") | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4682 | case True with fim show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4683 | apply (auto simp: ) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4684 | using fim continuous_on_const apply fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4685 | by (metis contf contractible_sing nullhomotopic_into_contractible) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4686 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4687 | case False | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4688 | 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 | 4689 | show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4690 | using fim unfolding eq | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4691 | 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 | 4692 |     apply (rule \<open>S \<noteq> {}\<close>)
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4693 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4694 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4695 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4696 | proposition Borsuk_map_essential_bounded_component: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4697 | fixes a :: "'a :: euclidean_space" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4698 | assumes "compact S" and "a \<notin> S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4699 | shows "bounded (connected_component_set (- S) a) \<longleftrightarrow> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4700 | ~(\<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 | 4701 | (\<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 | 4702 | (is "?lhs = ?rhs") | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4703 | proof (cases "S = {}")
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4704 | case True then show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4705 | by simp | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4706 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4707 | case False | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4708 | have "closed S" "bounded S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4709 | 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 | 4710 | 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 | 4711 | 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 | 4712 | 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 | 4713 | by (simp add: \<open>a \<notin> S\<close>) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4714 | 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 | 4715 | 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 | 4716 | have "~ ?rhs \<longleftrightarrow> ~ ?lhs" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4717 | proof | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4718 | assume notr: "~ ?rhs" | 
| 63497 
ef794d2e3754
HOL-Multivariate_Analysis: add amssymb to document generation; reintroduce \nexists (cf d51a0a772094)
 hoelzl parents: 
63493diff
changeset | 4719 | 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 | 4720 | 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 | 4721 | (\<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 | 4722 | if "bounded (connected_component_set (- S) a)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4723 | 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 | 4724 | 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 | 4725 | 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 | 4726 | "\<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 | 4727 | using notr | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4728 | by (auto simp add: nullhomotopic_into_sphere_extension | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4729 | [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 | 4730 | with \<open>a \<notin> S\<close> show "~ ?lhs" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4731 | apply (clarsimp simp: Borsuk_map_into_sphere [of a S, symmetric] dest!: nog) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4732 | apply (drule_tac x="g" in spec) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4733 | using continuous_on_subset by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4734 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4735 | assume "~ ?lhs" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4736 | 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 | 4737 | using bounded_iff linear by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4738 | then have bnot: "b \<notin> ball 0 r" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4739 | by simp | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4740 | 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 | 4741 | (\<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 | 4742 | apply (rule Borsuk_maps_homotopic_in_path_component) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4743 | 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 | 4744 | done | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4745 | moreover | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4746 | 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 | 4747 | (\<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 | 4748 | proof (rule nullhomotopic_from_contractible) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4749 | show "contractible (ball (0::'a) r)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4750 | by (metis convex_imp_contractible convex_ball) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4751 | 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 | 4752 | by (rule continuous_on_Borsuk_map [OF bnot]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4753 | 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 | 4754 | using bnot Borsuk_map_into_sphere by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4755 | qed blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4756 | 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 | 4757 | (\<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 | 4758 | 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 | 4759 | then show "~ ?rhs" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4760 | by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4761 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4762 | then show ?thesis by blast | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4763 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4764 | |
| 64791 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4765 | lemma homotopic_Borsuk_maps_in_bounded_component: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4766 | fixes a :: "'a :: euclidean_space" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4767 | 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 | 4768 | and boc: "bounded (connected_component_set (- S) a)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4769 | 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 | 4770 | (\<lambda>x. (x - a) /\<^sub>R norm (x - a)) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4771 | (\<lambda>x. (x - b) /\<^sub>R norm (x - b))" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4772 | shows "connected_component (- S) a b" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4773 | proof (rule ccontr) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4774 | assume notcc: "\<not> connected_component (- S) a b" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4775 | let ?T = "S \<union> connected_component_set (- S) a" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4776 | 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 | 4777 | 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 | 4778 | (\<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 | 4779 | 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 | 4780 | 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 | 4781 | "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 | 4782 | "\<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 | 4783 | proof (rule Borsuk_homotopy_extension_homotopic) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4784 | show "closedin (subtopology euclidean ?T) S" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4785 | 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 | 4786 | 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 | 4787 | 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 | 4788 | 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 | 4789 | 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 | 4790 | 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 | 4791 | (\<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 | 4792 | by (simp add: hom homotopic_with_symD) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4793 | qed (auto simp: ANR_sphere intro: that) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4794 | ultimately show False by blast | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4795 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4796 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4797 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4798 | lemma Borsuk_maps_homotopic_in_connected_component_eq: | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4799 | fixes a :: "'a :: euclidean_space" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4800 |   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 | 4801 | 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 | 4802 | (\<lambda>x. (x - a) /\<^sub>R norm (x - a)) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4803 | (\<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 | 4804 | connected_component (- S) a b)" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4805 | (is "?lhs = ?rhs") | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4806 | proof | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4807 | assume L: ?lhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4808 | show ?rhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4809 | proof (cases "bounded(connected_component_set (- S) a)") | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4810 | case True | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4811 | show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4812 | 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 | 4813 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4814 | case not_bo_a: False | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4815 | show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4816 | proof (cases "bounded(connected_component_set (- S) b)") | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4817 | case True | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4818 | show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4819 | using homotopic_Borsuk_maps_in_bounded_component [OF S] | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4820 | 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 | 4821 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4822 | case False | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4823 | then show ?thesis | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4824 | 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 | 4825 | 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 | 4826 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4827 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4828 | next | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4829 | assume R: ?rhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4830 | then have "path_component (- S) a b" | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4831 | 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 | 4832 | then show ?lhs | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4833 | by (simp add: Borsuk_maps_homotopic_in_path_component) | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4834 | qed | 
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4835 | |
| 
05a2b3b20664
facts about ANRs, ENRs, covering spaces
 paulson <lp15@cam.ac.uk> parents: 
64789diff
changeset | 4836 | |
| 66939 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4837 | subsection\<open>More extension theorems\<close> | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4838 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4839 | lemma extension_from_clopen: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4840 | 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 | 4841 | 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 | 4842 |       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 | 4843 | 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 | 4844 | proof (cases "U = {}")
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4845 | case True | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4846 | then show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4847 | 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 | 4848 | next | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4849 | case False | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4850 | 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 | 4851 | by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4852 | 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 | 4853 | 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 | 4854 | 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 | 4855 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4856 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4857 | 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 | 4858 | 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 | 4859 | 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 | 4860 | 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 | 4861 | by metis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4862 | 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 | 4863 | 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 | 4864 | 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 | 4865 | by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4866 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4867 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4868 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4869 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4870 | lemma extension_from_component: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4871 | 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 | 4872 | 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 | 4873 | 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 | 4874 | 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 | 4875 | proof - | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4876 | 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 | 4877 | 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 | 4878 | 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 | 4879 | 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 | 4880 | using S | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4881 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4882 | assume "locally connected S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4883 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4884 | 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 | 4885 | next | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4886 | assume "compact S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4887 | 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 | 4888 | 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 | 4889 | 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 | 4890 | 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 | 4891 | 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 | 4892 | 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 | 4893 | 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 | 4894 | 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 | 4895 | 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 | 4896 | 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 | 4897 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4898 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4899 | 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 | 4900 | 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 | 4901 | 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 | 4902 | 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 | 4903 | 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 | 4904 | 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 | 4905 | 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 | 4906 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4907 | 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 | 4908 | using extension_from_clopen | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4909 | 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 | 4910 | then show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4911 | 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 | 4912 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4913 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4914 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4915 | lemma tube_lemma: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4916 | 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 | 4917 |   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 | 4918 | 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 | 4919 | 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 | 4920 | proof - | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4921 |   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 | 4922 | 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 | 4923 | 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 | 4924 | 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 | 4925 | 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 | 4926 | 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 | 4927 | 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 | 4928 | 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 | 4929 | by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4930 | ultimately show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4931 | 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 | 4932 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4933 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4934 | lemma tube_lemma_gen: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4935 | 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 | 4936 |   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 | 4937 | 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 | 4938 | 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 | 4939 | proof - | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4940 | 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 | 4941 | 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 | 4942 | 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 | 4943 | by metis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4944 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4945 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4946 | 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 | 4947 | using F by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4948 | 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 | 4949 | using F by blast | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4950 | 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 | 4951 | using F by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4952 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4953 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4954 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4955 | proposition homotopic_neighbourhood_extension: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4956 | 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 | 4957 | 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 | 4958 | 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 | 4959 | 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 | 4960 | 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 | 4961 | 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 | 4962 | "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 | 4963 | proof - | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4964 | have "T \<subseteq> S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4965 | 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 | 4966 |   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 | 4967 |              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 | 4968 | 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 | 4969 | 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 | 4970 |   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 | 4971 |                              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 | 4972 | else h z" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4973 |   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 | 4974 |   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 | 4975 | unfolding h'_def | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4976 | 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 | 4977 |     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 | 4978 |          "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 | 4979 |       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 | 4980 |     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 | 4981 |          "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 | 4982 |       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 | 4983 | 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 | 4984 | 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 | 4985 | 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 | 4986 | 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 | 4987 | 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 | 4988 |   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 | 4989 | 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 | 4990 |   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 | 4991 | 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 | 4992 |   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 | 4993 | 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 | 4994 | ultimately | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 4995 |   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 | 4996 |                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 | 4997 | 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 | 4998 | 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 | 4999 |                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 | 5000 |     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 | 5001 | 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 | 5002 |               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 | 5003 |     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 | 5004 | 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 | 5005 | 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 | 5006 |     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 | 5007 | 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 | 5008 |     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 | 5009 | 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 | 5010 | have "T' \<subseteq> S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5011 | 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 | 5012 | 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 | 5013 | 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 | 5014 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5015 | ultimately show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5016 | by (blast intro: that) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5017 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5018 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5019 | text\<open> Homotopy on a union of closed-open sets.\<close> | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5020 | proposition homotopic_on_clopen_Union: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5021 | 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 | 5022 | 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 | 5023 | 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 | 5024 | 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 | 5025 | 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 | 5026 | proof - | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5027 | 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 | 5028 | 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 | 5029 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5030 |   proof (cases "\<V> = {}")
 | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5031 | case True | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5032 | then show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5033 | 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 | 5034 | next | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5035 | case False | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5036 | 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 | 5037 | 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 | 5038 | 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 | 5039 | 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 | 5040 | 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 | 5041 | using assms by auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5042 |     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 | 5043 |                   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 | 5044 | 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 | 5045 | 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 | 5046 | 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 | 5047 | 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 | 5048 | 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 | 5049 |     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 | 5050 |               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 | 5051 |                                    {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 | 5052 | 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 | 5053 |       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 | 5054 | 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 | 5055 |       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 | 5056 |                    ({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 | 5057 | 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 | 5058 | 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 | 5059 | 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 | 5060 | 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 | 5061 | 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 | 5062 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5063 |       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 | 5064 | 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 | 5065 |       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 | 5066 |                     {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 | 5067 | \<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 | 5068 | 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 | 5069 | qed auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5070 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5071 | 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 | 5072 |       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 | 5073 | 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 | 5074 |       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 | 5075 | proof clarsimp | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5076 | 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 | 5077 | 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 | 5078 | 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 | 5079 | 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 | 5080 | with him t show "\<zeta>(t, y) \<in> T" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5081 | by (subst eq) (force simp:)+ | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5082 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5083 | fix X y | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5084 | 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 | 5085 | 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 | 5086 | 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 | 5087 | 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 | 5088 | 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 | 5089 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5090 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5091 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5092 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5093 | proposition homotopic_on_components_eq: | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5094 | 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 | 5095 | 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 | 5096 | 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 | 5097 | (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 | 5098 | (\<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 | 5099 | (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 | 5100 | proof - | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5101 | 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 | 5102 | 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 | 5103 | 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 | 5104 | 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 | 5105 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5106 | assume ?lhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5107 | with that show ?rhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5108 | 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 | 5109 | next | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5110 | assume R: ?rhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5111 | 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 | 5112 | 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 | 5113 | 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 | 5114 | proof - | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5115 | have "C \<subseteq> S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5116 | 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 | 5117 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5118 | using S | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5119 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5120 | assume "locally connected S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5121 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5122 | proof (intro exI conjI) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5123 | 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 | 5124 | 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 | 5125 | 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 | 5126 | 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 | 5127 | 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 | 5128 | 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 | 5129 | qed auto | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5130 | next | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5131 | assume "compact S" | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5132 | 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 | 5133 | 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 | 5134 | 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 | 5135 | 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 | 5136 | 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 | 5137 | \<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 | 5138 | 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 | 5139 | 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 | 5140 | 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 | 5141 | 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 | 5142 | 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 | 5143 | 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 | 5144 | show ?thesis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5145 | proof (intro exI conjI) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5146 | 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 | 5147 | 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 | 5148 | 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 | 5149 | 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 | 5150 | 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 | 5151 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5152 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5153 | 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 | 5154 | 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 | 5155 | 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 | 5156 | 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 | 5157 | by metis | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5158 | 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 | 5159 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5160 | 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 | 5161 | 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 | 5162 | 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 | 5163 | 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 | 5164 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5165 | show ?lhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5166 | apply (subst Seq) | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5167 | 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 | 5168 | 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 | 5169 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5170 | 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 | 5171 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5172 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5173 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5174 | 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 | 5175 | 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 | 5176 | 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 | 5177 | shows | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5178 | "(\<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 | 5179 | 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 | 5180 | \<longleftrightarrow> | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5181 | (\<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 | 5182 | \<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 | 5183 | 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 | 5184 | (is "?lhs = ?rhs") | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5185 | proof | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5186 | 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 | 5187 | show ?rhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5188 | proof clarify | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5189 | fix C f g | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5190 | 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 | 5191 | 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 | 5192 | 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 | 5193 | 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 | 5194 | 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 | 5195 | 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 | 5196 | 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 | 5197 | 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 | 5198 | 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 | 5199 | 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 | 5200 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5201 | next | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5202 | 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 | 5203 | show ?lhs | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5204 | proof clarify | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5205 | fix f g | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5206 | 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 | 5207 | 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 | 5208 | 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 | 5209 | using R [OF that] | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5210 | 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 | 5211 | 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 | 5212 | 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 | 5213 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5214 | qed | 
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5215 | |
| 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 5216 | |
| 64122 | 5217 | subsection\<open>The complement of a set and path-connectedness\<close> | 
| 5218 | ||
| 5219 | text\<open>Complement in dimension N > 1 of set homeomorphic to any interval in | |
| 5220 | any dimension is (path-)connected. This naively generalizes the argument | |
| 5221 | in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer fixed point theorem", | |
| 5222 | American Mathematical Monthly 1984.\<close> | |
| 5223 | ||
| 5224 | lemma unbounded_components_complement_absolute_retract: | |
| 5225 | fixes S :: "'a::euclidean_space set" | |
| 5226 | assumes C: "C \<in> components(- S)" and S: "compact S" "AR S" | |
| 5227 | shows "\<not> bounded C" | |
| 5228 | proof - | |
| 5229 | obtain y where y: "C = connected_component_set (- S) y" and "y \<notin> S" | |
| 5230 | using C by (auto simp: components_def) | |
| 5231 | have "open(- S)" | |
| 5232 | using S by (simp add: closed_open compact_eq_bounded_closed) | |
| 5233 | have "S retract_of UNIV" | |
| 5234 | using S compact_AR by blast | |
| 5235 | then obtain r where contr: "continuous_on UNIV r" and ontor: "range r \<subseteq> S" | |
| 5236 | and r: "\<And>x. x \<in> S \<Longrightarrow> r x = x" | |
| 5237 | by (auto simp: retract_of_def retraction_def) | |
| 5238 | show ?thesis | |
| 5239 | proof | |
| 5240 | assume "bounded C" | |
| 5241 | have "connected_component_set (- S) y \<subseteq> S" | |
| 5242 | proof (rule frontier_subset_retraction) | |
| 5243 | show "bounded (connected_component_set (- S) y)" | |
| 5244 | using \<open>bounded C\<close> y by blast | |
| 5245 | show "frontier (connected_component_set (- S) y) \<subseteq> S" | |
| 5246 | using C \<open>compact S\<close> compact_eq_bounded_closed frontier_of_components_closed_complement y by blast | |
| 5247 | show "continuous_on (closure (connected_component_set (- S) y)) r" | |
| 5248 | by (blast intro: continuous_on_subset [OF contr]) | |
| 5249 | qed (use ontor r in auto) | |
| 5250 | with \<open>y \<notin> S\<close> show False by force | |
| 5251 | qed | |
| 5252 | qed | |
| 5253 | ||
| 5254 | lemma connected_complement_absolute_retract: | |
| 5255 | fixes S :: "'a::euclidean_space set" | |
| 5256 |   assumes S: "compact S" "AR S" and 2: "2 \<le> DIM('a)"
 | |
| 5257 | shows "connected(- S)" | |
| 5258 | proof - | |
| 5259 | have "S retract_of UNIV" | |
| 5260 | using S compact_AR by blast | |
| 5261 | show ?thesis | |
| 5262 | apply (clarsimp simp: connected_iff_connected_component_eq) | |
| 5263 | apply (rule cobounded_unique_unbounded_component [OF _ 2]) | |
| 5264 | apply (simp add: \<open>compact S\<close> compact_imp_bounded) | |
| 5265 | apply (meson ComplI S componentsI unbounded_components_complement_absolute_retract)+ | |
| 5266 | done | |
| 5267 | qed | |
| 5268 | ||
| 5269 | lemma path_connected_complement_absolute_retract: | |
| 5270 | fixes S :: "'a::euclidean_space set" | |
| 5271 |   assumes "compact S" "AR S" "2 \<le> DIM('a)"
 | |
| 5272 | shows "path_connected(- S)" | |
| 5273 | using connected_complement_absolute_retract [OF assms] | |
| 5274 | using \<open>compact S\<close> compact_eq_bounded_closed connected_open_path_connected by blast | |
| 5275 | ||
| 5276 | theorem connected_complement_homeomorphic_convex_compact: | |
| 5277 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | |
| 5278 |   assumes hom: "S homeomorphic T" and T: "convex T" "compact T" and 2: "2 \<le> DIM('a)"
 | |
| 5279 | shows "connected(- S)" | |
| 5280 | proof (cases "S = {}")
 | |
| 5281 | case True | |
| 5282 | then show ?thesis | |
| 5283 | by (simp add: connected_UNIV) | |
| 5284 | next | |
| 5285 | case False | |
| 5286 | show ?thesis | |
| 5287 | proof (rule connected_complement_absolute_retract) | |
| 5288 | show "compact S" | |
| 5289 | using \<open>compact T\<close> hom homeomorphic_compactness by auto | |
| 5290 | show "AR S" | |
| 5291 | by (meson AR_ANR False \<open>convex T\<close> convex_imp_ANR convex_imp_contractible hom homeomorphic_ANR_iff_ANR homeomorphic_contractible_eq) | |
| 5292 | qed (rule 2) | |
| 5293 | qed | |
| 5294 | ||
| 5295 | corollary path_connected_complement_homeomorphic_convex_compact: | |
| 5296 | fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" | |
| 5297 |   assumes hom: "S homeomorphic T" "convex T" "compact T" "2 \<le> DIM('a)"
 | |
| 5298 | shows "path_connected(- S)" | |
| 5299 | using connected_complement_homeomorphic_convex_compact [OF assms] | |
| 5300 | 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 | 5301 | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 5302 | end |