| author | hoelzl | 
| Tue, 05 Nov 2013 09:44:59 +0100 | |
| changeset 54260 | 6a967667fd45 | 
| parent 53846 | 2e4b435e17bc | 
| child 54775 | 2d3df8633dad | 
| permissions | -rw-r--r-- | 
| 53674 | 1 | (* Author: John Harrison | 
| 2 | Author: Robert Himmelmann, TU Muenchen (Translation from HOL light) | |
| 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 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 19 | header {* 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 | 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 | 
| 53674 | 22 | imports Convex_Euclidean_Space | 
| 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 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 25 | (** move this **) | 
| 53185 | 26 | lemma divide_nonneg_nonneg: | 
| 53846 | 27 | fixes a b :: real | 
| 53674 | 28 | assumes "a \<ge> 0" | 
| 29 | and "b \<ge> 0" | |
| 53846 | 30 | shows "0 \<le> a / b" | 
| 53674 | 31 | proof (cases "b = 0") | 
| 32 | case True | |
| 33 | then show ?thesis by auto | |
| 34 | next | |
| 35 | case False | |
| 36 | show ?thesis | |
| 37 | apply (rule divide_nonneg_pos) | |
| 38 | using assms False | |
| 39 | apply auto | |
| 40 | done | |
| 41 | qed | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 42 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 43 | lemma brouwer_compactness_lemma: | 
| 50884 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 hoelzl parents: 
50526diff
changeset | 44 | fixes f :: "'a::metric_space \<Rightarrow> 'b::euclidean_space" | 
| 53674 | 45 | assumes "compact s" | 
| 46 | and "continuous_on s f" | |
| 53688 | 47 | and "\<not> (\<exists>x\<in>s. f x = 0)" | 
| 53674 | 48 | obtains d where "0 < d" and "\<forall>x\<in>s. d \<le> norm (f x)" | 
| 53185 | 49 | proof (cases "s = {}")
 | 
| 53674 | 50 | case True | 
| 53688 | 51 | show thesis | 
| 52 | by (rule that [of 1]) (auto simp: True) | |
| 53674 | 53 | next | 
| 49374 | 54 | case False | 
| 55 | have "continuous_on s (norm \<circ> f)" | |
| 49555 | 56 | by (rule continuous_on_intros continuous_on_norm assms(2))+ | 
| 53674 | 57 | with False obtain x where x: "x \<in> s" "\<forall>y\<in>s. (norm \<circ> f) x \<le> (norm \<circ> f) y" | 
| 58 | using continuous_attains_inf[OF assms(1), of "norm \<circ> f"] | |
| 59 | unfolding o_def | |
| 60 | by auto | |
| 61 | have "(norm \<circ> f) x > 0" | |
| 62 | using assms(3) and x(1) | |
| 63 | by auto | |
| 64 | then show ?thesis | |
| 65 | by (rule that) (insert x(2), auto simp: o_def) | |
| 49555 | 66 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 67 | |
| 49555 | 68 | 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 | 69 | fixes P Q :: "'a::euclidean_space \<Rightarrow> bool" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 70 | assumes "(\<forall>x. P x \<longrightarrow> P (f x))" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 71 | 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 | 72 | 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 | 73 | (\<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 | 74 | (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (x\<bullet>i = 1) \<longrightarrow> (l x i = 1)) \<and> | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 75 | (\<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> | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 76 | (\<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 | 77 | proof - | 
| 78 | have and_forall_thm:"\<And>P Q. (\<forall>x. P x) \<and> (\<forall>x. Q x) \<longleftrightarrow> (\<forall>x. P x \<and> Q x)" | |
| 79 | by auto | |
| 49555 | 80 | 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)" | 
| 49374 | 81 | by auto | 
| 82 | show ?thesis | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 83 | unfolding and_forall_thm Ball_def | 
| 53185 | 84 | apply (subst choice_iff[symmetric])+ | 
| 49374 | 85 | apply rule | 
| 86 | apply rule | |
| 87 | proof - | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 88 | case (goal1 x) | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 89 | let ?R = "\<lambda>y. (P x \<and> Q xa \<and> x \<bullet> xa = 0 \<longrightarrow> y = (0::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 | 90 | (P x \<and> Q xa \<and> x \<bullet> xa = 1 \<longrightarrow> y = 1) \<and> | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 91 | (P x \<and> Q xa \<and> y = 0 \<longrightarrow> x \<bullet> xa \<le> f x \<bullet> xa) \<and> | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 92 | (P x \<and> Q xa \<and> y = 1 \<longrightarrow> f x \<bullet> xa \<le> x \<bullet> xa)" | 
| 49374 | 93 |     {
 | 
| 53688 | 94 | assume "P x" "Q xa" "xa \<in> 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 | 95 | then have "0 \<le> f x \<bullet> xa \<and> f x \<bullet> xa \<le> 1" | 
| 49374 | 96 | using assms(2)[rule_format,of "f x" xa] | 
| 97 | apply (drule_tac assms(1)[rule_format]) | |
| 98 | apply auto | |
| 99 | done | |
| 100 | } | |
| 53688 | 101 | then have "xa \<in> Basis \<Longrightarrow> ?R 0 \<or> ?R 1" by auto | 
| 49374 | 102 | then show ?case by auto | 
| 103 | qed | |
| 104 | qed | |
| 105 | ||
| 53185 | 106 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 107 | subsection {* The key "counting" observation, somewhat abstracted. *}
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 108 | |
| 49374 | 109 | lemma setsum_Un_disjoint': | 
| 53688 | 110 | assumes "finite A" | 
| 111 | and "finite B" | |
| 112 |     and "A \<inter> B = {}"
 | |
| 113 | and "A \<union> B = C" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 114 | shows "setsum g C = setsum g A + setsum g B" | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 115 | using setsum_Un_disjoint[OF assms(1-3)] and assms(4) by auto | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 116 | |
| 53252 | 117 | lemma kuhn_counting_lemma: | 
| 53688 | 118 | assumes "finite faces" | 
| 119 | and "finite simplices" | |
| 120 |     and "\<forall>f\<in>faces. bnd f \<longrightarrow> (card {s \<in> simplices. face f s} = 1)"
 | |
| 121 |     and "\<forall>f\<in>faces. \<not> bnd f \<longrightarrow> (card {s \<in> simplices. face f s} = 2)"
 | |
| 122 |     and "\<forall>s\<in>simplices. compo s \<longrightarrow> (card {f \<in> faces. face f s \<and> compo' f} = 1)"
 | |
| 123 | and "\<forall>s\<in>simplices. \<not> compo s \<longrightarrow> | |
| 53252 | 124 |       (card {f \<in> faces. face f s \<and> compo' f} = 0) \<or> (card {f \<in> faces. face f s \<and> compo' f} = 2)"
 | 
| 53688 | 125 |     and "odd(card {f \<in> faces. compo' f \<and> bnd f})"
 | 
| 49374 | 126 |   shows "odd(card {s \<in> simplices. compo s})"
 | 
| 127 | proof - | |
| 128 |   have "\<And>x. {f\<in>faces. compo' f \<and> bnd f \<and> face f x} \<union> {f\<in>faces. compo' f \<and> \<not>bnd f \<and> face f x} =
 | |
| 129 |       {f\<in>faces. compo' f \<and> face f x}"
 | |
| 130 |     "\<And>x. {f \<in> faces. compo' f \<and> bnd f \<and> face f x} \<inter> {f \<in> faces. compo' f \<and> \<not> bnd f \<and> face f x} = {}"
 | |
| 131 | by auto | |
| 53688 | 132 |   then have lem1: "setsum (\<lambda>s. (card {f \<in> faces. face f s \<and> compo' f})) simplices =
 | 
| 49374 | 133 |       setsum (\<lambda>s. card {f \<in> {f \<in> faces. compo' f \<and> bnd f}. face f s}) simplices +
 | 
| 134 |       setsum (\<lambda>s. card {f \<in> {f \<in> faces. compo' f \<and> \<not> (bnd f)}. face f s}) simplices"
 | |
| 53185 | 135 | unfolding setsum_addf[symmetric] | 
| 49374 | 136 | apply - | 
| 53688 | 137 | apply (rule setsum_cong2) | 
| 49374 | 138 | using assms(1) | 
| 139 | apply (auto simp add: card_Un_Int, auto simp add:conj_commute) | |
| 140 | done | |
| 53252 | 141 | have lem2: | 
| 142 |     "setsum (\<lambda>j. card {f \<in> {f \<in> faces. compo' f \<and> bnd f}. face f j}) simplices =
 | |
| 143 |       1 * card {f \<in> faces. compo' f \<and> bnd f}"
 | |
| 144 |     "setsum (\<lambda>j. card {f \<in> {f \<in> faces. compo' f \<and> \<not> bnd f}. face f j}) simplices =
 | |
| 145 |       2 * card {f \<in> faces. compo' f \<and> \<not> bnd f}"
 | |
| 53688 | 146 | apply (rule_tac[!] setsum_multicount) | 
| 49374 | 147 | using assms | 
| 148 | apply auto | |
| 149 | done | |
| 53252 | 150 | have lem3: | 
| 151 |     "setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f}) simplices =
 | |
| 152 |       setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f}) {s \<in> simplices.   compo s}+
 | |
| 153 |       setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f}) {s \<in> simplices. \<not> compo s}"
 | |
| 154 | apply (rule setsum_Un_disjoint') | |
| 155 | using assms(2) | |
| 156 | apply auto | |
| 157 | done | |
| 158 |   have lem4: "setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f}) {s \<in> simplices. compo s} =
 | |
| 159 |     setsum (\<lambda>s. 1) {s \<in> simplices. compo s}"
 | |
| 160 | apply (rule setsum_cong2) | |
| 161 | using assms(5) | |
| 162 | apply auto | |
| 163 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 164 |   have lem5: "setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f}) {s \<in> simplices. \<not> compo s} =
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 165 |     setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f})
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 166 |            {s \<in> simplices. (\<not> compo s) \<and> (card {f \<in> faces. face f s \<and> compo' f} = 0)} +
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 167 |     setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f})
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 168 |            {s \<in> simplices. (\<not> compo s) \<and> (card {f \<in> faces. face f s \<and> compo' f} = 2)}"
 | 
| 53252 | 169 | apply (rule setsum_Un_disjoint') | 
| 170 | using assms(2,6) | |
| 171 | apply auto | |
| 172 | done | |
| 173 |   have *: "int (\<Sum>s\<in>{s \<in> simplices. compo s}. card {f \<in> faces. face f s \<and> compo' f}) =
 | |
| 53185 | 174 |     int (card {f \<in> faces. compo' f \<and> bnd f} + 2 * card {f \<in> faces. compo' f \<and> \<not> bnd f}) -
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 175 |     int (card {s \<in> simplices. \<not> compo s \<and> card {f \<in> faces. face f s \<and> compo' f} = 2} * 2)"
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 176 | using lem1[unfolded lem3 lem2 lem5] by auto | 
| 49374 | 177 | have even_minus_odd:"\<And>x y. even x \<Longrightarrow> odd (y::int) \<Longrightarrow> odd (x - y)" | 
| 178 | using assms by auto | |
| 179 | have odd_minus_even:"\<And>x y. odd x \<Longrightarrow> even (y::int) \<Longrightarrow> odd (x - y)" | |
| 180 | using assms by auto | |
| 181 | show ?thesis | |
| 53185 | 182 | unfolding even_nat_def card_eq_setsum and lem4[symmetric] and *[unfolded card_eq_setsum] | 
| 183 | unfolding card_eq_setsum[symmetric] | |
| 49374 | 184 | apply (rule odd_minus_even) | 
| 185 | unfolding of_nat_add | |
| 186 | apply(rule odd_plus_even) | |
| 187 | apply(rule assms(7)[unfolded even_nat_def]) | |
| 188 | unfolding int_mult | |
| 189 | apply auto | |
| 190 | done | |
| 191 | qed | |
| 192 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 193 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 194 | subsection {* The odd/even result for faces of complete vertices, generalized. *}
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 195 | |
| 49374 | 196 | lemma card_1_exists: "card s = 1 \<longleftrightarrow> (\<exists>!x. x \<in> s)" | 
| 197 | unfolding One_nat_def | |
| 198 | apply rule | |
| 199 | apply (drule card_eq_SucD) | |
| 200 | defer | |
| 49555 | 201 | apply (erule ex1E) | 
| 49374 | 202 | proof - | 
| 53186 | 203 | fix x | 
| 204 | assume as: "x \<in> s" "\<forall>y. y \<in> s \<longrightarrow> y = x" | |
| 49374 | 205 |   have *: "s = insert x {}"
 | 
| 49555 | 206 | apply (rule set_eqI, rule) unfolding singleton_iff | 
| 49374 | 207 | apply (rule as(2)[rule_format]) using as(1) | 
| 208 | apply auto | |
| 209 | done | |
| 53248 | 210 | show "card s = Suc 0" | 
| 211 | unfolding * using card_insert by auto | |
| 49374 | 212 | 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 | 213 | |
| 53688 | 214 | 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))" | 
| 49374 | 215 | proof | 
| 216 | assume "card s = 2" | |
| 53688 | 217 |   then obtain x y where s: "s = {x, y}" "x \<noteq> y"
 | 
| 218 | unfolding numeral_2_eq_2 | |
| 53186 | 219 | apply - | 
| 220 | apply (erule exE conjE | drule card_eq_SucD)+ | |
| 221 | apply auto | |
| 222 | done | |
| 223 | show "\<exists>x\<in>s. \<exists>y\<in>s. x \<noteq> y \<and> (\<forall>z\<in>s. z = x \<or> z = y)" | |
| 224 | using s by auto | |
| 49555 | 225 | next | 
| 49374 | 226 | assume "\<exists>x\<in>s. \<exists>y\<in>s. x \<noteq> y \<and> (\<forall>z\<in>s. z = x \<or> z = y)" | 
| 53688 | 227 | then obtain x y where "x \<in> s" "y \<in> s" "x \<noteq> y" "\<forall>z\<in>s. z = x \<or> z = y" | 
| 53186 | 228 | by auto | 
| 53688 | 229 |   then have "s = {x, y}"
 | 
| 230 | by auto | |
| 231 | with `x \<noteq> y` show "card s = 2" | |
| 232 | by auto | |
| 49555 | 233 | qed | 
| 234 | ||
| 235 | lemma image_lemma_0: | |
| 236 |   assumes "card {a\<in>s. f ` (s - {a}) = t - {b}} = n"
 | |
| 237 |   shows "card {s'. \<exists>a\<in>s. (s' = s - {a}) \<and> (f ` s' = t - {b})} = n"
 | |
| 238 | proof - | |
| 53185 | 239 |   have *: "{s'. \<exists>a\<in>s. (s' = s - {a}) \<and> (f ` s' = t - {b})} =
 | 
| 240 |     (\<lambda>a. s - {a}) ` {a\<in>s. f ` (s - {a}) = t - {b}}"
 | |
| 49555 | 241 | by auto | 
| 53185 | 242 | show ?thesis | 
| 243 | unfolding * | |
| 244 | unfolding assms[symmetric] | |
| 245 | apply (rule card_image) | |
| 246 | unfolding inj_on_def | |
| 247 | apply (rule, rule, rule) | |
| 248 | unfolding mem_Collect_eq | |
| 249 | apply auto | |
| 250 | done | |
| 49555 | 251 | qed | 
| 252 | ||
| 253 | lemma image_lemma_1: | |
| 53688 | 254 | assumes "finite s" | 
| 255 | and "finite t" | |
| 256 | and "card s = card t" | |
| 257 | and "f ` s = t" | |
| 258 | and "b \<in> t" | |
| 49555 | 259 |   shows "card {s'. \<exists>a\<in>s. s' = s - {a} \<and>  f ` s' = t - {b}} = 1"
 | 
| 260 | proof - | |
| 53688 | 261 | obtain a where a: "b = f a" "a \<in> s" | 
| 262 | using assms(4-5) by auto | |
| 53185 | 263 | have inj: "inj_on f s" | 
| 264 | apply (rule eq_card_imp_inj_on) | |
| 53688 | 265 | using assms(1-4) | 
| 266 | apply auto | |
| 53185 | 267 | done | 
| 268 |   have *: "{a \<in> s. f ` (s - {a}) = t - {b}} = {a}"
 | |
| 269 | apply (rule set_eqI) | |
| 270 | unfolding singleton_iff | |
| 49555 | 271 | apply (rule, rule inj[unfolded inj_on_def, rule_format]) | 
| 53688 | 272 | unfolding a | 
| 273 | using a(2) and assms and inj[unfolded inj_on_def] | |
| 53252 | 274 | apply auto | 
| 49555 | 275 | done | 
| 53185 | 276 | show ?thesis | 
| 277 | apply (rule image_lemma_0) | |
| 278 | unfolding * | |
| 279 | apply auto | |
| 280 | done | |
| 49374 | 281 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 282 | |
| 49555 | 283 | lemma image_lemma_2: | 
| 284 | assumes "finite s" "finite t" "card s = card t" "f ` s \<subseteq> t" "f ` s \<noteq> t" "b \<in> t" | |
| 53688 | 285 |   shows "card {s'. \<exists>a\<in>s. (s' = s - {a}) \<and> f ` s' = t - {b}} = 0 \<or>
 | 
| 286 |          card {s'. \<exists>a\<in>s. (s' = s - {a}) \<and> f ` s' = t - {b}} = 2"
 | |
| 49555 | 287 | proof (cases "{a\<in>s. f ` (s - {a}) = t - {b}} = {}")
 | 
| 288 | case True | |
| 289 | then show ?thesis | |
| 290 | apply - | |
| 53185 | 291 | apply (rule disjI1, rule image_lemma_0) | 
| 53252 | 292 | using assms(1) | 
| 293 | apply auto | |
| 53185 | 294 | done | 
| 49555 | 295 | next | 
| 296 |   let ?M = "{a\<in>s. f ` (s - {a}) = t - {b}}"
 | |
| 297 | case False | |
| 53688 | 298 | then obtain a where "a \<in> ?M" | 
| 299 | by auto | |
| 300 |   then have a: "a \<in> s" "f ` (s - {a}) = t - {b}"
 | |
| 301 | by auto | |
| 302 |   have "f a \<in> t - {b}"
 | |
| 303 | using a and assms by auto | |
| 53186 | 304 |   then have "\<exists>c \<in> s - {a}. f a = f c"
 | 
| 53688 | 305 | unfolding image_iff[symmetric] and a | 
| 306 | by auto | |
| 307 | then obtain c where c: "c \<in> s" "a \<noteq> c" "f a = f c" | |
| 308 | by auto | |
| 49555 | 309 |   then have *: "f ` (s - {c}) = f ` (s - {a})"
 | 
| 310 | apply - | |
| 53688 | 311 | apply (rule set_eqI) | 
| 312 | apply rule | |
| 49555 | 313 | proof - | 
| 314 | fix x | |
| 315 |     assume "x \<in> f ` (s - {a})"
 | |
| 53688 | 316 |     then obtain y where y: "f y = x" "y \<in> s - {a}"
 | 
| 317 | by auto | |
| 49555 | 318 |     then show "x \<in> f ` (s - {c})"
 | 
| 319 | unfolding image_iff | |
| 320 | apply (rule_tac x = "if y = c then a else y" in bexI) | |
| 53688 | 321 | using c a | 
| 322 | apply auto | |
| 323 | done | |
| 49555 | 324 | qed auto | 
| 53688 | 325 | have "c \<in> ?M" | 
| 53186 | 326 | unfolding mem_Collect_eq and * | 
| 53688 | 327 | using a and c(1) | 
| 328 | by auto | |
| 49555 | 329 | show ?thesis | 
| 53688 | 330 | apply (rule disjI2) | 
| 331 | apply (rule image_lemma_0) | |
| 332 | unfolding card_2_exists | |
| 333 | apply (rule bexI[OF _ `a\<in>?M`]) | |
| 334 | apply (rule bexI[OF _ `c\<in>?M`]) | |
| 335 | apply rule | |
| 336 | apply (rule `a \<noteq> c`) | |
| 337 | apply rule | |
| 338 | apply (unfold mem_Collect_eq) | |
| 339 | apply (erule conjE) | |
| 340 | proof - | |
| 49555 | 341 | fix z | 
| 342 |     assume as: "z \<in> s" "f ` (s - {z}) = t - {b}"
 | |
| 343 |     have inj: "inj_on f (s - {z})"
 | |
| 344 | apply (rule eq_card_imp_inj_on) | |
| 53688 | 345 | unfolding as | 
| 346 | using as(1) and assms | |
| 53186 | 347 | apply auto | 
| 49555 | 348 | done | 
| 349 | show "z = a \<or> z = c" | |
| 350 | proof (rule ccontr) | |
| 351 | assume "\<not> ?thesis" | |
| 352 | then show False | |
| 353 | using inj[unfolded inj_on_def, THEN bspec[where x=a], THEN bspec[where x=c]] | |
| 53688 | 354 | using `a \<in> s` `c \<in> s` `f a = f c` `a \<noteq> c` | 
| 53186 | 355 | apply auto | 
| 49555 | 356 | done | 
| 357 | qed | |
| 358 | qed | |
| 359 | qed | |
| 360 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 361 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 362 | subsection {* Combine this with the basic counting lemma. *}
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 363 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 364 | lemma kuhn_complete_lemma: | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 365 | assumes "finite simplices" | 
| 53688 | 366 |     and "\<forall>f s. face f s \<longleftrightarrow> (\<exists>a\<in>s. f = s - {a})"
 | 
| 367 | and "\<forall>s\<in>simplices. card s = n + 2" | |
| 368 |     and "\<forall>s\<in>simplices. (rl ` s) \<subseteq> {0..n+1}"
 | |
| 369 |     and "\<forall>f\<in>{f. \<exists>s\<in>simplices. face f s}. bnd f  \<longrightarrow> card {s\<in>simplices. face f s} = 1"
 | |
| 370 |     and "\<forall>f\<in>{f. \<exists>s\<in>simplices. face f s}. \<not> bnd f \<longrightarrow> card {s\<in>simplices. face f s} = 2"
 | |
| 371 |     and "odd (card {f\<in>{f. \<exists>s\<in>simplices. face f s}. rl ` f = {0..n} \<and> bnd f})"
 | |
| 53185 | 372 |   shows "odd (card {s\<in>simplices. (rl ` s = {0..n+1})})"
 | 
| 49555 | 373 | apply (rule kuhn_counting_lemma) | 
| 374 | defer | |
| 375 | apply (rule assms)+ | |
| 376 | prefer 3 | |
| 377 | apply (rule assms) | |
| 53688 | 378 | apply (rule_tac[1-2] ballI impI)+ | 
| 379 | proof - | |
| 49555 | 380 |   have *: "{f. \<exists>s\<in>simplices. \<exists>a\<in>s. f = s - {a}} = (\<Union>s\<in>simplices. {f. \<exists>a\<in>s. (f = s - {a})})"
 | 
| 381 | by auto | |
| 382 | have **: "\<forall>s\<in>simplices. card s = n + 2 \<and> finite s" | |
| 383 | using assms(3) by (auto intro: card_ge_0_finite) | |
| 384 |   show "finite {f. \<exists>s\<in>simplices. face f s}"
 | |
| 385 | unfolding assms(2)[rule_format] and * | |
| 53252 | 386 | apply (rule finite_UN_I[OF assms(1)]) | 
| 387 | using ** | |
| 388 | apply auto | |
| 49555 | 389 | done | 
| 390 |   have *: "\<And>P f s. s\<in>simplices \<Longrightarrow> (f \<in> {f. \<exists>s\<in>simplices. \<exists>a\<in>s. f = s - {a}}) \<and>
 | |
| 53688 | 391 |     (\<exists>a\<in>s. (f = s - {a})) \<and> P f \<longleftrightarrow> (\<exists>a\<in>s. (f = s - {a}) \<and> P f)"
 | 
| 392 | by auto | |
| 393 | fix s | |
| 394 | assume s: "s \<in> simplices" | |
| 49555 | 395 |   let ?S = "{f \<in> {f. \<exists>s\<in>simplices. face f s}. face f s \<and> rl ` f = {0..n}}"
 | 
| 53688 | 396 |   have "{0..n + 1} - {n + 1} = {0..n}"
 | 
| 397 | by auto | |
| 49555 | 398 |   then have S: "?S = {s'. \<exists>a\<in>s. s' = s - {a} \<and> rl ` s' = {0..n + 1} - {n + 1}}"
 | 
| 399 | apply - | |
| 400 | apply (rule set_eqI) | |
| 53185 | 401 | unfolding assms(2)[rule_format] mem_Collect_eq | 
| 402 |     unfolding *[OF s, unfolded mem_Collect_eq, where P="\<lambda>x. rl ` x = {0..n}"]
 | |
| 49555 | 403 | apply auto | 
| 404 | done | |
| 53688 | 405 |   show "rl ` s = {0..n+1} \<Longrightarrow> card ?S = 1" and "rl ` s \<noteq> {0..n+1} \<Longrightarrow> card ?S = 0 \<or> card ?S = 2"
 | 
| 49555 | 406 | unfolding S | 
| 53688 | 407 | apply (rule_tac[!] image_lemma_1 image_lemma_2) | 
| 53252 | 408 | using ** assms(4) and s | 
| 409 | apply auto | |
| 49555 | 410 | done | 
| 411 | qed | |
| 412 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 413 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 414 | subsection {*We use the following notion of ordering rather than pointwise indexing. *}
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 415 | |
| 53688 | 416 | definition "kle n x y \<longleftrightarrow> (\<exists>k\<subseteq>{1..n::nat}. \<forall>j. y j = x j + (if j \<in> k then (1::nat) else 0))"
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 417 | |
| 53185 | 418 | lemma kle_refl [intro]: "kle n x x" | 
| 419 | unfolding kle_def 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 | 420 | |
| 53688 | 421 | lemma kle_antisym: "kle n x y \<and> kle n y x \<longleftrightarrow> x = y" | 
| 53185 | 422 | unfolding kle_def | 
| 423 | apply rule | |
| 424 | apply auto | |
| 425 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 426 | |
| 53185 | 427 | lemma pointwise_minimal_pointwise_maximal: | 
| 428 | fixes s :: "(nat \<Rightarrow> nat) set" | |
| 53688 | 429 | assumes "finite s" | 
| 430 |     and "s \<noteq> {}"
 | |
| 431 | and "\<forall>x\<in>s. \<forall>y\<in>s. (\<forall>j. x j \<le> y j) \<or> (\<forall>j. y j \<le> x j)" | |
| 432 | shows "\<exists>a\<in>s. \<forall>x\<in>s. \<forall>j. a j \<le> x j" | |
| 433 | and "\<exists>a\<in>s. \<forall>x\<in>s. \<forall>j. x j \<le> a j" | |
| 434 | using assms | |
| 435 | unfolding atomize_conj | |
| 53185 | 436 | proof (induct s rule: finite_induct) | 
| 53688 | 437 | fix x | 
| 438 | fix F :: "(nat \<Rightarrow> nat) set" | |
| 439 | assume as: "finite F" "x \<notin> F" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 440 |     "\<lbrakk>F \<noteq> {}; \<forall>x\<in>F. \<forall>y\<in>F. (\<forall>j. x j \<le> y j) \<or> (\<forall>j. y j \<le> x j)\<rbrakk>
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 441 |         \<Longrightarrow> (\<exists>a\<in>F. \<forall>x\<in>F. \<forall>j. a j \<le> x j) \<and> (\<exists>a\<in>F. \<forall>x\<in>F. \<forall>j. x j \<le> a j)" "insert x F \<noteq> {}"
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 442 | "\<forall>xa\<in>insert x F. \<forall>y\<in>insert x F. (\<forall>j. xa j \<le> y j) \<or> (\<forall>j. y j \<le> xa j)" | 
| 49555 | 443 | show "(\<exists>a\<in>insert x F. \<forall>x\<in>insert x F. \<forall>j. a j \<le> x j) \<and> (\<exists>a\<in>insert x F. \<forall>x\<in>insert x F. \<forall>j. x j \<le> a j)" | 
| 444 |   proof (cases "F = {}")
 | |
| 445 | case True | |
| 446 | then show ?thesis | |
| 447 | apply - | |
| 448 | apply (rule, rule_tac[!] x=x in bexI) | |
| 449 | apply auto | |
| 450 | done | |
| 451 | next | |
| 452 | case False | |
| 53186 | 453 | obtain a b where a: "a\<in>insert x F" "\<forall>x\<in>F. \<forall>j. a j \<le> x j" | 
| 454 | and b: "b \<in> insert x F" "\<forall>x\<in>F. \<forall>j. x j \<le> b j" | |
| 455 | using as(3)[OF False] using as(5) by auto | |
| 456 | have "\<exists>a \<in> insert x F. \<forall>x \<in> insert x F. \<forall>j. a j \<le> x j" | |
| 49555 | 457 | using as(5)[rule_format,OF a(1) insertI1] | 
| 458 | apply - | |
| 459 | proof (erule disjE) | |
| 460 | assume "\<forall>j. a j \<le> x j" | |
| 461 | then show ?thesis | |
| 53185 | 462 | apply (rule_tac x=a in bexI) | 
| 53688 | 463 | using a | 
| 464 | apply auto | |
| 53185 | 465 | done | 
| 49555 | 466 | next | 
| 467 | assume "\<forall>j. x j \<le> a j" | |
| 468 | then show ?thesis | |
| 469 | apply (rule_tac x=x in bexI) | |
| 53185 | 470 | apply (rule, rule) | 
| 471 | apply (insert a) | |
| 49555 | 472 | apply (erule_tac x=xa in ballE) | 
| 473 | apply (erule_tac x=j in allE)+ | |
| 474 | apply auto | |
| 475 | done | |
| 476 | qed | |
| 477 | moreover | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 478 | have "\<exists>b\<in>insert x F. \<forall>x\<in>insert x F. \<forall>j. x j \<le> b j" | 
| 53186 | 479 | using as(5)[rule_format,OF b(1) insertI1] | 
| 480 | apply - | |
| 49555 | 481 | proof (erule disjE) | 
| 482 | assume "\<forall>j. x j \<le> b j" | |
| 483 | then show ?thesis | |
| 484 | apply(rule_tac x=b in bexI) using b | |
| 485 | apply auto | |
| 486 | done | |
| 487 | next | |
| 488 | assume "\<forall>j. b j \<le> x j" | |
| 489 | then show ?thesis | |
| 490 | apply (rule_tac x=x in bexI) | |
| 53185 | 491 | apply (rule, rule) | 
| 492 | apply (insert b) | |
| 49555 | 493 | apply (erule_tac x=xa in ballE) | 
| 494 | apply (erule_tac x=j in allE)+ | |
| 495 | apply auto | |
| 496 | done | |
| 497 | qed | |
| 498 | ultimately show ?thesis by auto | |
| 499 | qed | |
| 500 | qed auto | |
| 501 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 502 | |
| 53688 | 503 | lemma kle_imp_pointwise: "kle n x y \<Longrightarrow> \<forall>j. x j \<le> y j" | 
| 53186 | 504 | unfolding kle_def 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 | 505 | |
| 49555 | 506 | lemma pointwise_antisym: | 
| 507 | fixes x :: "nat \<Rightarrow> nat" | |
| 53252 | 508 | shows "(\<forall>j. x j \<le> y j) \<and> (\<forall>j. y j \<le> x j) \<longleftrightarrow> x = y" | 
| 53688 | 509 | apply rule | 
| 510 | apply (rule ext) | |
| 511 | apply (erule conjE) | |
| 53252 | 512 | apply (erule_tac x = xa in allE)+ | 
| 49555 | 513 | apply auto | 
| 514 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 515 | |
| 49555 | 516 | lemma kle_trans: | 
| 53688 | 517 | assumes "kle n x y" | 
| 518 | and "kle n y z" | |
| 519 | and "kle n x z \<or> kle n z x" | |
| 49555 | 520 | shows "kle n x z" | 
| 521 | using assms | |
| 53688 | 522 | apply - | 
| 523 | apply (erule disjE) | |
| 524 | apply assumption | |
| 49555 | 525 | proof - | 
| 526 | case goal1 | |
| 527 | then have "x = z" | |
| 528 | apply - | |
| 529 | apply (rule ext) | |
| 530 | apply (drule kle_imp_pointwise)+ | |
| 531 | apply (erule_tac x=xa in allE)+ | |
| 532 | apply auto | |
| 533 | done | |
| 534 | then show ?case by auto | |
| 535 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 536 | |
| 49555 | 537 | lemma kle_strict: | 
| 53688 | 538 | assumes "kle n x y" | 
| 539 | and "x \<noteq> y" | |
| 540 | shows "\<forall>j. x j \<le> y j" | |
| 541 | and "\<exists>k. 1 \<le> k \<and> k \<le> n \<and> x k < y k" | |
| 49555 | 542 | apply (rule kle_imp_pointwise[OF assms(1)]) | 
| 543 | proof - | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 544 | guess k using assms(1)[unfolded kle_def] .. note k = this | 
| 53688 | 545 | show "\<exists>k. 1 \<le> k \<and> k \<le> n \<and> x k < y k" | 
| 53186 | 546 | proof (cases "k = {}")
 | 
| 49555 | 547 | case True | 
| 548 | then have "x = y" | |
| 549 | apply - | |
| 550 | apply (rule ext) | |
| 53688 | 551 | using k | 
| 552 | apply auto | |
| 49555 | 553 | done | 
| 53688 | 554 | then show ?thesis | 
| 555 | using assms(2) by auto | |
| 49555 | 556 | next | 
| 557 | case False | |
| 558 | then have "(SOME k'. k' \<in> k) \<in> k" | |
| 559 | apply - | |
| 560 | apply (rule someI_ex) | |
| 561 | apply auto | |
| 562 | done | |
| 563 | then show ?thesis | |
| 564 | apply (rule_tac x = "SOME k'. k' \<in> k" in exI) | |
| 53688 | 565 | using k | 
| 566 | apply auto | |
| 49555 | 567 | done | 
| 568 | qed | |
| 569 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 570 | |
| 53185 | 571 | lemma kle_minimal: | 
| 53688 | 572 | assumes "finite s" | 
| 573 |     and "s \<noteq> {}"
 | |
| 574 | and "\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x" | |
| 53185 | 575 | shows "\<exists>a\<in>s. \<forall>x\<in>s. kle n a x" | 
| 576 | proof - | |
| 577 | have "\<exists>a\<in>s. \<forall>x\<in>s. \<forall>j. a j \<le> x j" | |
| 578 | apply (rule pointwise_minimal_pointwise_maximal(1)[OF assms(1-2)]) | |
| 53688 | 579 | apply rule | 
| 580 | apply rule | |
| 581 | apply (drule_tac assms(3)[rule_format]) | |
| 582 | apply assumption | |
| 53252 | 583 | using kle_imp_pointwise | 
| 584 | apply auto | |
| 53185 | 585 | done | 
| 586 | then guess a .. note a = this | |
| 587 | show ?thesis | |
| 53252 | 588 | apply (rule_tac x = a in bexI) | 
| 53185 | 589 | proof | 
| 590 | fix x | |
| 591 | assume "x \<in> s" | |
| 592 | show "kle n a x" | |
| 593 | using assms(3)[rule_format,OF a(1) `x\<in>s`] | |
| 594 | apply - | |
| 595 | proof (erule disjE) | |
| 596 | assume "kle n x a" | |
| 53252 | 597 | then have "x = a" | 
| 53185 | 598 | apply - | 
| 599 | unfolding pointwise_antisym[symmetric] | |
| 600 | apply (drule kle_imp_pointwise) | |
| 53252 | 601 | using a(2)[rule_format,OF `x\<in>s`] | 
| 602 | apply auto | |
| 53185 | 603 | done | 
| 53252 | 604 | then show ?thesis using kle_refl by auto | 
| 53185 | 605 | qed | 
| 606 | qed (insert a, auto) | |
| 607 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 608 | |
| 53185 | 609 | lemma kle_maximal: | 
| 53688 | 610 | assumes "finite s" | 
| 611 |     and "s \<noteq> {}"
 | |
| 612 | and "\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x" | |
| 53185 | 613 | shows "\<exists>a\<in>s. \<forall>x\<in>s. kle n x a" | 
| 614 | proof - | |
| 615 | have "\<exists>a\<in>s. \<forall>x\<in>s. \<forall>j. a j \<ge> x j" | |
| 616 | apply (rule pointwise_minimal_pointwise_maximal(2)[OF assms(1-2)]) | |
| 53688 | 617 | apply rule | 
| 618 | apply rule | |
| 53185 | 619 | apply (drule_tac assms(3)[rule_format],assumption) | 
| 53688 | 620 | using kle_imp_pointwise | 
| 621 | apply auto | |
| 53185 | 622 | done | 
| 623 | then guess a .. note a = this | |
| 624 | show ?thesis | |
| 53186 | 625 | apply (rule_tac x = a in bexI) | 
| 53185 | 626 | proof | 
| 627 | fix x | |
| 628 | assume "x \<in> s" | |
| 629 | show "kle n x a" | |
| 630 | using assms(3)[rule_format,OF a(1) `x\<in>s`] | |
| 631 | apply - | |
| 632 | proof (erule disjE) | |
| 633 | assume "kle n a x" | |
| 53674 | 634 | then have "x = a" | 
| 53185 | 635 | apply - | 
| 636 | unfolding pointwise_antisym[symmetric] | |
| 637 | apply (drule kle_imp_pointwise) | |
| 53688 | 638 | using a(2)[rule_format,OF `x\<in>s`] | 
| 639 | apply auto | |
| 53185 | 640 | done | 
| 53688 | 641 | then show ?thesis | 
| 642 | using kle_refl by auto | |
| 53185 | 643 | qed | 
| 644 | qed (insert a, auto) | |
| 645 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 646 | |
| 53185 | 647 | lemma kle_strict_set: | 
| 53688 | 648 | assumes "kle n x y" | 
| 649 | and "x \<noteq> y" | |
| 53185 | 650 |   shows "1 \<le> card {k\<in>{1..n}. x k < y k}"
 | 
| 651 | proof - | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 652 | guess i using kle_strict(2)[OF assms] .. | 
| 53674 | 653 |   then have "card {i} \<le> card {k\<in>{1..n}. x k < y k}"
 | 
| 53185 | 654 | apply - | 
| 655 | apply (rule card_mono) | |
| 656 | apply auto | |
| 657 | done | |
| 53688 | 658 | then show ?thesis | 
| 659 | by auto | |
| 53185 | 660 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 661 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 662 | lemma kle_range_combine: | 
| 53688 | 663 | assumes "kle n x y" | 
| 664 | and "kle n y z" | |
| 665 | and "kle n x z \<or> kle n z x" | |
| 666 |     and "m1 \<le> card {k\<in>{1..n}. x k < y k}"
 | |
| 667 |     and "m2 \<le> card {k\<in>{1..n}. y k < z k}"
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 668 |   shows "kle n x z \<and> m1 + m2 \<le> card {k\<in>{1..n}. x k < z k}"
 | 
| 53688 | 669 | apply rule | 
| 670 | apply (rule kle_trans[OF assms(1-3)]) | |
| 53185 | 671 | proof - | 
| 672 | have "\<And>j. x j < y j \<Longrightarrow> x j < z j" | |
| 673 | apply (rule less_le_trans) | |
| 53252 | 674 | using kle_imp_pointwise[OF assms(2)] | 
| 675 | apply auto | |
| 53185 | 676 | done | 
| 677 | moreover | |
| 678 | have "\<And>j. y j < z j \<Longrightarrow> x j < z j" | |
| 679 | apply (rule le_less_trans) | |
| 53252 | 680 | using kle_imp_pointwise[OF assms(1)] | 
| 681 | apply auto | |
| 53185 | 682 | done | 
| 683 | ultimately | |
| 684 |   have *: "{k\<in>{1..n}. x k < y k} \<union> {k\<in>{1..n}. y k < z k} = {k\<in>{1..n}. x k < z k}"
 | |
| 685 | by auto | |
| 686 |   have **: "{k \<in> {1..n}. x k < y k} \<inter> {k \<in> {1..n}. y k < z k} = {}"
 | |
| 687 | unfolding disjoint_iff_not_equal | |
| 53688 | 688 | apply rule | 
| 689 | apply rule | |
| 690 | apply (unfold mem_Collect_eq) | |
| 691 | apply (rule notI) | |
| 53185 | 692 | apply (erule conjE)+ | 
| 693 | proof - | |
| 694 | fix i j | |
| 53688 | 695 |     assume as: "i \<in> {1..n}" "x i < y i" "j \<in> {1..n}" "y j < z j" "i = j"
 | 
| 53185 | 696 | guess kx using assms(1)[unfolded kle_def] .. note kx = this | 
| 53688 | 697 | have "x i < y i" | 
| 698 | using as by auto | |
| 699 | then have "i \<in> kx" | |
| 700 | using as(1) kx | |
| 701 | apply - | |
| 702 | apply (rule ccontr) | |
| 53185 | 703 | apply auto | 
| 704 | done | |
| 53688 | 705 | then have "x i + 1 = y i" | 
| 706 | using kx by auto | |
| 53185 | 707 | moreover | 
| 708 | guess ky using assms(2)[unfolded kle_def] .. note ky = this | |
| 53688 | 709 | have "y i < z i" | 
| 710 | using as by auto | |
| 711 | then have "i \<in> ky" | |
| 712 | using as(1) ky | |
| 713 | apply - | |
| 714 | apply (rule ccontr) | |
| 53185 | 715 | apply auto | 
| 716 | done | |
| 53688 | 717 | then have "y i + 1 = z i" | 
| 718 | using ky by auto | |
| 53185 | 719 | ultimately | 
| 53688 | 720 | have "z i = x i + 2" | 
| 721 | by auto | |
| 722 | then show False | |
| 723 | using assms(3) | |
| 724 | unfolding kle_def | |
| 53185 | 725 | by (auto simp add: split_if_eq1) | 
| 726 | qed | |
| 53688 | 727 |   have fin: "\<And>P. finite {x\<in>{1..n::nat}. P x}"
 | 
| 728 | by auto | |
| 53185 | 729 |   have "m1 + m2 \<le> card {k\<in>{1..n}. x k < y k} + card {k\<in>{1..n}. y k < z k}"
 | 
| 730 | using assms(4-5) by auto | |
| 731 |   also have "\<dots> \<le> card {k\<in>{1..n}. x k < z k}"
 | |
| 53688 | 732 | unfolding card_Un_Int[OF fin fin] | 
| 733 | unfolding * ** | |
| 734 | by auto | |
| 735 |   finally show "m1 + m2 \<le> card {k \<in> {1..n}. x k < z k}"
 | |
| 736 | by auto | |
| 53185 | 737 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 738 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 739 | lemma kle_range_combine_l: | 
| 53185 | 740 | assumes "kle n x y" | 
| 741 | and "kle n y z" | |
| 742 | and "kle n x z \<or> kle n z x" | |
| 743 |     and "m \<le> card {k\<in>{1..n}. y(k) < z(k)}"
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 744 |   shows "kle n x z \<and> m \<le> card {k\<in>{1..n}. x(k) < z(k)}"
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 745 | using kle_range_combine[OF assms(1-3) _ assms(4), of 0] by auto | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 746 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 747 | lemma kle_range_combine_r: | 
| 53185 | 748 | assumes "kle n x y" | 
| 749 | and "kle n y z" | |
| 750 |     and "kle n x z \<or> kle n z x" "m \<le> card {k\<in>{1..n}. x k < y k}"
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 751 |   shows "kle n x z \<and> m \<le> card {k\<in>{1..n}. x(k) < z(k)}"
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 752 | using kle_range_combine[OF assms(1-3) assms(4), of 0] by auto | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 753 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 754 | lemma kle_range_induct: | 
| 53185 | 755 | assumes "card s = Suc m" | 
| 756 | and "\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x" | |
| 757 |   shows "\<exists>x\<in>s. \<exists>y\<in>s. kle n x y \<and> m \<le> card {k\<in>{1..n}. x k < y k}"
 | |
| 758 | proof - | |
| 53688 | 759 |   have "finite s" and "s \<noteq> {}"
 | 
| 760 | using assms(1) | |
| 53185 | 761 | by (auto intro: card_ge_0_finite) | 
| 53688 | 762 | then show ?thesis | 
| 763 | using assms | |
| 53185 | 764 | proof (induct m arbitrary: s) | 
| 765 | case 0 | |
| 53688 | 766 | then show ?case | 
| 767 | using kle_refl by auto | |
| 53185 | 768 | next | 
| 769 | case (Suc m) | |
| 770 | then obtain a where a: "a \<in> s" "\<forall>x\<in>s. kle n a x" | |
| 771 | using kle_minimal[of s n] by auto | |
| 772 | show ?case | |
| 773 |     proof (cases "s \<subseteq> {a}")
 | |
| 774 | case False | |
| 53688 | 775 |       then have "card (s - {a}) = Suc m" and "s - {a} \<noteq> {}"
 | 
| 776 | using card_Diff_singleton[OF _ a(1)] Suc(4) `finite s` | |
| 777 | by auto | |
| 778 | then obtain x b where xb: | |
| 779 |         "x \<in> s - {a}"
 | |
| 780 |         "b \<in> s - {a}"
 | |
| 781 | "kle n x b" | |
| 782 |         "m \<le> card {k \<in> {1..n}. x k < b k}"
 | |
| 783 |         using Suc(1)[of "s - {a}"]
 | |
| 784 | using Suc(5) `finite s` | |
| 785 | by auto | |
| 786 |       have "1 \<le> card {k \<in> {1..n}. a k < x k}" and "m \<le> card {k \<in> {1..n}. x k < b k}"
 | |
| 53185 | 787 | apply (rule kle_strict_set) | 
| 788 | apply (rule a(2)[rule_format]) | |
| 53248 | 789 | using a and xb | 
| 790 | apply auto | |
| 53185 | 791 | done | 
| 53674 | 792 | then show ?thesis | 
| 53688 | 793 | apply (rule_tac x=a in bexI) | 
| 794 | apply (rule_tac x=b in bexI) | |
| 53185 | 795 | using kle_range_combine[OF a(2)[rule_format] xb(3) Suc(5)[rule_format], of 1 "m"] | 
| 53248 | 796 | using a(1) xb(1-2) | 
| 797 | apply auto | |
| 53185 | 798 | done | 
| 799 | next | |
| 800 | case True | |
| 53688 | 801 |       then have "s = {a}"
 | 
| 802 | using Suc(3) by auto | |
| 803 | then have "card s = 1" | |
| 804 | by auto | |
| 805 | then have False | |
| 806 | using Suc(4) `finite s` by auto | |
| 807 | then show ?thesis | |
| 808 | by auto | |
| 53185 | 809 | qed | 
| 810 | qed | |
| 811 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 812 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 813 | lemma kle_Suc: "kle n x y \<Longrightarrow> kle (n + 1) x y" | 
| 53185 | 814 | unfolding kle_def | 
| 815 | apply (erule exE) | |
| 816 | apply (rule_tac x=k in exI) | |
| 817 | apply auto | |
| 818 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 819 | |
| 53185 | 820 | lemma kle_trans_1: | 
| 821 | assumes "kle n x y" | |
| 53688 | 822 | shows "x j \<le> y j" | 
| 823 | and "y j \<le> x j + 1" | |
| 53185 | 824 | using assms[unfolded kle_def] 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 | 825 | |
| 53185 | 826 | lemma kle_trans_2: | 
| 53688 | 827 | assumes "kle n a b" | 
| 828 | and "kle n b c" | |
| 829 | and "\<forall>j. c j \<le> a j + 1" | |
| 53185 | 830 | shows "kle n a c" | 
| 831 | proof - | |
| 832 | guess kk1 using assms(1)[unfolded kle_def] .. note kk1 = this | |
| 833 | guess kk2 using assms(2)[unfolded kle_def] .. note kk2 = this | |
| 834 | show ?thesis | |
| 835 | unfolding kle_def | |
| 836 | apply (rule_tac x="kk1 \<union> kk2" in exI) | |
| 53252 | 837 | apply rule | 
| 838 | defer | |
| 53185 | 839 | proof | 
| 840 | fix i | |
| 841 | show "c i = a i + (if i \<in> kk1 \<union> kk2 then 1 else 0)" | |
| 842 | proof (cases "i \<in> kk1 \<union> kk2") | |
| 843 | case True | |
| 53674 | 844 | then have "c i \<ge> a i + (if i \<in> kk1 \<union> kk2 then 1 else 0)" | 
| 53185 | 845 | unfolding kk1[THEN conjunct2,rule_format,of i] kk2[THEN conjunct2,rule_format,of i] | 
| 846 | by auto | |
| 847 | moreover | |
| 848 | have "c i \<le> a i + (if i \<in> kk1 \<union> kk2 then 1 else 0)" | |
| 849 | using True assms(3) by auto | |
| 850 | ultimately show ?thesis by auto | |
| 851 | next | |
| 852 | case False | |
| 53688 | 853 | then show ?thesis | 
| 854 | using kk1 kk2 by auto | |
| 53185 | 855 | qed | 
| 856 | qed (insert kk1 kk2, auto) | |
| 857 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 858 | |
| 53185 | 859 | lemma kle_between_r: | 
| 53688 | 860 | assumes "kle n a b" | 
| 861 | and "kle n b c" | |
| 862 | and "kle n a x" | |
| 863 | and "kle n c x" | |
| 53185 | 864 | shows "kle n b x" | 
| 865 | apply (rule kle_trans_2[OF assms(2,4)]) | |
| 866 | proof | |
| 53688 | 867 | have *: "\<And>c b x::nat. x \<le> c + 1 \<Longrightarrow> c \<le> b \<Longrightarrow> x \<le> b + 1" | 
| 868 | by auto | |
| 53185 | 869 | fix j | 
| 870 | show "x j \<le> b j + 1" | |
| 871 | apply (rule *) | |
| 53186 | 872 | using kle_trans_1[OF assms(1),of j] kle_trans_1[OF assms(3), of j] | 
| 873 | apply auto | |
| 53185 | 874 | done | 
| 875 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 876 | |
| 53185 | 877 | lemma kle_between_l: | 
| 53688 | 878 | assumes "kle n a b" | 
| 879 | and "kle n b c" | |
| 880 | and "kle n x a" | |
| 881 | and "kle n x c" | |
| 53185 | 882 | shows "kle n x b" | 
| 883 | apply (rule kle_trans_2[OF assms(3,1)]) | |
| 884 | proof | |
| 885 | have *: "\<And>c b x::nat. c \<le> x + 1 \<Longrightarrow> b \<le> c \<Longrightarrow> b \<le> x + 1" | |
| 886 | by auto | |
| 887 | fix j | |
| 888 | show "b j \<le> x j + 1" | |
| 889 | apply (rule *) | |
| 53186 | 890 | using kle_trans_1[OF assms(2),of j] kle_trans_1[OF assms(4), of j] | 
| 891 | apply auto | |
| 53185 | 892 | done | 
| 893 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 894 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 895 | lemma kle_adjacent: | 
| 53688 | 896 | assumes "\<forall>j. b j = (if j = k then a(j) + 1 else a j)" | 
| 897 | and "kle n a x" | |
| 898 | and "kle n x b" | |
| 53252 | 899 | shows "x = a \<or> x = b" | 
| 53185 | 900 | proof (cases "x k = a k") | 
| 901 | case True | |
| 902 | show ?thesis | |
| 53186 | 903 | apply (rule disjI1) | 
| 904 | apply (rule ext) | |
| 53185 | 905 | proof - | 
| 906 | fix j | |
| 53688 | 907 | have "x j \<le> a j" | 
| 908 | using kle_imp_pointwise[OF assms(3),THEN spec[where x=j]] | |
| 53186 | 909 | unfolding assms(1)[rule_format] | 
| 910 | apply - | |
| 911 | apply(cases "j = k") | |
| 912 | using True | |
| 913 | apply auto | |
| 914 | done | |
| 53252 | 915 | then show "x j = a j" | 
| 53688 | 916 | using kle_imp_pointwise[OF assms(2),THEN spec[where x=j]] | 
| 917 | by auto | |
| 53185 | 918 | qed | 
| 919 | next | |
| 920 | case False | |
| 53688 | 921 | show ?thesis | 
| 922 | apply (rule disjI2) | |
| 923 | apply (rule ext) | |
| 53185 | 924 | proof - | 
| 925 | fix j | |
| 53248 | 926 | have "x j \<ge> b j" | 
| 927 | using kle_imp_pointwise[OF assms(2),THEN spec[where x=j]] | |
| 928 | unfolding assms(1)[rule_format] | |
| 929 | apply - | |
| 53688 | 930 | apply (cases "j = k") | 
| 931 | using False | |
| 932 | apply auto | |
| 933 | done | |
| 53252 | 934 | then show "x j = b j" | 
| 53248 | 935 | using kle_imp_pointwise[OF assms(3),THEN spec[where x=j]] | 
| 53185 | 936 | by auto | 
| 937 | qed | |
| 938 | qed | |
| 939 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 940 | |
| 53688 | 941 | subsection {* Kuhn's notion of a simplex (a reformulation to avoid so much indexing) *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 942 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 943 | definition "ksimplex p n (s::(nat \<Rightarrow> nat) set) \<longleftrightarrow> | 
| 53674 | 944 | card s = n + 1 \<and> | 
| 945 | (\<forall>x\<in>s. \<forall>j. x j \<le> p) \<and> | |
| 946 |   (\<forall>x\<in>s. \<forall>j. j \<notin> {1..n} \<longrightarrow> x j = p) \<and>
 | |
| 947 | (\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x)" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 948 | |
| 53185 | 949 | lemma ksimplexI: | 
| 950 | "card s = n + 1 \<Longrightarrow> | |
| 951 | \<forall>x\<in>s. \<forall>j. x j \<le> p \<Longrightarrow> | |
| 952 |   \<forall>x\<in>s. \<forall>j. j \<notin> {1..n} \<longrightarrow> x j = p \<Longrightarrow>
 | |
| 953 | \<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x \<Longrightarrow> | |
| 954 | ksimplex p n s" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 955 | unfolding ksimplex_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 | 956 | |
| 53185 | 957 | lemma ksimplex_eq: | 
| 958 | "ksimplex p n (s::(nat \<Rightarrow> nat) set) \<longleftrightarrow> | |
| 53688 | 959 | card s = n + 1 \<and> | 
| 960 | finite s \<and> | |
| 53185 | 961 | (\<forall>x\<in>s. \<forall>j. x(j) \<le> p) \<and> | 
| 53688 | 962 |     (\<forall>x\<in>s. \<forall>j. j\<notin>{1..n} \<longrightarrow> x j = p) \<and>
 | 
| 963 | (\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x)" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 964 | unfolding ksimplex_def by (auto intro: card_ge_0_finite) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 965 | |
| 53185 | 966 | lemma ksimplex_extrema: | 
| 967 | assumes "ksimplex p n s" | |
| 53688 | 968 | obtains a b where "a \<in> s" | 
| 969 | and "b \<in> s" | |
| 970 | and "\<forall>x\<in>s. kle n a x \<and> kle n x b" | |
| 971 |     and "\<forall>i. b i = (if i \<in> {1..n} then a i + 1 else a i)"
 | |
| 53185 | 972 | proof (cases "n = 0") | 
| 973 | case True | |
| 974 |   obtain x where *: "s = {x}"
 | |
| 975 | using assms[unfolded ksimplex_eq True,THEN conjunct1] | |
| 53688 | 976 | unfolding add_0_left card_1_exists | 
| 977 | by auto | |
| 53185 | 978 | show ?thesis | 
| 53688 | 979 | apply (rule that[of x x]) | 
| 980 | unfolding * True | |
| 53185 | 981 | apply auto | 
| 982 | done | |
| 983 | next | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 984 | note assm = assms[unfolded ksimplex_eq] | 
| 53185 | 985 | case False | 
| 53688 | 986 |   have "s \<noteq> {}"
 | 
| 987 | using assm by auto | |
| 53185 | 988 | obtain a where a: "a \<in> s" "\<forall>x\<in>s. kle n a x" | 
| 53688 | 989 |     using `s \<noteq> {}` assm
 | 
| 990 | using kle_minimal[of s n] | |
| 991 | by auto | |
| 53185 | 992 | obtain b where b: "b \<in> s" "\<forall>x\<in>s. kle n x b" | 
| 53688 | 993 |     using `s \<noteq> {}` assm
 | 
| 994 | using kle_maximal[of s n] | |
| 995 | by auto | |
| 53248 | 996 |   obtain c d where c_d: "c \<in> s" "d \<in> s" "kle n c d" "n \<le> card {k \<in> {1..n}. c k < d k}"
 | 
| 53688 | 997 | using kle_range_induct[of s n n] | 
| 998 | using assm | |
| 999 | by auto | |
| 53185 | 1000 |   have "kle n c b \<and> n \<le> card {k \<in> {1..n}. c k < b k}"
 | 
| 1001 | apply (rule kle_range_combine_r[where y=d]) | |
| 53252 | 1002 | using c_d a b | 
| 1003 | apply auto | |
| 53185 | 1004 | done | 
| 53674 | 1005 |   then have "kle n a b \<and> n \<le> card {k\<in>{1..n}. a(k) < b(k)}"
 | 
| 53185 | 1006 | apply - | 
| 1007 | apply (rule kle_range_combine_l[where y=c]) | |
| 53252 | 1008 | using a `c \<in> s` `b \<in> s` | 
| 1009 | apply auto | |
| 53185 | 1010 | done | 
| 1011 | moreover | |
| 1012 |   have "card {1..n} \<ge> card {k\<in>{1..n}. a(k) < b(k)}"
 | |
| 1013 | by (rule card_mono) auto | |
| 1014 | ultimately | |
| 1015 |   have *: "{k\<in>{1 .. n}. a k < b k} = {1..n}"
 | |
| 1016 | apply - | |
| 1017 | apply (rule card_subset_eq) | |
| 1018 | apply auto | |
| 1019 | done | |
| 1020 | show ?thesis | |
| 1021 | apply (rule that[OF a(1) b(1)]) | |
| 1022 | defer | |
| 53688 | 1023 | apply (subst *[symmetric]) | 
| 1024 | unfolding mem_Collect_eq | |
| 53185 | 1025 | proof | 
| 1026 | guess k using a(2)[rule_format,OF b(1),unfolded kle_def] .. note k = this | |
| 1027 | fix i | |
| 1028 |     show "b i = (if i \<in> {1..n} \<and> a i < b i then a i + 1 else a i)"
 | |
| 1029 |     proof (cases "i \<in> {1..n}")
 | |
| 1030 | case True | |
| 53688 | 1031 | then show ?thesis | 
| 1032 | unfolding k[THEN conjunct2,rule_format] by auto | |
| 53185 | 1033 | next | 
| 1034 | case False | |
| 53688 | 1035 | have "a i = p" | 
| 1036 | using assm and False `a\<in>s` by auto | |
| 1037 | moreover have "b i = p" | |
| 1038 | using assm and False `b\<in>s` by auto | |
| 1039 | ultimately show ?thesis | |
| 1040 | by auto | |
| 53185 | 1041 | qed | 
| 53688 | 1042 | qed (insert a(2) b(2) assm, auto) | 
| 53185 | 1043 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1044 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1045 | lemma ksimplex_extrema_strong: | 
| 53688 | 1046 | assumes "ksimplex p n s" | 
| 1047 | and "n \<noteq> 0" | |
| 1048 | obtains a b where "a \<in> s" | |
| 1049 | and "b \<in> s" | |
| 1050 | and "a \<noteq> b" | |
| 1051 | and "\<forall>x\<in>s. kle n a x \<and> kle n x b" | |
| 1052 |     and "\<forall>i. b i = (if i \<in> {1..n} then a(i) + 1 else a i)"
 | |
| 53185 | 1053 | proof - | 
| 1054 | obtain a b where ab: "a \<in> s" "b \<in> s" | |
| 1055 |     "\<forall>x\<in>s. kle n a x \<and> kle n x b" "\<forall>i. b(i) = (if i \<in> {1..n} then a(i) + 1 else a(i))"
 | |
| 1056 | apply (rule ksimplex_extrema[OF assms(1)]) | |
| 1057 | apply auto | |
| 1058 | done | |
| 1059 | have "a \<noteq> b" | |
| 1060 | apply (rule notI) | |
| 1061 | apply (drule cong[of _ _ 1 1]) | |
| 53688 | 1062 | using ab(4) assms(2) | 
| 1063 | apply auto | |
| 53185 | 1064 | done | 
| 53674 | 1065 | then show ?thesis | 
| 53185 | 1066 | apply (rule_tac that[of a b]) | 
| 53688 | 1067 | using ab | 
| 1068 | apply auto | |
| 53185 | 1069 | done | 
| 1070 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1071 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1072 | lemma ksimplexD: | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1073 | assumes "ksimplex p n s" | 
| 53688 | 1074 | shows "card s = n + 1" | 
| 1075 | and "finite s" | |
| 1076 | and "card s = n + 1" | |
| 1077 | and "\<forall>x\<in>s. \<forall>j. x j \<le> p" | |
| 1078 |     and "\<forall>x\<in>s. \<forall>j. j \<notin> {1..n} \<longrightarrow> x j = p"
 | |
| 1079 | and "\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x" | |
| 53185 | 1080 | using assms unfolding ksimplex_eq 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 | 1081 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1082 | lemma ksimplex_successor: | 
| 53688 | 1083 | assumes "ksimplex p n s" | 
| 1084 | and "a \<in> s" | |
| 1085 |   shows "(\<forall>x\<in>s. kle n x a) \<or> (\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. y j = (if j = k then a j + 1 else a j))"
 | |
| 53185 | 1086 | proof (cases "\<forall>x\<in>s. kle n x a") | 
| 1087 | case True | |
| 53674 | 1088 | then show ?thesis by auto | 
| 53185 | 1089 | next | 
| 1090 | note assm = ksimplexD[OF assms(1)] | |
| 1091 | case False | |
| 53688 | 1092 |   then obtain b where b: "b \<in> s" "\<not> kle n b a" "\<forall>x\<in>{x \<in> s. \<not> kle n x a}. kle n b x"
 | 
| 1093 |     using kle_minimal[of "{x\<in>s. \<not> kle n x a}" n] and assm
 | |
| 1094 | by auto | |
| 53252 | 1095 |   then have **: "1 \<le> card {k\<in>{1..n}. a k < b k}"
 | 
| 53185 | 1096 | apply - | 
| 1097 | apply (rule kle_strict_set) | |
| 53252 | 1098 | using assm(6) and `a\<in>s` | 
| 53688 | 1099 | apply (auto simp add: kle_refl) | 
| 53185 | 1100 | done | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1101 | |
| 53185 | 1102 |   let ?kle1 = "{x \<in> s. \<not> kle n x a}"
 | 
| 1103 | have "card ?kle1 > 0" | |
| 1104 | apply (rule ccontr) | |
| 53252 | 1105 | using assm(2) and False | 
| 1106 | apply auto | |
| 53185 | 1107 | done | 
| 53674 | 1108 | then have sizekle1: "card ?kle1 = Suc (card ?kle1 - 1)" | 
| 53185 | 1109 | using assm(2) by auto | 
| 1110 | obtain c d where c_d: "c \<in> s" "\<not> kle n c a" "d \<in> s" "\<not> kle n d a" | |
| 1111 |     "kle n c d" "card ?kle1 - 1 \<le> card {k \<in> {1..n}. c k < d k}"
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1112 | using kle_range_induct[OF sizekle1, of n] using assm by auto | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1113 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1114 |   let ?kle2 = "{x \<in> s. kle n x a}"
 | 
| 53185 | 1115 | have "card ?kle2 > 0" | 
| 1116 | apply (rule ccontr) | |
| 53252 | 1117 | using assm(6)[rule_format,of a a] and `a\<in>s` and assm(2) | 
| 1118 | apply auto | |
| 53185 | 1119 | done | 
| 53674 | 1120 | then have sizekle2: "card ?kle2 = Suc (card ?kle2 - 1)" | 
| 53185 | 1121 | using assm(2) by auto | 
| 1122 | obtain e f where e_f: "e \<in> s" "kle n e a" "f \<in> s" "kle n f a" | |
| 1123 |     "kle n e f" "card ?kle2 - 1 \<le> card {k \<in> {1..n}. e k < f k}"
 | |
| 53688 | 1124 | using kle_range_induct[OF sizekle2, of n] | 
| 1125 | using assm | |
| 1126 | 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 | 1127 | |
| 53185 | 1128 |   have "card {k\<in>{1..n}. a k < b k} = 1"
 | 
| 1129 | proof (rule ccontr) | |
| 53688 | 1130 | assume "\<not> ?thesis" | 
| 53674 | 1131 |     then have as: "card {k\<in>{1..n}. a k < b k} \<ge> 2"
 | 
| 53185 | 1132 | using ** by auto | 
| 1133 |     have *: "finite ?kle2" "finite ?kle1" "?kle2 \<union> ?kle1 = s" "?kle2 \<inter> ?kle1 = {}"
 | |
| 1134 | using assm(2) by auto | |
| 1135 | have "(card ?kle2 - 1) + 2 + (card ?kle1 - 1) = card ?kle2 + card ?kle1" | |
| 1136 | using sizekle1 sizekle2 by auto | |
| 1137 | also have "\<dots> = n + 1" | |
| 53688 | 1138 | unfolding card_Un_Int[OF *(1-2)] *(3-) | 
| 1139 | using assm(3) | |
| 1140 | by auto | |
| 53185 | 1141 | finally have n: "(card ?kle2 - 1) + (2 + (card ?kle1 - 1)) = n + 1" | 
| 1142 | 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 | 1143 |     have "kle n e a \<and> card {x \<in> s. kle n x a} - 1 \<le> card {k \<in> {1..n}. e k < a k}"
 | 
| 53185 | 1144 | apply (rule kle_range_combine_r[where y=f]) | 
| 53688 | 1145 | using e_f using `a \<in> s` assm(6) | 
| 53252 | 1146 | apply auto | 
| 53185 | 1147 | done | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1148 |     moreover have "kle n b d \<and> card {x \<in> s. \<not> kle n x a} - 1 \<le> card {k \<in> {1..n}. b k < d k}"
 | 
| 53185 | 1149 | apply (rule kle_range_combine_l[where y=c]) | 
| 53252 | 1150 | using c_d using assm(6) and b | 
| 1151 | apply auto | |
| 53185 | 1152 | done | 
| 53674 | 1153 |     then have "kle n a d \<and> 2 + (card {x \<in> s. \<not> kle n x a} - 1) \<le> card {k \<in> {1..n}. a k < d k}"
 | 
| 53185 | 1154 | apply - | 
| 53688 | 1155 | apply (rule kle_range_combine[where y=b]) using as and b assm(6) `a \<in> s` `d \<in> s` | 
| 53185 | 1156 | apply blast+ | 
| 1157 | done | |
| 1158 | ultimately | |
| 1159 |     have "kle n e d \<and> (card ?kle2 - 1) + (2 + (card ?kle1 - 1)) \<le> card {k\<in>{1..n}. e k < d k}"
 | |
| 1160 | apply - | |
| 53688 | 1161 | apply (rule kle_range_combine[where y=a]) using assm(6)[rule_format, OF `e \<in> s` `d \<in> s`] | 
| 53185 | 1162 | apply blast+ | 
| 1163 | done | |
| 1164 |     moreover have "card {k \<in> {1..n}. e k < d k} \<le> card {1..n}"
 | |
| 1165 | by (rule card_mono) auto | |
| 53688 | 1166 | ultimately show False | 
| 1167 | unfolding n by auto | |
| 53185 | 1168 | qed | 
| 1169 | then guess k unfolding card_1_exists .. note k = this[unfolded mem_Collect_eq] | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1170 | |
| 53185 | 1171 | show ?thesis | 
| 1172 | apply (rule disjI2) | |
| 53688 | 1173 | apply (rule_tac x=b in bexI) | 
| 1174 | apply (rule_tac x=k in bexI) | |
| 53185 | 1175 | proof | 
| 1176 | fix j :: nat | |
| 1177 | have "kle n a b" | |
| 53688 | 1178 | using b and assm(6)[rule_format, OF `a\<in>s` `b\<in>s`] | 
| 1179 | by auto | |
| 53185 | 1180 | then guess kk unfolding kle_def .. note kk_raw = this | 
| 1181 | note kk = this[THEN conjunct2, rule_format] | |
| 1182 | have kkk: "k \<in> kk" | |
| 1183 | apply (rule ccontr) | |
| 1184 | using k(1) | |
| 53252 | 1185 | unfolding kk | 
| 1186 | apply auto | |
| 53185 | 1187 | done | 
| 1188 | show "b j = (if j = k then a j + 1 else a j)" | |
| 1189 | proof (cases "j \<in> kk") | |
| 1190 | case True | |
| 53252 | 1191 | then have "j = k" | 
| 1192 | apply - | |
| 1193 | apply (rule k(2)[rule_format]) | |
| 1194 | using kk_raw kkk | |
| 1195 | apply auto | |
| 1196 | done | |
| 53688 | 1197 | then show ?thesis | 
| 1198 | unfolding kk using kkk by auto | |
| 53185 | 1199 | next | 
| 1200 | case False | |
| 53252 | 1201 | then have "j \<noteq> k" | 
| 53688 | 1202 | using k(2)[rule_format, of j k] and kk_raw kkk | 
| 1203 | by auto | |
| 1204 | then show ?thesis | |
| 1205 | unfolding kk | |
| 1206 | using kkk and False | |
| 53185 | 1207 | by auto | 
| 1208 | qed | |
| 53688 | 1209 | qed (insert k(1) `b \<in> s`, auto) | 
| 53185 | 1210 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1211 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1212 | lemma ksimplex_predecessor: | 
| 53688 | 1213 | assumes "ksimplex p n s" | 
| 1214 | and "a \<in> s" | |
| 1215 |   shows "(\<forall>x\<in>s. kle n a x) \<or> (\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. a j = (if j = k then y j + 1 else y j))"
 | |
| 53185 | 1216 | proof (cases "\<forall>x\<in>s. kle n a x") | 
| 1217 | case True | |
| 53674 | 1218 | then show ?thesis by auto | 
| 53185 | 1219 | next | 
| 1220 | note assm = ksimplexD[OF assms(1)] | |
| 1221 | case False | |
| 53688 | 1222 |   then obtain b where b: "b \<in> s" "\<not> kle n a b" "\<forall>x\<in>{x \<in> s. \<not> kle n a x}. kle n x b"
 | 
| 1223 |     using kle_maximal[of "{x\<in>s. \<not> kle n a x}" n] and assm
 | |
| 1224 | by auto | |
| 53674 | 1225 |   then have **: "1 \<le> card {k\<in>{1..n}. a k > b k}"
 | 
| 53185 | 1226 | apply - | 
| 1227 | apply (rule kle_strict_set) | |
| 53688 | 1228 | using assm(6) and `a \<in> s` | 
| 53252 | 1229 | apply (auto simp add: kle_refl) | 
| 53185 | 1230 | done | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1231 | |
| 53185 | 1232 |   let ?kle1 = "{x \<in> s. \<not> kle n a x}"
 | 
| 1233 | have "card ?kle1 > 0" | |
| 1234 | apply (rule ccontr) | |
| 53252 | 1235 | using assm(2) and False | 
| 1236 | apply auto | |
| 53185 | 1237 | done | 
| 53674 | 1238 | then have sizekle1: "card ?kle1 = Suc (card ?kle1 - 1)" | 
| 53185 | 1239 | using assm(2) by auto | 
| 1240 | obtain c d where c_d: "c \<in> s" "\<not> kle n a c" "d \<in> s" "\<not> kle n a d" | |
| 1241 |     "kle n d c" "card ?kle1 - 1 \<le> card {k \<in> {1..n}. c k > d k}"
 | |
| 53688 | 1242 | using kle_range_induct[OF sizekle1, of n] | 
| 1243 | using assm | |
| 1244 | 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 | 1245 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1246 |   let ?kle2 = "{x \<in> s. kle n a x}"
 | 
| 53185 | 1247 | have "card ?kle2 > 0" | 
| 1248 | apply (rule ccontr) | |
| 53688 | 1249 | using assm(6)[rule_format,of a a] and `a \<in> s` and assm(2) | 
| 53252 | 1250 | apply auto | 
| 53185 | 1251 | done | 
| 53688 | 1252 | then have sizekle2: "card ?kle2 = Suc (card ?kle2 - 1)" | 
| 1253 | using assm(2) | |
| 1254 | by auto | |
| 53185 | 1255 | obtain e f where e_f: "e \<in> s" "kle n a e" "f \<in> s" "kle n a f" | 
| 1256 |     "kle n f e" "card ?kle2 - 1 \<le> card {k \<in> {1..n}. e k > f k}"
 | |
| 53688 | 1257 | using kle_range_induct[OF sizekle2, of n] | 
| 1258 | using assm | |
| 1259 | 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 | 1260 | |
| 53185 | 1261 |   have "card {k\<in>{1..n}. a k > b k} = 1"
 | 
| 1262 | proof (rule ccontr) | |
| 53688 | 1263 | assume "\<not> ?thesis" | 
| 53674 | 1264 |     then have as: "card {k\<in>{1..n}. a k > b k} \<ge> 2"
 | 
| 53185 | 1265 | using ** by auto | 
| 1266 |     have *: "finite ?kle2" "finite ?kle1" "?kle2 \<union> ?kle1 = s" "?kle2 \<inter> ?kle1 = {}"
 | |
| 1267 | using assm(2) by auto | |
| 1268 | have "(card ?kle2 - 1) + 2 + (card ?kle1 - 1) = card ?kle2 + card ?kle1" | |
| 1269 | using sizekle1 sizekle2 by auto | |
| 1270 | also have "\<dots> = n + 1" | |
| 53688 | 1271 | unfolding card_Un_Int[OF *(1-2)] *(3-) | 
| 1272 | using assm(3) by auto | |
| 53185 | 1273 | finally have n: "(card ?kle1 - 1) + 2 + (card ?kle2 - 1) = n + 1" | 
| 1274 | 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 | 1275 |     have "kle n a e \<and> card {x \<in> s. kle n a x} - 1 \<le> card {k \<in> {1..n}. e k > a k}"
 | 
| 53185 | 1276 | apply (rule kle_range_combine_l[where y=f]) | 
| 53252 | 1277 | using e_f and `a\<in>s` assm(6) | 
| 1278 | apply auto | |
| 53185 | 1279 | done | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1280 |     moreover have "kle n d b \<and> card {x \<in> s. \<not> kle n a x} - 1 \<le> card {k \<in> {1..n}. b k > d k}"
 | 
| 53185 | 1281 | apply (rule kle_range_combine_r[where y=c]) | 
| 53252 | 1282 | using c_d and assm(6) and b | 
| 1283 | apply auto | |
| 53185 | 1284 | done | 
| 53674 | 1285 |     then have "kle n d a \<and> (card {x \<in> s. \<not> kle n a x} - 1) + 2 \<le> card {k \<in> {1..n}. a k > d k}"
 | 
| 53185 | 1286 | apply - | 
| 53688 | 1287 | apply (rule kle_range_combine[where y=b]) using as and b assm(6) `a \<in> s` `d \<in> s` | 
| 53185 | 1288 | apply blast+ | 
| 1289 | done | |
| 1290 |     ultimately have "kle n d e \<and> (card ?kle1 - 1 + 2) + (card ?kle2 - 1) \<le> card {k\<in>{1..n}. e k > d k}"
 | |
| 1291 | apply - | |
| 1292 | apply (rule kle_range_combine[where y=a]) | |
| 53252 | 1293 | using assm(6)[rule_format,OF `e\<in>s` `d\<in>s`] | 
| 1294 | apply blast+ | |
| 53185 | 1295 | done | 
| 1296 |     moreover have "card {k \<in> {1..n}. e k > d k} \<le> card {1..n}"
 | |
| 1297 | by (rule card_mono) auto | |
| 53688 | 1298 | ultimately show False | 
| 1299 | unfolding n by auto | |
| 53185 | 1300 | qed | 
| 1301 | then guess k unfolding card_1_exists .. note k = this[unfolded mem_Collect_eq] | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1302 | |
| 53185 | 1303 | show ?thesis | 
| 1304 | apply (rule disjI2) | |
| 53688 | 1305 | apply (rule_tac x=b in bexI) | 
| 1306 | apply (rule_tac x=k in bexI) | |
| 53185 | 1307 | proof | 
| 1308 | fix j :: nat | |
| 1309 | have "kle n b a" | |
| 1310 | using b and assm(6)[rule_format, OF `a\<in>s` `b\<in>s`] by auto | |
| 1311 | then guess kk unfolding kle_def .. note kk_raw = this | |
| 1312 | note kk = this[THEN conjunct2,rule_format] | |
| 1313 | have kkk: "k \<in> kk" | |
| 1314 | apply (rule ccontr) | |
| 1315 | using k(1) | |
| 1316 | unfolding kk | |
| 1317 | apply auto | |
| 1318 | done | |
| 1319 | show "a j = (if j = k then b j + 1 else b j)" | |
| 1320 | proof (cases "j \<in> kk") | |
| 1321 | case True | |
| 53674 | 1322 | then have "j = k" | 
| 53185 | 1323 | apply - | 
| 1324 | apply (rule k(2)[rule_format]) | |
| 53252 | 1325 | using kk_raw kkk | 
| 1326 | apply auto | |
| 53185 | 1327 | done | 
| 53688 | 1328 | then show ?thesis | 
| 1329 | unfolding kk using kkk by auto | |
| 53185 | 1330 | next | 
| 1331 | case False | |
| 53688 | 1332 | then have "j \<noteq> k" | 
| 1333 | using k(2)[rule_format, of j k] | |
| 1334 | using kk_raw kkk | |
| 1335 | by auto | |
| 1336 | then show ?thesis | |
| 1337 | unfolding kk | |
| 1338 | using kkk and False | |
| 1339 | by auto | |
| 53185 | 1340 | qed | 
| 53186 | 1341 | qed (insert k(1) `b\<in>s`, auto) | 
| 53185 | 1342 | qed | 
| 1343 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1344 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1345 | subsection {* The lemmas about simplices that we need. *}
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1346 | |
| 53185 | 1347 | (* FIXME: These are clones of lemmas in Library/FuncSet *) | 
| 1348 | lemma card_funspace': | |
| 53688 | 1349 | assumes "finite s" | 
| 1350 | and "finite t" | |
| 1351 | and "card s = m" | |
| 1352 | and "card t = n" | |
| 1353 |   shows "card {f. (\<forall>x\<in>s. f x \<in> t) \<and> (\<forall>x\<in>UNIV - s. f x = d)} = n ^ m"
 | |
| 1354 | (is "card (?M s) = _") | |
| 53185 | 1355 | using assms | 
| 1356 | proof (induct m arbitrary: s) | |
| 53252 | 1357 | case 0 | 
| 1358 |   have [simp]: "{f. \<forall>x. f x = d} = {\<lambda>x. d}"
 | |
| 53688 | 1359 | apply (rule set_eqI) | 
| 1360 | apply rule | |
| 53185 | 1361 | unfolding mem_Collect_eq | 
| 53688 | 1362 | apply rule | 
| 1363 | apply (rule ext) | |
| 53185 | 1364 | apply auto | 
| 1365 | done | |
| 53688 | 1366 | from 0 show ?case | 
| 1367 | by auto | |
| 53185 | 1368 | next | 
| 1369 | case (Suc m) | |
| 1370 | guess a using card_eq_SucD[OF Suc(4)] .. | |
| 1371 | then guess s0 by (elim exE conjE) note as0 = this | |
| 1372 | have **: "card s0 = m" | |
| 53688 | 1373 | using as0 using Suc(2) Suc(4) | 
| 1374 | by auto | |
| 53185 | 1375 | let ?l = "(\<lambda>(b, g) x. if x = a then b else g x)" | 
| 1376 |   have *: "?M (insert a s0) = ?l ` {(b,g). b\<in>t \<and> g\<in>?M s0}"
 | |
| 1377 | apply (rule set_eqI, rule) | |
| 1378 | unfolding mem_Collect_eq image_iff | |
| 1379 | apply (erule conjE) | |
| 1380 | apply (rule_tac x="(x a, \<lambda>y. if y\<in>s0 then x y else d)" in bexI) | |
| 1381 | apply (rule ext) | |
| 1382 | prefer 3 | |
| 1383 | apply rule | |
| 1384 | defer | |
| 53688 | 1385 | apply (erule bexE) | 
| 1386 | apply rule | |
| 53185 | 1387 | unfolding mem_Collect_eq | 
| 1388 | apply (erule splitE)+ | |
| 1389 | apply (erule conjE)+ | |
| 1390 | proof - | |
| 1391 | fix x xa xb xc y | |
| 53688 | 1392 | assume as: | 
| 1393 | "x = (\<lambda>(b, g) x. if x = a then b else g x) xa" | |
| 53846 | 1394 | "xb \<in> UNIV - insert a s0" | 
| 1395 | "xa = (xc, y)" | |
| 53688 | 1396 | "xc \<in> t" | 
| 1397 | "\<forall>x\<in>s0. y x \<in> t" | |
| 1398 | "\<forall>x\<in>UNIV - s0. y x = d" | |
| 1399 | then show "x xb = d" | |
| 1400 | unfolding as by auto | |
| 53185 | 1401 | qed auto | 
| 1402 |   have inj: "inj_on ?l {(b,g). b\<in>t \<and> g\<in>?M s0}"
 | |
| 1403 | unfolding inj_on_def | |
| 53688 | 1404 | apply rule | 
| 1405 | apply rule | |
| 1406 | apply rule | |
| 1407 | unfolding mem_Collect_eq | |
| 1408 | apply (erule splitE conjE)+ | |
| 53185 | 1409 | proof - | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1410 | case goal1 note as = this(1,4-)[unfolded goal1 split_conv] | 
| 53688 | 1411 | have "xa = xb" | 
| 1412 | using as(1)[THEN cong[of _ _ a]] by auto | |
| 53185 | 1413 | moreover have "ya = yb" | 
| 1414 | proof (rule ext) | |
| 1415 | fix x | |
| 1416 | show "ya x = yb x" | |
| 1417 | proof (cases "x = a") | |
| 1418 | case False | |
| 53688 | 1419 | then show ?thesis | 
| 1420 | using as(1)[THEN cong[of _ _ x x]] by auto | |
| 53185 | 1421 | next | 
| 1422 | case True | |
| 53688 | 1423 | then show ?thesis | 
| 1424 | using as(5,7) using as0(2) by auto | |
| 53185 | 1425 | qed | 
| 1426 | qed | |
| 53688 | 1427 | ultimately show ?case | 
| 1428 | unfolding goal1 by auto | |
| 53185 | 1429 | qed | 
| 53688 | 1430 | have "finite s0" | 
| 1431 | using `finite s` unfolding as0 by simp | |
| 53185 | 1432 | show ?case | 
| 1433 | unfolding as0 * card_image[OF inj] | |
| 1434 | using assms | |
| 1435 | unfolding SetCompr_Sigma_eq | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1436 | unfolding card_cartesian_product | 
| 53185 | 1437 | using Suc(1)[OF `finite s0` `finite t` ** `card t = n`] | 
| 1438 | 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 | 1439 | qed | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1440 | |
| 53185 | 1441 | lemma card_funspace: | 
| 53688 | 1442 | assumes "finite s" | 
| 1443 | and "finite t" | |
| 1444 |   shows "card {f. (\<forall>x\<in>s. f x \<in> t) \<and> (\<forall>x\<in>UNIV - s. f x = d)} = card t ^ card s"
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1445 | using assms by (auto intro: card_funspace') | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1446 | |
| 53185 | 1447 | lemma finite_funspace: | 
| 53688 | 1448 | assumes "finite s" | 
| 1449 | and "finite t" | |
| 1450 |   shows "finite {f. (\<forall>x\<in>s. f x \<in> t) \<and> (\<forall>x\<in>UNIV - s. f x = d)}"
 | |
| 1451 | (is "finite ?S") | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1452 | proof (cases "card t > 0") | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1453 | case True | 
| 53688 | 1454 | have "card ?S = card t ^ card s" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1455 | using assms by (auto intro!: card_funspace) | 
| 53688 | 1456 | then show ?thesis | 
| 1457 | using True by (rule_tac card_ge_0_finite) simp | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1458 | next | 
| 53688 | 1459 | case False | 
| 1460 |   then have "t = {}"
 | |
| 1461 | using `finite t` 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 | 1462 | show ?thesis | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1463 |   proof (cases "s = {}")
 | 
| 53185 | 1464 | case True | 
| 53688 | 1465 |     have *: "{f. \<forall>x. f x = d} = {\<lambda>x. d}"
 | 
| 1466 | by auto | |
| 1467 | from True show ?thesis | |
| 1468 |       using `t = {}` by (auto simp: *)
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1469 | next | 
| 53185 | 1470 | case False | 
| 53688 | 1471 | then show ?thesis | 
| 1472 |       using `t = {}` by simp
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1473 | qed | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1474 | qed | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1475 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1476 | lemma finite_simplices: "finite {s. ksimplex p n s}"
 | 
| 53185 | 1477 |   apply (rule finite_subset[of _ "{s. s\<subseteq>{f. (\<forall>i\<in>{1..n}. f i \<in> {0..p}) \<and> (\<forall>i\<in>UNIV-{1..n}. f i = p)}}"])
 | 
| 1478 | unfolding ksimplex_def | |
| 1479 | defer | |
| 1480 | apply (rule finite_Collect_subsets) | |
| 1481 | apply (rule finite_funspace) | |
| 1482 | apply auto | |
| 1483 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1484 | |
| 53185 | 1485 | lemma simplex_top_face: | 
| 53688 | 1486 | assumes "0 < p" | 
| 1487 | and "\<forall>x\<in>f. x (n + 1) = p" | |
| 1488 |   shows "(\<exists>s a. ksimplex p (n + 1) s \<and> a \<in> s \<and> (f = s - {a})) \<longleftrightarrow> ksimplex p n f"
 | |
| 1489 | (is "?ls = ?rs") | |
| 53185 | 1490 | proof | 
| 1491 | assume ?ls | |
| 1492 | then guess s .. | |
| 1493 | then guess a by (elim exE conjE) note sa = this | |
| 1494 | show ?rs | |
| 1495 | unfolding ksimplex_def sa(3) | |
| 1496 | apply rule | |
| 1497 | defer | |
| 1498 | apply rule | |
| 1499 | defer | |
| 1500 | apply (rule, rule, rule, rule) | |
| 1501 | defer | |
| 1502 | apply (rule, rule) | |
| 1503 | proof - | |
| 1504 | fix x y | |
| 1505 |     assume as: "x \<in>s - {a}" "y \<in>s - {a}"
 | |
| 1506 | have xyp: "x (n + 1) = y (n + 1)" | |
| 1507 | using as(1)[unfolded sa(3)[symmetric], THEN assms(2)[rule_format]] | |
| 1508 | using as(2)[unfolded sa(3)[symmetric], THEN assms(2)[rule_format]] | |
| 1509 | by auto | |
| 1510 | show "kle n x y \<or> kle n y x" | |
| 1511 | proof (cases "kle (n + 1) x y") | |
| 1512 | case True | |
| 1513 | then guess k unfolding kle_def .. note k = this | |
| 53674 | 1514 | then have *: "n + 1 \<notin> k" using xyp by auto | 
| 53688 | 1515 |       have "\<not> (\<exists>x\<in>k. x \<notin> {1..n})"
 | 
| 53252 | 1516 | apply (rule notI) | 
| 1517 | apply (erule bexE) | |
| 53185 | 1518 | proof - | 
| 1519 | fix x | |
| 1520 |         assume as: "x \<in> k" "x \<notin> {1..n}"
 | |
| 53688 | 1521 | have "x \<noteq> n + 1" | 
| 1522 | using as and * by auto | |
| 1523 | then show False | |
| 1524 | using as and k[THEN conjunct1,unfolded subset_eq] by auto | |
| 53185 | 1525 | qed | 
| 53674 | 1526 | then show ?thesis | 
| 53185 | 1527 | apply - | 
| 1528 | apply (rule disjI1) | |
| 1529 | unfolding kle_def | |
| 1530 | using k | |
| 1531 | apply (rule_tac x=k in exI) | |
| 1532 | apply auto | |
| 1533 | done | |
| 1534 | next | |
| 1535 | case False | |
| 53674 | 1536 | then have "kle (n + 1) y x" | 
| 53688 | 1537 | using ksimplexD(6)[OF sa(1),rule_format, of x y] and as | 
| 1538 | by auto | |
| 53185 | 1539 | then guess k unfolding kle_def .. note k = this | 
| 53688 | 1540 | then have *: "n + 1 \<notin> k" | 
| 1541 | using xyp by auto | |
| 1542 |       then have "\<not> (\<exists>x\<in>k. x \<notin> {1..n})"
 | |
| 53185 | 1543 | apply - | 
| 53252 | 1544 | apply (rule notI) | 
| 53185 | 1545 | apply (erule bexE) | 
| 1546 | proof - | |
| 1547 | fix x | |
| 1548 |         assume as: "x \<in> k" "x \<notin> {1..n}"
 | |
| 53688 | 1549 | have "x \<noteq> n + 1" | 
| 1550 | using as and * by auto | |
| 1551 | then show False | |
| 1552 | using as and k[THEN conjunct1,unfolded subset_eq] | |
| 1553 | by auto | |
| 53185 | 1554 | qed | 
| 53252 | 1555 | then show ?thesis | 
| 53185 | 1556 | apply - | 
| 1557 | apply (rule disjI2) | |
| 1558 | unfolding kle_def | |
| 53252 | 1559 | using k | 
| 1560 | apply (rule_tac x = k in exI) | |
| 1561 | apply auto | |
| 1562 | done | |
| 53185 | 1563 | qed | 
| 1564 | next | |
| 1565 | fix x j | |
| 53688 | 1566 |     assume as: "x \<in> s - {a}" "j \<notin> {1..n}"
 | 
| 53674 | 1567 | then show "x j = p" | 
| 53185 | 1568 | using as(1)[unfolded sa(3)[symmetric], THEN assms(2)[rule_format]] | 
| 53688 | 1569 | apply (cases "j = n + 1") | 
| 53185 | 1570 | using sa(1)[unfolded ksimplex_def] | 
| 1571 | apply auto | |
| 1572 | done | |
| 1573 | qed (insert sa ksimplexD[OF sa(1)], auto) | |
| 1574 | next | |
| 1575 | assume ?rs note rs=ksimplexD[OF this] | |
| 53252 | 1576 | guess a b by (rule ksimplex_extrema[OF `?rs`]) note ab = this | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1577 | def c \<equiv> "\<lambda>i. if i = (n + 1) then p - 1 else a i" | 
| 53185 | 1578 | have "c \<notin> f" | 
| 53252 | 1579 | apply (rule notI) | 
| 53185 | 1580 | apply (drule assms(2)[rule_format]) | 
| 1581 | unfolding c_def | |
| 53252 | 1582 | using assms(1) | 
| 1583 | apply auto | |
| 53185 | 1584 | done | 
| 53674 | 1585 | then show ?ls | 
| 53688 | 1586 | apply (rule_tac x = "insert c f" in exI) | 
| 1587 | apply (rule_tac x = c in exI) | |
| 53185 | 1588 | unfolding ksimplex_def conj_assoc | 
| 1589 | apply (rule conjI) | |
| 1590 | defer | |
| 1591 | apply (rule conjI) | |
| 1592 | defer | |
| 1593 | apply (rule conjI) | |
| 1594 | defer | |
| 1595 | apply (rule conjI) | |
| 1596 | defer | |
| 1597 | proof (rule_tac[3-5] ballI allI)+ | |
| 1598 | fix x j | |
| 1599 | assume x: "x \<in> insert c f" | |
| 53674 | 1600 | then show "x j \<le> p" | 
| 53688 | 1601 | proof (cases "x = c") | 
| 53185 | 1602 | case True | 
| 1603 | show ?thesis | |
| 1604 | unfolding True c_def | |
| 53688 | 1605 | apply (cases "j = n + 1") | 
| 53252 | 1606 | using ab(1) and rs(4) | 
| 1607 | apply auto | |
| 53185 | 1608 | done | 
| 53688 | 1609 | next | 
| 1610 | case False | |
| 1611 | with insert x rs(4) show ?thesis | |
| 1612 | by (auto simp add: c_def) | |
| 1613 | qed | |
| 53185 | 1614 |     show "j \<notin> {1..n + 1} \<longrightarrow> x j = p"
 | 
| 1615 | apply (cases "x = c") | |
| 53252 | 1616 | using x ab(1) rs(5) | 
| 1617 | unfolding c_def | |
| 1618 | apply auto | |
| 1619 | done | |
| 53185 | 1620 |     {
 | 
| 1621 | fix z | |
| 1622 | assume z: "z \<in> insert c f" | |
| 53674 | 1623 | then have "kle (n + 1) c z" | 
| 53688 | 1624 | proof (cases "z = c") | 
| 53185 | 1625 | case False | 
| 53688 | 1626 | then have "z \<in> f" | 
| 1627 | using z by auto | |
| 53185 | 1628 | then guess k | 
| 1629 | apply (drule_tac ab(3)[THEN bspec[where x=z], THEN conjunct1]) | |
| 1630 | unfolding kle_def | |
| 1631 | apply (erule exE) | |
| 1632 | done | |
| 53674 | 1633 | then show "kle (n + 1) c z" | 
| 53185 | 1634 | unfolding kle_def | 
| 1635 | apply (rule_tac x="insert (n + 1) k" in exI) | |
| 1636 | unfolding c_def | |
| 1637 | using ab | |
| 1638 | using rs(5)[rule_format,OF ab(1),of "n + 1"] assms(1) | |
| 1639 | apply auto | |
| 1640 | done | |
| 53688 | 1641 | next | 
| 1642 | case True | |
| 1643 | then show ?thesis by auto | |
| 1644 | qed | |
| 53185 | 1645 | } note * = this | 
| 1646 | fix y | |
| 1647 | assume y: "y \<in> insert c f" | |
| 1648 | show "kle (n + 1) x y \<or> kle (n + 1) y x" | |
| 1649 | proof (cases "x = c \<or> y = c") | |
| 53688 | 1650 | case False | 
| 1651 | then have **: "x \<in> f" "y \<in> f" | |
| 1652 | using x y by auto | |
| 1653 | show ?thesis | |
| 1654 | using rs(6)[rule_format,OF **] | |
| 53252 | 1655 | by (auto dest: kle_Suc) | 
| 53185 | 1656 | qed (insert * x y, auto) | 
| 1657 | qed (insert rs, auto) | |
| 1658 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1659 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1660 | lemma ksimplex_fix_plane: | 
| 53688 | 1661 | fixes a a0 a1 :: "nat \<Rightarrow> nat" | 
| 1662 | assumes "a \<in> s" | |
| 1663 |     and "j \<in> {1..n}"
 | |
| 1664 |     and "\<forall>x\<in>s - {a}. x j = q"
 | |
| 1665 | and "a0 \<in> s" | |
| 1666 | and "a1 \<in> s" | |
| 1667 |     and "\<forall>i. a1 i = (if i \<in> {1..n} then a0 i + 1 else a0 i)"
 | |
| 1668 | shows "a = a0 \<or> a = a1" | |
| 53185 | 1669 | proof - | 
| 1670 | have *: "\<And>P A x y. \<forall>x\<in>A. P x \<Longrightarrow> x\<in>A \<Longrightarrow> y\<in>A \<Longrightarrow> P x \<and> P y" | |
| 1671 | by auto | |
| 1672 | show ?thesis | |
| 1673 | apply (rule ccontr) | |
| 1674 | using *[OF assms(3), of a0 a1] | |
| 1675 | unfolding assms(6)[THEN spec[where x=j]] | |
| 53252 | 1676 | using assms(1-2,4-5) | 
| 1677 | apply auto | |
| 1678 | done | |
| 53185 | 1679 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1680 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1681 | lemma ksimplex_fix_plane_0: | 
| 53688 | 1682 | fixes a a0 a1 :: "nat \<Rightarrow> nat" | 
| 1683 | assumes "a \<in> s" | |
| 1684 |     and "j \<in> {1..n}"
 | |
| 1685 |     and "\<forall>x\<in>s - {a}. x j = 0"
 | |
| 1686 | and "a0 \<in> s" | |
| 1687 | and "a1 \<in> s" | |
| 1688 |     and "\<forall>i. a1 i = (if i\<in>{1..n} then a0 i + 1 else a0 i)"
 | |
| 53185 | 1689 | shows "a = a1" | 
| 1690 | apply (rule ccontr) | |
| 1691 | using ksimplex_fix_plane[OF assms] | |
| 1692 | using assms(3)[THEN bspec[where x=a1]] | |
| 1693 | using assms(2,5) | |
| 1694 | unfolding assms(6)[THEN spec[where x=j]] | |
| 1695 | apply simp | |
| 1696 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1697 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1698 | lemma ksimplex_fix_plane_p: | 
| 53688 | 1699 | assumes "ksimplex p n s" | 
| 1700 | and "a \<in> s" | |
| 1701 |     and "j \<in> {1..n}"
 | |
| 1702 |     and "\<forall>x\<in>s - {a}. x j = p"
 | |
| 1703 | and "a0 \<in> s" | |
| 1704 | and "a1 \<in> s" | |
| 1705 |     and "\<forall>i. a1 i = (if i\<in>{1..n} then a0 i + 1 else a0 i)"
 | |
| 53185 | 1706 | shows "a = a0" | 
| 1707 | proof (rule ccontr) | |
| 1708 | note s = ksimplexD[OF assms(1),rule_format] | |
| 53688 | 1709 | assume as: "\<not> ?thesis" | 
| 53252 | 1710 |   then have *: "a0 \<in> s - {a}"
 | 
| 53185 | 1711 | using assms(5) by auto | 
| 53252 | 1712 | then have "a1 = a" | 
| 53185 | 1713 | using ksimplex_fix_plane[OF assms(2-)] by auto | 
| 53252 | 1714 | then show False | 
| 53185 | 1715 | using as and assms(3,5) and assms(7)[rule_format,of j] | 
| 1716 | unfolding assms(4)[rule_format,OF *] | |
| 1717 | using s(4)[OF assms(6), of j] | |
| 1718 | by auto | |
| 1719 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1720 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1721 | lemma ksimplex_replace_0: | 
| 53688 | 1722 | assumes "ksimplex p n s" "a \<in> s" | 
| 1723 | and "n \<noteq> 0" | |
| 1724 |     and "j \<in> {1..n}"
 | |
| 1725 |     and "\<forall>x\<in>s - {a}. x j = 0"
 | |
| 53185 | 1726 |   shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 1"
 | 
| 1727 | proof - | |
| 53688 | 1728 |   have *: "\<And>s' a a'. s' - {a'} = s - {a} \<Longrightarrow> a' = a \<Longrightarrow> a' \<in> s' \<Longrightarrow> a \<in> s \<Longrightarrow> s' = s"
 | 
| 53185 | 1729 | by auto | 
| 1730 |   have **: "\<And>s' a'. ksimplex p n s' \<Longrightarrow> a' \<in> s' \<Longrightarrow> s' - {a'} = s - {a} \<Longrightarrow> s' = s"
 | |
| 1731 | proof - | |
| 1732 | case goal1 | |
| 1733 | guess a0 a1 by (rule ksimplex_extrema_strong[OF assms(1,3)]) note exta = this[rule_format] | |
| 53688 | 1734 | have a: "a = a1" | 
| 53185 | 1735 | apply (rule ksimplex_fix_plane_0[OF assms(2,4-5)]) | 
| 53252 | 1736 | using exta(1-2,5) | 
| 1737 | apply auto | |
| 53185 | 1738 | done | 
| 1739 | moreover | |
| 1740 | guess b0 b1 by (rule ksimplex_extrema_strong[OF goal1(1) assms(3)]) | |
| 1741 | note extb = this[rule_format] | |
| 1742 | have a': "a' = b1" | |
| 1743 | apply (rule ksimplex_fix_plane_0[OF goal1(2) assms(4), of b0]) | |
| 1744 | unfolding goal1(3) | |
| 53252 | 1745 | using assms extb goal1 | 
| 1746 | apply auto | |
| 1747 | done | |
| 53185 | 1748 | moreover | 
| 1749 | have "b0 = a0" | |
| 1750 | unfolding kle_antisym[symmetric, of b0 a0 n] | |
| 1751 | using exta extb and goal1(3) | |
| 1752 | unfolding a a' by blast | |
| 53674 | 1753 | then have "b1 = a1" | 
| 53185 | 1754 | apply - | 
| 1755 | apply (rule ext) | |
| 1756 | unfolding exta(5) extb(5) | |
| 1757 | apply auto | |
| 1758 | done | |
| 1759 | ultimately | |
| 1760 | show "s' = s" | |
| 1761 | apply - | |
| 1762 | apply (rule *[of _ a1 b1]) | |
| 53252 | 1763 | using exta(1-2) extb(1-2) goal1 | 
| 1764 | apply auto | |
| 53185 | 1765 | done | 
| 1766 | qed | |
| 1767 | show ?thesis | |
| 1768 | unfolding card_1_exists | |
| 1769 | apply - | |
| 1770 | apply(rule ex1I[of _ s]) | |
| 1771 | unfolding mem_Collect_eq | |
| 1772 | defer | |
| 1773 | apply (erule conjE bexE)+ | |
| 1774 | apply (rule_tac a'=b in **) | |
| 53252 | 1775 | using assms(1,2) | 
| 1776 | apply auto | |
| 53185 | 1777 | done | 
| 1778 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1779 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1780 | lemma ksimplex_replace_1: | 
| 53688 | 1781 | assumes "ksimplex p n s" | 
| 1782 | and "a \<in> s" | |
| 1783 | and "n \<noteq> 0" | |
| 1784 |     and "j \<in> {1..n}"
 | |
| 1785 |     and "\<forall>x\<in>s - {a}. x j = p"
 | |
| 53186 | 1786 |   shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 1"
 | 
| 1787 | proof - | |
| 1788 |   have lem: "\<And>a a' s'. s' - {a'} = s - {a} \<Longrightarrow> a' = a \<Longrightarrow> a' \<in> s' \<Longrightarrow> a \<in> s \<Longrightarrow> s' = s"
 | |
| 1789 | by auto | |
| 1790 |   have lem: "\<And>s' a'. ksimplex p n s' \<Longrightarrow> a'\<in>s' \<Longrightarrow> s' - {a'} = s - {a} \<Longrightarrow> s' = s"
 | |
| 1791 | proof - | |
| 1792 | case goal1 | |
| 1793 | guess a0 a1 by (rule ksimplex_extrema_strong[OF assms(1,3)]) note exta = this [rule_format] | |
| 1794 | have a: "a = a0" | |
| 1795 | apply (rule ksimplex_fix_plane_p[OF assms(1-2,4-5) exta(1,2)]) | |
| 1796 | unfolding exta | |
| 1797 | apply auto | |
| 1798 | done | |
| 1799 | moreover | |
| 1800 | guess b0 b1 by (rule ksimplex_extrema_strong[OF goal1(1) assms(3)]) | |
| 1801 | note extb = this [rule_format] | |
| 1802 | have a': "a' = b0" | |
| 1803 | apply (rule ksimplex_fix_plane_p[OF goal1(1-2) assms(4), of _ b1]) | |
| 1804 | unfolding goal1 extb | |
| 1805 | using extb(1,2) assms(5) | |
| 1806 | apply auto | |
| 1807 | done | |
| 1808 | moreover | |
| 1809 | have *: "b1 = a1" | |
| 1810 | unfolding kle_antisym[symmetric, of b1 a1 n] | |
| 1811 | using exta extb | |
| 1812 | using goal1(3) | |
| 1813 | unfolding a a' | |
| 1814 | by blast | |
| 1815 | moreover | |
| 1816 | have "a0 = b0" | |
| 53688 | 1817 | proof (rule ext) | 
| 1818 | fix x | |
| 53186 | 1819 | show "a0 x = b0 x" | 
| 1820 | using *[THEN cong, of x x] | |
| 1821 | unfolding exta extb | |
| 53688 | 1822 |         by (cases "x \<in> {1..n}") auto
 | 
| 53186 | 1823 | qed | 
| 1824 | ultimately | |
| 1825 | show "s' = s" | |
| 1826 | apply - | |
| 1827 | apply (rule lem[OF goal1(3) _ goal1(2) assms(2)]) | |
| 1828 | apply auto | |
| 1829 | done | |
| 1830 | qed | |
| 1831 | show ?thesis | |
| 1832 | unfolding card_1_exists | |
| 1833 | apply (rule ex1I[of _ s]) | |
| 1834 | unfolding mem_Collect_eq | |
| 53688 | 1835 | apply rule | 
| 1836 | apply (rule assms(1)) | |
| 53186 | 1837 | apply (rule_tac x = a in bexI) | 
| 1838 | prefer 3 | |
| 1839 | apply (erule conjE bexE)+ | |
| 1840 | apply (rule_tac a'=b in lem) | |
| 1841 | using assms(1-2) | |
| 1842 | apply auto | |
| 1843 | done | |
| 1844 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1845 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1846 | lemma ksimplex_replace_2: | 
| 53688 | 1847 | assumes "ksimplex p n s" | 
| 1848 | and "a \<in> s" | |
| 1849 | and "n \<noteq> 0" | |
| 1850 |     and "\<not> (\<exists>j\<in>{1..n}. \<forall>x\<in>s - {a}. x j = 0)"
 | |
| 1851 |     and "\<not> (\<exists>j\<in>{1..n}. \<forall>x\<in>s - {a}. x j = p)"
 | |
| 53186 | 1852 |   shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 2"
 | 
| 53688 | 1853 | (is "card ?A = 2") | 
| 53186 | 1854 | proof - | 
| 1855 |   have lem1: "\<And>a a' s s'. s' - {a'} = s - {a} \<Longrightarrow> a' = a \<Longrightarrow> a' \<in> s' \<Longrightarrow> a \<in> s \<Longrightarrow> s' = s"
 | |
| 1856 | by auto | |
| 53688 | 1857 |   have lem2: "\<And>a b. a \<in> s \<Longrightarrow> b \<noteq> a \<Longrightarrow> s \<noteq> insert b (s - {a})"
 | 
| 53186 | 1858 | proof | 
| 1859 | case goal1 | |
| 53688 | 1860 |     then have "a \<in> insert b (s - {a})"
 | 
| 1861 | by auto | |
| 1862 |     then have "a \<in> s - {a}"
 | |
| 1863 | unfolding insert_iff | |
| 1864 | using goal1 | |
| 1865 | by auto | |
| 1866 | then show False | |
| 1867 | by auto | |
| 53186 | 1868 | qed | 
| 1869 | guess a0 a1 by (rule ksimplex_extrema_strong[OF assms(1,3)]) note a0a1 = this | |
| 1870 |   {
 | |
| 1871 | assume "a = a0" | |
| 53688 | 1872 | have *: "\<And>P Q. P \<or> Q \<Longrightarrow> \<not> P \<Longrightarrow> Q" | 
| 1873 | by auto | |
| 53186 | 1874 | have "\<exists>x\<in>s. \<not> kle n x a0" | 
| 1875 | apply (rule_tac x=a1 in bexI) | |
| 1876 | proof | |
| 1877 | assume as: "kle n a1 a0" | |
| 1878 | show False | |
| 1879 | using kle_imp_pointwise[OF as,THEN spec[where x=1]] | |
| 1880 | unfolding a0a1(5)[THEN spec[where x=1]] | |
| 53688 | 1881 | using assms(3) | 
| 1882 | by auto | |
| 53186 | 1883 | qed (insert a0a1, auto) | 
| 53674 | 1884 |     then have "\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. y j = (if j = k then a0 j + 1 else a0 j)"
 | 
| 53186 | 1885 | apply (rule_tac *[OF ksimplex_successor[OF assms(1-2),unfolded `a=a0`]]) | 
| 1886 | apply auto | |
| 1887 | done | |
| 1888 | then guess a2 .. | |
| 1889 | from this(2) guess k .. note k = this note a2 =`a2 \<in> s` | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1890 | def a3 \<equiv> "\<lambda>j. if j = k then a1 j + 1 else a1 j" | 
| 53186 | 1891 | have "a3 \<notin> s" | 
| 1892 | proof | |
| 1893 | assume "a3\<in>s" | |
| 53674 | 1894 | then have "kle n a3 a1" | 
| 53688 | 1895 | using a0a1(4) by auto | 
| 53674 | 1896 | then show False | 
| 53688 | 1897 | apply (drule_tac kle_imp_pointwise) | 
| 1898 | unfolding a3_def | |
| 53186 | 1899 | apply (erule_tac x = k in allE) | 
| 1900 | apply auto | |
| 1901 | done | |
| 1902 | qed | |
| 53688 | 1903 | then have "a3 \<noteq> a0" and "a3 \<noteq> a1" | 
| 1904 | using a0a1 by auto | |
| 1905 | have "a2 \<noteq> a0" | |
| 1906 | using k(2)[THEN spec[where x=k]] by auto | |
| 1907 |     have lem3: "\<And>x. x \<in> (s - {a0}) \<Longrightarrow> kle n a2 x"
 | |
| 53186 | 1908 | proof (rule ccontr) | 
| 1909 | case goal1 | |
| 53688 | 1910 | then have as: "x \<in> s" "x \<noteq> a0" | 
| 1911 | by auto | |
| 53186 | 1912 | have "kle n a2 x \<or> kle n x a2" | 
| 53688 | 1913 | using ksimplexD(6)[OF assms(1)] and as `a2 \<in> s` | 
| 1914 | by auto | |
| 53186 | 1915 | moreover | 
| 53688 | 1916 | have "kle n a0 x" | 
| 1917 | using a0a1(4) as by auto | |
| 53186 | 1918 | ultimately have "x = a0 \<or> x = a2" | 
| 1919 | apply - | |
| 1920 | apply (rule kle_adjacent[OF k(2)]) | |
| 1921 | using goal1(2) | |
| 1922 | apply auto | |
| 1923 | done | |
| 53688 | 1924 | then have "x = a2" | 
| 1925 | using as by auto | |
| 1926 | then show False | |
| 1927 | using goal1(2) | |
| 1928 | using kle_refl | |
| 1929 | by auto | |
| 53186 | 1930 | qed | 
| 1931 |     let ?s = "insert a3 (s - {a0})"
 | |
| 1932 | have "ksimplex p n ?s" | |
| 1933 | apply (rule ksimplexI) | |
| 53688 | 1934 | apply (rule_tac[2-] ballI) | 
| 1935 | apply (rule_tac[4] ballI) | |
| 1936 | proof - | |
| 53186 | 1937 | show "card ?s = n + 1" | 
| 1938 | using ksimplexD(2-3)[OF assms(1)] | |
| 53688 | 1939 | using `a3 \<noteq> a0` `a3 \<notin> s` `a0 \<in> s` | 
| 1940 | by (auto simp add: card_insert_if) | |
| 53186 | 1941 | fix x | 
| 1942 |       assume x: "x \<in> insert a3 (s - {a0})"
 | |
| 1943 | show "\<forall>j. x j \<le> p" | |
| 53688 | 1944 | proof | 
| 53186 | 1945 | fix j | 
| 53688 | 1946 | show "x j \<le> p" | 
| 1947 | proof (cases "x = a3") | |
| 53186 | 1948 | case False | 
| 53688 | 1949 | then show ?thesis | 
| 1950 | using x ksimplexD(4)[OF assms(1)] by auto | |
| 53186 | 1951 | next | 
| 1952 | case True | |
| 53688 | 1953 | show ?thesis unfolding True | 
| 1954 | proof (cases "j = k") | |
| 1955 | case False | |
| 1956 | then show "a3 j \<le> p" | |
| 1957 | unfolding True a3_def | |
| 1958 | using `a1 \<in> s` ksimplexD(4)[OF assms(1)] | |
| 1959 | by auto | |
| 1960 | next | |
| 1961 | guess a4 using assms(5)[unfolded bex_simps ball_simps,rule_format,OF k(1)] .. | |
| 1962 | note a4 = this | |
| 1963 | have "a2 k \<le> a4 k" | |
| 1964 | using lem3[OF a4(1)[unfolded `a = a0`],THEN kle_imp_pointwise] | |
| 1965 | by auto | |
| 1966 | also have "\<dots> < p" | |
| 1967 | using ksimplexD(4)[OF assms(1),rule_format,of a4 k] | |
| 1968 | using a4 by auto | |
| 1969 | finally have *: "a0 k + 1 < p" | |
| 1970 | unfolding k(2)[rule_format] | |
| 1971 | by auto | |
| 1972 | case True | |
| 1973 | then show "a3 j \<le>p" | |
| 1974 | unfolding a3_def | |
| 1975 | unfolding a0a1(5)[rule_format] | |
| 1976 | using k(1) k(2)assms(5) | |
| 1977 | using * | |
| 1978 | by simp | |
| 1979 | qed | |
| 53186 | 1980 | qed | 
| 1981 | qed | |
| 1982 |       show "\<forall>j. j \<notin> {1..n} \<longrightarrow> x j = p"
 | |
| 53688 | 1983 | proof (rule, rule) | 
| 53186 | 1984 | fix j :: nat | 
| 1985 |         assume j: "j \<notin> {1..n}"
 | |
| 53688 | 1986 | show "x j = p" | 
| 1987 | proof (cases "x = a3") | |
| 53186 | 1988 | case False | 
| 53688 | 1989 | then show ?thesis | 
| 1990 | using j x ksimplexD(5)[OF assms(1)] | |
| 1991 | by auto | |
| 1992 | next | |
| 1993 | case True | |
| 1994 | show ?thesis | |
| 1995 | unfolding True a3_def | |
| 1996 | using j k(1) | |
| 1997 | using ksimplexD(5)[OF assms(1),rule_format,OF `a1\<in>s` j] | |
| 1998 | by auto | |
| 1999 | qed | |
| 53186 | 2000 | qed | 
| 2001 | fix y | |
| 2002 |       assume y: "y \<in> insert a3 (s - {a0})"
 | |
| 53688 | 2003 | have lem4: "\<And>x. x\<in>s \<Longrightarrow> x \<noteq> a0 \<Longrightarrow> kle n x a3" | 
| 53186 | 2004 | proof - | 
| 2005 | case goal1 | |
| 2006 | guess kk using a0a1(4)[rule_format, OF `x\<in>s`,THEN conjunct2,unfolded kle_def] | |
| 2007 | by (elim exE conjE) | |
| 2008 | note kk = this | |
| 2009 | have "k \<notin> kk" | |
| 2010 | proof | |
| 2011 | assume "k \<in> kk" | |
| 53688 | 2012 | then have "a1 k = x k + 1" | 
| 2013 | using kk by auto | |
| 2014 | then have "a0 k = x k" | |
| 2015 | unfolding a0a1(5)[rule_format] using k(1) by auto | |
| 2016 | then have "a2 k = x k + 1" | |
| 2017 | unfolding k(2)[rule_format] by auto | |
| 53186 | 2018 | moreover | 
| 53688 | 2019 | have "a2 k \<le> x k" | 
| 2020 | using lem3[of x,THEN kle_imp_pointwise] goal1 by auto | |
| 2021 | ultimately show False | |
| 2022 | by auto | |
| 53186 | 2023 | qed | 
| 53674 | 2024 | then show ?case | 
| 53186 | 2025 | unfolding kle_def | 
| 2026 | apply (rule_tac x="insert k kk" in exI) | |
| 2027 | using kk(1) | |
| 2028 | unfolding a3_def kle_def kk(2)[rule_format] | |
| 2029 | using k(1) | |
| 2030 | apply auto | |
| 2031 | done | |
| 2032 | qed | |
| 2033 | show "kle n x y \<or> kle n y x" | |
| 2034 | proof (cases "y = a3") | |
| 2035 | case True | |
| 2036 | show ?thesis | |
| 2037 | unfolding True | |
| 2038 | apply (cases "x = a3") | |
| 2039 | defer | |
| 2040 | apply (rule disjI1, rule lem4) | |
| 2041 | using x | |
| 2042 | apply auto | |
| 2043 | done | |
| 2044 | next | |
| 2045 | case False | |
| 2046 | show ?thesis | |
| 2047 | proof (cases "x = a3") | |
| 2048 | case True | |
| 2049 | show ?thesis | |
| 2050 | unfolding True | |
| 53688 | 2051 | apply (rule disjI2) | 
| 2052 | apply (rule lem4) | |
| 53186 | 2053 | using y False | 
| 2054 | apply auto | |
| 2055 | done | |
| 2056 | next | |
| 2057 | case False | |
| 53674 | 2058 | then show ?thesis | 
| 53186 | 2059 | apply (rule_tac ksimplexD(6)[OF assms(1),rule_format]) | 
| 2060 | using x y `y \<noteq> a3` | |
| 2061 | apply auto | |
| 2062 | done | |
| 2063 | qed | |
| 2064 | qed | |
| 2065 | qed | |
| 53674 | 2066 |     then have "insert a3 (s - {a0}) \<in> ?A"
 | 
| 53186 | 2067 | unfolding mem_Collect_eq | 
| 2068 | apply - | |
| 53688 | 2069 | apply rule | 
| 2070 | apply assumption | |
| 53186 | 2071 | apply (rule_tac x = "a3" in bexI) | 
| 2072 | unfolding `a = a0` | |
| 2073 | using `a3 \<notin> s` | |
| 2074 | apply auto | |
| 2075 | done | |
| 2076 | moreover | |
| 53688 | 2077 | have "s \<in> ?A" | 
| 2078 | using assms(1,2) by auto | |
| 2079 |     ultimately have "?A \<supseteq> {s, insert a3 (s - {a0})}"
 | |
| 2080 | by auto | |
| 53186 | 2081 | moreover | 
| 2082 |     have "?A \<subseteq> {s, insert a3 (s - {a0})}"
 | |
| 2083 | apply rule | |
| 2084 | unfolding mem_Collect_eq | |
| 2085 | proof (erule conjE) | |
| 2086 | fix s' | |
| 2087 |       assume as: "ksimplex p n s'" and "\<exists>b\<in>s'. s' - {b} = s - {a}"
 | |
| 2088 | from this(2) guess a' .. note a' = this | |
| 2089 | guess a_min a_max by (rule ksimplex_extrema_strong[OF as assms(3)]) note min_max = this | |
| 2090 |       have *: "\<forall>x\<in>s' - {a'}. x k = a2 k"
 | |
| 2091 | unfolding a' | |
| 2092 | proof | |
| 2093 | fix x | |
| 2094 |         assume x: "x \<in> s - {a}"
 | |
| 53674 | 2095 | then have "kle n a2 x" | 
| 53186 | 2096 | apply - | 
| 2097 | apply (rule lem3) | |
| 2098 | using `a = a0` | |
| 2099 | apply auto | |
| 2100 | done | |
| 53674 | 2101 | then have "a2 k \<le> x k" | 
| 53186 | 2102 | apply (drule_tac kle_imp_pointwise) | 
| 2103 | apply auto | |
| 2104 | done | |
| 2105 | moreover | |
| 2106 | have "x k \<le> a2 k" | |
| 2107 | unfolding k(2)[rule_format] | |
| 2108 | using a0a1(4)[rule_format,of x, THEN conjunct1] | |
| 53688 | 2109 | unfolding kle_def using x | 
| 2110 | by auto | |
| 2111 | ultimately show "x k = a2 k" | |
| 2112 | by auto | |
| 53186 | 2113 | qed | 
| 2114 | have **: "a' = a_min \<or> a' = a_max" | |
| 2115 | apply (rule ksimplex_fix_plane[OF a'(1) k(1) *]) | |
| 2116 | using min_max | |
| 2117 | apply auto | |
| 2118 | done | |
| 2119 |       show "s' \<in> {s, insert a3 (s - {a0})}"
 | |
| 2120 | proof (cases "a' = a_min") | |
| 2121 | case True | |
| 2122 | have "a_max = a1" | |
| 2123 | unfolding kle_antisym[symmetric,of a_max a1 n] | |
| 2124 | apply rule | |
| 2125 | apply (rule a0a1(4)[rule_format,THEN conjunct2]) | |
| 2126 | defer | |
| 2127 | proof (rule min_max(4)[rule_format,THEN conjunct2]) | |
| 53688 | 2128 | show "a1 \<in> s'" | 
| 2129 | using a' | |
| 2130 | unfolding `a = a0` | |
| 2131 | using a0a1 | |
| 2132 | by auto | |
| 53186 | 2133 | show "a_max \<in> s" | 
| 2134 | proof (rule ccontr) | |
| 53688 | 2135 | assume "\<not> ?thesis" | 
| 2136 | then have "a_max = a'" | |
| 2137 | using a' min_max by auto | |
| 2138 | then show False | |
| 2139 | unfolding True using min_max by auto | |
| 53186 | 2140 | qed | 
| 2141 | qed | |
| 53688 | 2142 | then have "\<forall>i. a_max i = a1 i" | 
| 2143 | by auto | |
| 2144 | then have "a' = a" | |
| 2145 | unfolding True `a = a0` | |
| 53186 | 2146 | apply - | 
| 53688 | 2147 | apply (subst fun_eq_iff) | 
| 2148 | apply rule | |
| 53186 | 2149 | apply (erule_tac x=x in allE) | 
| 2150 | unfolding a0a1(5)[rule_format] min_max(5)[rule_format] | |
| 2151 | proof - | |
| 2152 | case goal1 | |
| 53688 | 2153 | then show ?case | 
| 2154 |             by (cases "x \<in> {1..n}") auto
 | |
| 53186 | 2155 | qed | 
| 53674 | 2156 | then have "s' = s" | 
| 53186 | 2157 | apply - | 
| 2158 | apply (rule lem1[OF a'(2)]) | |
| 53688 | 2159 | using `a \<in> s` `a' \<in> s'` | 
| 53186 | 2160 | apply auto | 
| 2161 | done | |
| 53688 | 2162 | then show ?thesis | 
| 2163 | by auto | |
| 53186 | 2164 | next | 
| 2165 | case False | |
| 53688 | 2166 | then have as: "a' = a_max" | 
| 2167 | using ** by auto | |
| 2168 | have "a_min = a2" | |
| 2169 | unfolding kle_antisym[symmetric, of _ _ n] | |
| 53186 | 2170 | apply rule | 
| 2171 | apply (rule min_max(4)[rule_format,THEN conjunct1]) | |
| 2172 | defer | |
| 2173 | proof (rule lem3) | |
| 2174 |           show "a_min \<in> s - {a0}"
 | |
| 2175 | unfolding a'(2)[symmetric,unfolded `a = a0`] | |
| 53688 | 2176 | unfolding as | 
| 2177 | using min_max(1-3) | |
| 2178 | by auto | |
| 53186 | 2179 | have "a2 \<noteq> a" | 
| 53688 | 2180 | unfolding `a = a0` | 
| 2181 | using k(2)[rule_format,of k] | |
| 2182 | by auto | |
| 53674 | 2183 |           then have "a2 \<in> s - {a}"
 | 
| 53186 | 2184 | using a2 by auto | 
| 53688 | 2185 | then show "a2 \<in> s'" | 
| 2186 | unfolding a'(2)[symmetric] by auto | |
| 53186 | 2187 | qed | 
| 53688 | 2188 | then have "\<forall>i. a_min i = a2 i" | 
| 2189 | by auto | |
| 53674 | 2190 | then have "a' = a3" | 
| 53186 | 2191 | unfolding as `a = a0` | 
| 53688 | 2192 | apply (subst fun_eq_iff) | 
| 2193 | apply rule | |
| 53186 | 2194 | apply (erule_tac x=x in allE) | 
| 2195 | unfolding a0a1(5)[rule_format] min_max(5)[rule_format] | |
| 2196 | unfolding a3_def k(2)[rule_format] | |
| 2197 | unfolding a0a1(5)[rule_format] | |
| 2198 | proof - | |
| 2199 | case goal1 | |
| 2200 | show ?case | |
| 2201 | unfolding goal1 | |
| 53688 | 2202 |             apply (cases "x \<in> {1..n}")
 | 
| 53186 | 2203 | defer | 
| 2204 | apply (cases "x = k") | |
| 53688 | 2205 |             using `k \<in> {1..n}`
 | 
| 53186 | 2206 | apply auto | 
| 2207 | done | |
| 2208 | qed | |
| 53674 | 2209 |         then have "s' = insert a3 (s - {a0})"
 | 
| 53186 | 2210 | apply - | 
| 2211 | apply (rule lem1) | |
| 2212 | defer | |
| 2213 | apply assumption | |
| 2214 | apply (rule a'(1)) | |
| 2215 | unfolding a' `a = a0` | |
| 2216 | using `a3 \<notin> s` | |
| 2217 | apply auto | |
| 2218 | done | |
| 53674 | 2219 | then show ?thesis by auto | 
| 53186 | 2220 | qed | 
| 2221 | qed | |
| 53688 | 2222 |     ultimately have *: "?A = {s, insert a3 (s - {a0})}"
 | 
| 2223 | by blast | |
| 2224 |     have "s \<noteq> insert a3 (s - {a0})"
 | |
| 53846 | 2225 | using `a3 \<notin> s` by auto | 
| 53688 | 2226 | then have ?thesis | 
| 2227 | unfolding * by auto | |
| 53186 | 2228 | } | 
| 2229 | moreover | |
| 2230 |   {
 | |
| 2231 | assume "a = a1" | |
| 53688 | 2232 | have *: "\<And>P Q. P \<or> Q \<Longrightarrow> \<not> P \<Longrightarrow> Q" | 
| 2233 | by auto | |
| 53186 | 2234 | have "\<exists>x\<in>s. \<not> kle n a1 x" | 
| 2235 | apply (rule_tac x=a0 in bexI) | |
| 2236 | proof | |
| 2237 | assume as: "kle n a1 a0" | |
| 2238 | show False | |
| 2239 | using kle_imp_pointwise[OF as,THEN spec[where x=1]] | |
| 2240 | unfolding a0a1(5)[THEN spec[where x=1]] | |
| 2241 | using assms(3) | |
| 2242 | by auto | |
| 2243 | qed (insert a0a1, auto) | |
| 53674 | 2244 |     then have "\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. a1 j = (if j = k then y j + 1 else y j)"
 | 
| 53186 | 2245 | apply (rule_tac *[OF ksimplex_predecessor[OF assms(1-2),unfolded `a=a1`]]) | 
| 2246 | apply auto | |
| 2247 | done | |
| 2248 | then guess a2 .. from this(2) guess k .. note k=this note a2 = `a2 \<in> s` | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2249 | def a3 \<equiv> "\<lambda>j. if j = k then a0 j - 1 else a0 j" | 
| 53688 | 2250 | have "a2 \<noteq> a1" | 
| 2251 | using k(2)[THEN spec[where x=k]] by auto | |
| 2252 |     have lem3: "\<And>x. x \<in> s - {a1} \<Longrightarrow> kle n x a2"
 | |
| 53186 | 2253 | proof (rule ccontr) | 
| 2254 | case goal1 | |
| 53688 | 2255 | then have as: "x \<in> s" "x \<noteq> a1" by auto | 
| 53186 | 2256 | have "kle n a2 x \<or> kle n x a2" | 
| 53688 | 2257 | using ksimplexD(6)[OF assms(1)] and as `a2\<in>s` | 
| 2258 | by auto | |
| 53186 | 2259 | moreover | 
| 53688 | 2260 | have "kle n x a1" | 
| 2261 | using a0a1(4) as by auto | |
| 53186 | 2262 | ultimately have "x = a2 \<or> x = a1" | 
| 2263 | apply - | |
| 2264 | apply (rule kle_adjacent[OF k(2)]) | |
| 2265 | using goal1(2) | |
| 2266 | apply auto | |
| 2267 | done | |
| 53688 | 2268 | then have "x = a2" | 
| 2269 | using as by auto | |
| 2270 | then show False | |
| 2271 | using goal1(2) using kle_refl by auto | |
| 53186 | 2272 | qed | 
| 2273 | have "a0 k \<noteq> 0" | |
| 2274 | proof - | |
| 2275 |       guess a4 using assms(4)[unfolded bex_simps ball_simps,rule_format,OF `k\<in>{1..n}`] ..
 | |
| 2276 | note a4 = this | |
| 53688 | 2277 | have "a4 k \<le> a2 k" | 
| 2278 | using lem3[OF a4(1)[unfolded `a=a1`],THEN kle_imp_pointwise] | |
| 2279 | by auto | |
| 2280 | moreover have "a4 k > 0" | |
| 2281 | using a4 by auto | |
| 2282 | ultimately have "a2 k > 0" | |
| 53186 | 2283 | by auto | 
| 53688 | 2284 | then have "a1 k > 1" | 
| 2285 | unfolding k(2)[rule_format] by simp | |
| 2286 | then show ?thesis | |
| 2287 | unfolding a0a1(5)[rule_format] using k(1) by simp | |
| 53186 | 2288 | qed | 
| 53688 | 2289 | then have lem4: "\<forall>j. a0 j = (if j = k then a3 j + 1 else a3 j)" | 
| 53186 | 2290 | unfolding a3_def by simp | 
| 2291 | have "\<not> kle n a0 a3" | |
| 53688 | 2292 | apply (rule notI) | 
| 53186 | 2293 | apply (drule kle_imp_pointwise) | 
| 2294 | unfolding lem4[rule_format] | |
| 2295 | apply (erule_tac x=k in allE) | |
| 2296 | apply auto | |
| 2297 | done | |
| 53688 | 2298 | then have "a3 \<notin> s" | 
| 2299 | using a0a1(4) by auto | |
| 2300 | then have "a3 \<noteq> a1" "a3 \<noteq> a0" | |
| 2301 | using a0a1 by auto | |
| 53186 | 2302 |     let ?s = "insert a3 (s - {a1})"
 | 
| 2303 | have "ksimplex p n ?s" | |
| 2304 | apply (rule ksimplexI) | |
| 2305 | proof (rule_tac[2-] ballI,rule_tac[4] ballI) | |
| 2306 | show "card ?s = n+1" | |
| 2307 | using ksimplexD(2-3)[OF assms(1)] | |
| 53688 | 2308 | using `a3 \<noteq> a0` `a3 \<notin> s` `a1 \<in> s` | 
| 2309 | by (auto simp add:card_insert_if) | |
| 53186 | 2310 | fix x | 
| 2311 |       assume x: "x \<in> insert a3 (s - {a1})"
 | |
| 53688 | 2312 | show "\<forall>j. x j \<le> p" | 
| 2313 | proof | |
| 53186 | 2314 | fix j | 
| 53688 | 2315 | show "x j \<le> p" | 
| 2316 | proof (cases "x = a3") | |
| 53186 | 2317 | case False | 
| 53688 | 2318 | then show ?thesis | 
| 2319 | using x ksimplexD(4)[OF assms(1)] by auto | |
| 53186 | 2320 | next | 
| 53688 | 2321 | case True | 
| 2322 | show ?thesis | |
| 2323 | unfolding True | |
| 2324 | proof (cases "j = k") | |
| 2325 | case False | |
| 2326 | then show "a3 j \<le> p" | |
| 2327 | unfolding True a3_def | |
| 2328 | using `a0 \<in> s` ksimplexD(4)[OF assms(1)] | |
| 2329 | by auto | |
| 2330 | next | |
| 2331 | guess a4 using assms(5)[unfolded bex_simps ball_simps,rule_format,OF k(1)] .. | |
| 2332 | note a4 = this | |
| 2333 | case True have "a3 k \<le> a0 k" | |
| 2334 | unfolding lem4[rule_format] by auto | |
| 2335 | also have "\<dots> \<le> p" | |
| 2336 | using ksimplexD(4)[OF assms(1),rule_format, of a0 k] a0a1 | |
| 2337 | by auto | |
| 2338 | finally show "a3 j \<le> p" | |
| 2339 | unfolding True by auto | |
| 2340 | qed | |
| 53186 | 2341 | qed | 
| 2342 | qed | |
| 2343 |       show "\<forall>j. j \<notin> {1..n} \<longrightarrow> x j = p"
 | |
| 53688 | 2344 | proof (rule, rule) | 
| 53186 | 2345 | fix j :: nat | 
| 2346 |         assume j: "j \<notin> {1..n}"
 | |
| 53688 | 2347 | show "x j = p" | 
| 2348 | proof (cases "x = a3") | |
| 53186 | 2349 | case False | 
| 53688 | 2350 | then show ?thesis | 
| 2351 | using j x ksimplexD(5)[OF assms(1)] by auto | |
| 53186 | 2352 | next | 
| 2353 | case True | |
| 53688 | 2354 | show ?thesis | 
| 2355 | unfolding True a3_def | |
| 2356 | using j k(1) | |
| 2357 | using ksimplexD(5)[OF assms(1),rule_format,OF `a0\<in>s` j] | |
| 2358 | by auto | |
| 2359 | qed | |
| 53186 | 2360 | qed | 
| 2361 | fix y | |
| 53688 | 2362 |       assume y: "y \<in> insert a3 (s - {a1})"
 | 
| 2363 | have lem4: "\<And>x. x \<in> s \<Longrightarrow> x \<noteq> a1 \<Longrightarrow> kle n a3 x" | |
| 53186 | 2364 | proof - | 
| 2365 | case goal1 | |
| 53688 | 2366 |         then have *: "x\<in>s - {a1}"
 | 
| 2367 | by auto | |
| 53186 | 2368 | have "kle n a3 a2" | 
| 2369 | proof - | |
| 2370 | have "kle n a0 a1" | |
| 2371 | using a0a1 by auto then guess kk unfolding kle_def .. | |
| 53674 | 2372 | then show ?thesis | 
| 53186 | 2373 | unfolding kle_def | 
| 2374 | apply (rule_tac x=kk in exI) | |
| 2375 | unfolding lem4[rule_format] k(2)[rule_format] | |
| 2376 | apply rule | |
| 2377 | defer | |
| 2378 | proof rule | |
| 2379 | case goal1 | |
| 53674 | 2380 | then show ?case | 
| 53186 | 2381 | apply - | 
| 2382 | apply (erule conjE) | |
| 2383 | apply (erule_tac[!] x=j in allE) | |
| 2384 | apply (cases "j \<in> kk") | |
| 2385 | apply (case_tac[!] "j=k") | |
| 2386 | apply auto | |
| 2387 | done | |
| 2388 | qed auto | |
| 2389 | qed | |
| 2390 | moreover | |
| 2391 | have "kle n a3 a0" | |
| 2392 | unfolding kle_def lem4[rule_format] | |
| 2393 |           apply (rule_tac x="{k}" in exI)
 | |
| 2394 | using k(1) | |
| 2395 | apply auto | |
| 2396 | done | |
| 2397 | ultimately | |
| 2398 | show ?case | |
| 2399 | apply - | |
| 2400 | apply (rule kle_between_l[of _ a0 _ a2]) | |
| 2401 | using lem3[OF *] | |
| 2402 | using a0a1(4)[rule_format,OF goal1(1)] | |
| 2403 | apply auto | |
| 2404 | done | |
| 2405 | qed | |
| 2406 | show "kle n x y \<or> kle n y x" | |
| 2407 | proof (cases "y = a3") | |
| 2408 | case True | |
| 2409 | show ?thesis | |
| 2410 | unfolding True | |
| 2411 | apply (cases "x = a3") | |
| 2412 | defer | |
| 53688 | 2413 | apply (rule disjI2) | 
| 2414 | apply (rule lem4) | |
| 53186 | 2415 | using x | 
| 2416 | apply auto | |
| 2417 | done | |
| 2418 | next | |
| 2419 | case False | |
| 2420 | show ?thesis | |
| 2421 | proof (cases "x = a3") | |
| 2422 | case True | |
| 2423 | show ?thesis | |
| 2424 | unfolding True | |
| 53688 | 2425 | apply (rule disjI1) | 
| 2426 | apply (rule lem4) | |
| 53186 | 2427 | using y False | 
| 2428 | apply auto | |
| 2429 | done | |
| 2430 | next | |
| 2431 | case False | |
| 53674 | 2432 | then show ?thesis | 
| 53186 | 2433 | apply (rule_tac ksimplexD(6)[OF assms(1),rule_format]) | 
| 53688 | 2434 | using x y `y \<noteq> a3` | 
| 53186 | 2435 | apply auto | 
| 2436 | done | |
| 2437 | qed | |
| 2438 | qed | |
| 2439 | qed | |
| 53674 | 2440 |     then have "insert a3 (s - {a1}) \<in> ?A"
 | 
| 53186 | 2441 | unfolding mem_Collect_eq | 
| 2442 | apply - | |
| 2443 | apply (rule, assumption) | |
| 2444 | apply (rule_tac x = "a3" in bexI) | |
| 2445 | unfolding `a = a1` | |
| 2446 | using `a3 \<notin> s` | |
| 2447 | apply auto | |
| 2448 | done | |
| 2449 | moreover | |
| 53688 | 2450 | have "s \<in> ?A" | 
| 2451 | using assms(1,2) by auto | |
| 2452 |     ultimately have "?A \<supseteq> {s, insert a3 (s - {a1})}"
 | |
| 2453 | by auto | |
| 53186 | 2454 |     moreover have "?A \<subseteq> {s, insert a3 (s - {a1})}"
 | 
| 2455 | apply rule | |
| 2456 | unfolding mem_Collect_eq | |
| 2457 | proof (erule conjE) | |
| 2458 | fix s' | |
| 2459 |       assume as: "ksimplex p n s'" and "\<exists>b\<in>s'. s' - {b} = s - {a}"
 | |
| 2460 | from this(2) guess a' .. note a' = this | |
| 2461 | guess a_min a_max by (rule ksimplex_extrema_strong[OF as assms(3)]) note min_max = this | |
| 2462 |       have *: "\<forall>x\<in>s' - {a'}. x k = a2 k" unfolding a'
 | |
| 2463 | proof | |
| 2464 | fix x | |
| 2465 |         assume x: "x \<in> s - {a}"
 | |
| 53674 | 2466 | then have "kle n x a2" | 
| 53186 | 2467 | apply - | 
| 2468 | apply (rule lem3) | |
| 2469 | using `a = a1` | |
| 2470 | apply auto | |
| 2471 | done | |
| 53674 | 2472 | then have "x k \<le> a2 k" | 
| 53186 | 2473 | apply (drule_tac kle_imp_pointwise) | 
| 2474 | apply auto | |
| 2475 | done | |
| 2476 | moreover | |
| 2477 |         {
 | |
| 2478 | have "a2 k \<le> a0 k" | |
| 2479 | using k(2)[rule_format,of k] | |
| 2480 | unfolding a0a1(5)[rule_format] | |
| 2481 | using k(1) | |
| 2482 | by simp | |
| 2483 | also have "\<dots> \<le> x k" | |
| 2484 | using a0a1(4)[rule_format,of x,THEN conjunct1,THEN kle_imp_pointwise] x | |
| 2485 | by auto | |
| 2486 | finally have "a2 k \<le> x k" . | |
| 2487 | } | |
| 53688 | 2488 | ultimately show "x k = a2 k" | 
| 2489 | by auto | |
| 53186 | 2490 | qed | 
| 2491 | have **: "a' = a_min \<or> a' = a_max" | |
| 2492 | apply (rule ksimplex_fix_plane[OF a'(1) k(1) *]) | |
| 2493 | using min_max | |
| 2494 | apply auto | |
| 2495 | done | |
| 2496 | have "a2 \<noteq> a1" | |
| 2497 | proof | |
| 2498 | assume as: "a2 = a1" | |
| 2499 | show False | |
| 2500 | using k(2) | |
| 2501 | unfolding as | |
| 2502 | apply (erule_tac x = k in allE) | |
| 2503 | apply auto | |
| 2504 | done | |
| 2505 | qed | |
| 53674 | 2506 |       then have a2': "a2 \<in> s' - {a'}"
 | 
| 53186 | 2507 | unfolding a' | 
| 2508 | using a2 | |
| 2509 | unfolding `a = a1` | |
| 2510 | by auto | |
| 2511 |       show "s' \<in> {s, insert a3 (s - {a1})}"
 | |
| 2512 | proof (cases "a' = a_min") | |
| 2513 | case True | |
| 2514 |         have "a_max \<in> s - {a1}"
 | |
| 2515 | using min_max | |
| 2516 | unfolding a'(2)[unfolded `a=a1`,symmetric] True | |
| 2517 | by auto | |
| 53674 | 2518 | then have "a_max = a2" | 
| 53186 | 2519 | unfolding kle_antisym[symmetric,of a_max a2 n] | 
| 2520 | apply - | |
| 2521 | apply rule | |
| 53688 | 2522 | apply (rule lem3) | 
| 2523 | apply assumption | |
| 53186 | 2524 | apply (rule min_max(4)[rule_format,THEN conjunct2]) | 
| 2525 | using a2' | |
| 2526 | apply auto | |
| 2527 | done | |
| 53688 | 2528 | then have a_max: "\<forall>i. a_max i = a2 i" | 
| 2529 | by auto | |
| 2530 |         have *: "\<forall>j. a2 j = (if j \<in> {1..n} then a3 j + 1 else a3 j)"
 | |
| 53186 | 2531 | using k(2) | 
| 2532 | unfolding lem4[rule_format] a0a1(5)[rule_format] | |
| 2533 | apply - | |
| 53688 | 2534 | apply rule | 
| 2535 | apply (erule_tac x=j in allE) | |
| 53186 | 2536 | proof - | 
| 2537 | case goal1 | |
| 53688 | 2538 |           then show ?case by (cases "j \<in> {1..n}", case_tac[!] "j = k") auto
 | 
| 53186 | 2539 | qed | 
| 2540 | have "\<forall>i. a_min i = a3 i" | |
| 2541 | using a_max | |
| 2542 | apply - | |
| 53688 | 2543 | apply rule | 
| 2544 | apply (erule_tac x=i in allE) | |
| 53186 | 2545 | unfolding min_max(5)[rule_format] *[rule_format] | 
| 2546 | proof - | |
| 2547 | case goal1 | |
| 53688 | 2548 | then show ?case | 
| 2549 |             by (cases "i \<in> {1..n}") auto
 | |
| 53186 | 2550 | qed | 
| 53688 | 2551 | then have "a_min = a3" | 
| 2552 | unfolding fun_eq_iff . | |
| 53674 | 2553 |         then have "s' = insert a3 (s - {a1})"
 | 
| 53688 | 2554 | using a' | 
| 2555 | unfolding `a = a1` True | |
| 2556 | by auto | |
| 2557 | then show ?thesis | |
| 2558 | by auto | |
| 53186 | 2559 | next | 
| 53688 | 2560 | case False | 
| 2561 | then have as: "a' = a_max" | |
| 2562 | using ** by auto | |
| 53186 | 2563 | have "a_min = a0" | 
| 2564 | unfolding kle_antisym[symmetric,of _ _ n] | |
| 2565 | apply rule | |
| 2566 | apply (rule min_max(4)[rule_format,THEN conjunct1]) | |
| 2567 | defer | |
| 2568 | apply (rule a0a1(4)[rule_format,THEN conjunct1]) | |
| 2569 | proof - | |
| 2570 |           have "a_min \<in> s - {a1}"
 | |
| 2571 | using min_max(1,3) | |
| 2572 | unfolding a'(2)[symmetric,unfolded `a=a1`] as | |
| 2573 | by auto | |
| 53688 | 2574 | then show "a_min \<in> s" | 
| 2575 | by auto | |
| 2576 |           have "a0 \<in> s - {a1}"
 | |
| 2577 | using a0a1(1-3) by auto | |
| 2578 | then show "a0 \<in> s'" | |
| 2579 | unfolding a'(2)[symmetric,unfolded `a=a1`] | |
| 2580 | by auto | |
| 53186 | 2581 | qed | 
| 53674 | 2582 | then have "\<forall>i. a_max i = a1 i" | 
| 53688 | 2583 | unfolding a0a1(5)[rule_format] min_max(5)[rule_format] | 
| 2584 | by auto | |
| 53674 | 2585 | then have "s' = s" | 
| 53186 | 2586 | apply - | 
| 2587 | apply (rule lem1[OF a'(2)]) | |
| 2588 | using `a \<in> s` `a' \<in> s'` | |
| 2589 | unfolding as `a = a1` | |
| 2590 | unfolding fun_eq_iff | |
| 2591 | apply auto | |
| 2592 | done | |
| 53688 | 2593 | then show ?thesis | 
| 2594 | by auto | |
| 53186 | 2595 | qed | 
| 2596 | qed | |
| 53688 | 2597 |     ultimately have *: "?A = {s, insert a3 (s - {a1})}"
 | 
| 2598 | by blast | |
| 2599 |     have "s \<noteq> insert a3 (s - {a1})"
 | |
| 2600 | using `a3\<notin>s` by auto | |
| 2601 | then have ?thesis | |
| 2602 | unfolding * by auto | |
| 53186 | 2603 | } | 
| 2604 | moreover | |
| 2605 |   {
 | |
| 53688 | 2606 | assume as: "a \<noteq> a0" "a \<noteq> a1" | 
| 2607 | have "\<not> (\<forall>x\<in>s. kle n a x)" | |
| 53186 | 2608 | proof | 
| 2609 | case goal1 | |
| 2610 | have "a = a0" | |
| 2611 | unfolding kle_antisym[symmetric,of _ _ n] | |
| 2612 | apply rule | |
| 2613 | using goal1 a0a1 assms(2) | |
| 2614 | apply auto | |
| 2615 | done | |
| 53688 | 2616 | then show False | 
| 2617 | using as by auto | |
| 53186 | 2618 | qed | 
| 53674 | 2619 |     then have "\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. a j = (if j = k then y j + 1 else y j)"
 | 
| 53688 | 2620 | using ksimplex_predecessor[OF assms(1-2)] | 
| 2621 | by blast | |
| 53186 | 2622 | then guess u .. from this(2) guess k .. note k = this[rule_format] | 
| 2623 | note u = `u \<in> s` | |
| 2624 | have "\<not> (\<forall>x\<in>s. kle n x a)" | |
| 2625 | proof | |
| 2626 | case goal1 | |
| 2627 | have "a = a1" | |
| 2628 | unfolding kle_antisym[symmetric,of _ _ n] | |
| 2629 | apply rule | |
| 2630 | using goal1 a0a1 assms(2) | |
| 2631 | apply auto | |
| 2632 | done | |
| 53688 | 2633 | then show False | 
| 2634 | using as by auto | |
| 53186 | 2635 | qed | 
| 53674 | 2636 |     then have "\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. y j = (if j = k then a j + 1 else a j)"
 | 
| 53186 | 2637 | using ksimplex_successor[OF assms(1-2)] by blast | 
| 2638 | then guess v .. from this(2) guess l .. | |
| 2639 | note l = this[rule_format] | |
| 53688 | 2640 | note v = `v \<in> s` | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2641 | def a' \<equiv> "\<lambda>j. if j = l then u j + 1 else u j" | 
| 53186 | 2642 | have kl: "k \<noteq> l" | 
| 2643 | proof | |
| 2644 | assume "k = l" | |
| 53688 | 2645 | have *: "\<And>P. (if P then (1::nat) else 0) \<noteq> 2" | 
| 2646 | by auto | |
| 2647 | then show False | |
| 2648 | using ksimplexD(6)[OF assms(1),rule_format,OF u v] | |
| 2649 | unfolding kle_def | |
| 53186 | 2650 | unfolding l(2) k(2) `k = l` | 
| 2651 | apply - | |
| 2652 | apply (erule disjE) | |
| 2653 | apply (erule_tac[!] exE conjE)+ | |
| 2654 | apply (erule_tac[!] x = l in allE)+ | |
| 2655 | apply (auto simp add: *) | |
| 2656 | done | |
| 2657 | qed | |
| 53674 | 2658 | then have aa': "a' \<noteq> a" | 
| 53186 | 2659 | apply - | 
| 2660 | apply rule | |
| 2661 | unfolding fun_eq_iff | |
| 2662 | unfolding a'_def k(2) | |
| 2663 | apply (erule_tac x=l in allE) | |
| 2664 | apply auto | |
| 2665 | done | |
| 2666 | have "a' \<notin> s" | |
| 2667 | apply rule | |
| 2668 | apply (drule ksimplexD(6)[OF assms(1),rule_format,OF `a\<in>s`]) | |
| 2669 | proof (cases "kle n a a'") | |
| 2670 | case goal2 | |
| 53688 | 2671 | then have "kle n a' a" | 
| 2672 | by auto | |
| 53674 | 2673 | then show False | 
| 53186 | 2674 | apply (drule_tac kle_imp_pointwise) | 
| 2675 | apply (erule_tac x=l in allE) | |
| 2676 | unfolding a'_def k(2) | |
| 2677 | using kl | |
| 2678 | apply auto | |
| 2679 | done | |
| 2680 | next | |
| 2681 | case True | |
| 53674 | 2682 | then show False | 
| 53186 | 2683 | apply (drule_tac kle_imp_pointwise) | 
| 2684 | apply (erule_tac x=k in allE) | |
| 2685 | unfolding a'_def k(2) | |
| 2686 | using kl | |
| 2687 | apply auto | |
| 2688 | done | |
| 2689 | qed | |
| 2690 | have kle_uv: "kle n u a" "kle n u a'" "kle n a v" "kle n a' v" | |
| 2691 | unfolding kle_def | |
| 2692 | apply - | |
| 2693 |       apply (rule_tac[1] x="{k}" in exI,rule_tac[2] x="{l}" in exI)
 | |
| 2694 |       apply (rule_tac[3] x="{l}" in exI,rule_tac[4] x="{k}" in exI)
 | |
| 2695 | unfolding l(2) k(2) a'_def | |
| 2696 | using l(1) k(1) | |
| 2697 | apply auto | |
| 2698 | done | |
| 2699 | have uxv: "\<And>x. kle n u x \<Longrightarrow> kle n x v \<Longrightarrow> x = u \<or> x = a \<or> x = a' \<or> x = v" | |
| 2700 | proof - | |
| 2701 | case goal1 | |
| 53674 | 2702 | then show ?case | 
| 53186 | 2703 | proof (cases "x k = u k", case_tac[!] "x l = u l") | 
| 2704 | assume as: "x l = u l" "x k = u k" | |
| 53688 | 2705 | have "x = u" | 
| 2706 | unfolding fun_eq_iff | |
| 2707 | using goal1(2)[THEN kle_imp_pointwise,unfolded l(2)] | |
| 2708 | unfolding k(2) | |
| 53186 | 2709 | apply - | 
| 2710 | using goal1(1)[THEN kle_imp_pointwise] | |
| 2711 | apply - | |
| 2712 | apply rule | |
| 2713 | apply (erule_tac x=xa in allE)+ | |
| 2714 | proof - | |
| 2715 | case goal1 | |
| 53674 | 2716 | then show ?case | 
| 53186 | 2717 | apply (cases "x = l") | 
| 2718 | apply (case_tac[!] "x = k") | |
| 53688 | 2719 | using as | 
| 2720 | by auto | |
| 53186 | 2721 | qed | 
| 53688 | 2722 | then show ?case | 
| 2723 | by auto | |
| 53186 | 2724 | next | 
| 2725 | assume as: "x l \<noteq> u l" "x k = u k" | |
| 2726 | have "x = a'" | |
| 2727 | unfolding fun_eq_iff | |
| 2728 | unfolding a'_def | |
| 2729 | using goal1(2)[THEN kle_imp_pointwise] | |
| 2730 | unfolding l(2) k(2) | |
| 2731 | using goal1(1)[THEN kle_imp_pointwise] | |
| 2732 | apply - | |
| 2733 | apply rule | |
| 2734 | apply (erule_tac x = xa in allE)+ | |
| 2735 | proof - | |
| 2736 | case goal1 | |
| 53674 | 2737 | then show ?case | 
| 53186 | 2738 | apply (cases "x = l") | 
| 2739 | apply (case_tac[!] "x = k") | |
| 2740 | using as | |
| 2741 | apply auto | |
| 2742 | done | |
| 2743 | qed | |
| 53674 | 2744 | then show ?case by auto | 
| 53186 | 2745 | next | 
| 2746 | assume as: "x l = u l" "x k \<noteq> u k" | |
| 2747 | have "x = a" | |
| 2748 | unfolding fun_eq_iff | |
| 2749 | using goal1(2)[THEN kle_imp_pointwise] | |
| 2750 | unfolding l(2) k(2) | |
| 2751 | using goal1(1)[THEN kle_imp_pointwise] | |
| 2752 | apply - | |
| 2753 | apply rule | |
| 2754 | apply (erule_tac x=xa in allE)+ | |
| 2755 | proof - | |
| 2756 | case goal1 | |
| 53674 | 2757 | then show ?case | 
| 53186 | 2758 | apply (cases "x = l") | 
| 2759 | apply (case_tac[!] "x = k") | |
| 2760 | using as | |
| 2761 | apply auto | |
| 2762 | done | |
| 2763 | qed | |
| 53688 | 2764 | then show ?case | 
| 2765 | by auto | |
| 53186 | 2766 | next | 
| 2767 | assume as: "x l \<noteq> u l" "x k \<noteq> u k" | |
| 2768 | have "x = v" | |
| 2769 | unfolding fun_eq_iff | |
| 2770 | using goal1(2)[THEN kle_imp_pointwise] | |
| 2771 | unfolding l(2) k(2) | |
| 2772 | using goal1(1)[THEN kle_imp_pointwise] | |
| 2773 | apply - | |
| 2774 | apply rule | |
| 2775 | apply (erule_tac x=xa in allE)+ | |
| 2776 | proof - | |
| 2777 | case goal1 | |
| 53674 | 2778 | then show ?case | 
| 53186 | 2779 | apply (cases "x = l") | 
| 2780 | apply (case_tac[!] "x = k") | |
| 2781 | using as `k \<noteq> l` | |
| 2782 | apply auto | |
| 2783 | done | |
| 2784 | qed | |
| 53674 | 2785 | then show ?case by auto | 
| 53186 | 2786 | qed | 
| 2787 | qed | |
| 2788 | have uv: "kle n u v" | |
| 2789 | apply (rule kle_trans[OF kle_uv(1,3)]) | |
| 2790 | using ksimplexD(6)[OF assms(1)] | |
| 2791 | using u v | |
| 2792 | apply auto | |
| 2793 | done | |
| 2794 | have lem3: "\<And>x. x \<in> s \<Longrightarrow> kle n v x \<Longrightarrow> kle n a' x" | |
| 2795 | apply (rule kle_between_r[of _ u _ v]) | |
| 2796 | prefer 3 | |
| 2797 | apply (rule kle_trans[OF uv]) | |
| 2798 | defer | |
| 2799 | apply (rule ksimplexD(6)[OF assms(1), rule_format]) | |
| 53688 | 2800 | using kle_uv `u \<in> s` | 
| 53186 | 2801 | apply auto | 
| 2802 | done | |
| 53688 | 2803 | have lem4: "\<And>x. x \<in> s \<Longrightarrow> kle n x u \<Longrightarrow> kle n x a'" | 
| 53186 | 2804 | apply (rule kle_between_l[of _ u _ v]) | 
| 2805 | prefer 4 | |
| 2806 | apply (rule kle_trans[OF _ uv]) | |
| 2807 | defer | |
| 2808 | apply (rule ksimplexD(6)[OF assms(1), rule_format]) | |
| 53688 | 2809 | using kle_uv `v \<in> s` | 
| 53186 | 2810 | apply auto | 
| 2811 | done | |
| 2812 | have lem5: "\<And>x. x \<in> s \<Longrightarrow> x \<noteq> a \<Longrightarrow> kle n x a' \<or> kle n a' x" | |
| 2813 | proof - | |
| 2814 | case goal1 | |
| 53674 | 2815 | then show ?case | 
| 53186 | 2816 | proof (cases "kle n v x \<or> kle n x u") | 
| 2817 | case True | |
| 53688 | 2818 | then show ?thesis | 
| 2819 | using goal1 by (auto intro: lem3 lem4) | |
| 53186 | 2820 | next | 
| 2821 | case False | |
| 53674 | 2822 | then have *: "kle n u x" "kle n x v" | 
| 53186 | 2823 | using ksimplexD(6)[OF assms(1)] | 
| 53688 | 2824 | using goal1 `u \<in> s` `v \<in> s` | 
| 53186 | 2825 | by auto | 
| 2826 | show ?thesis | |
| 2827 | using uxv[OF *] | |
| 2828 | using kle_uv | |
| 2829 | using goal1 | |
| 2830 | by auto | |
| 2831 | qed | |
| 2832 | qed | |
| 2833 |     have "ksimplex p n (insert a' (s - {a}))"
 | |
| 2834 | apply (rule ksimplexI) | |
| 53688 | 2835 | apply (rule_tac[2-] ballI) | 
| 2836 | apply (rule_tac[4] ballI) | |
| 2837 | proof - | |
| 53186 | 2838 |       show "card (insert a' (s - {a})) = n + 1"
 | 
| 2839 | using ksimplexD(2-3)[OF assms(1)] | |
| 2840 | using `a' \<noteq> a` `a' \<notin> s` `a \<in> s` | |
| 2841 | by (auto simp add:card_insert_if) | |
| 2842 | fix x | |
| 2843 |       assume x: "x \<in> insert a' (s - {a})"
 | |
| 2844 | show "\<forall>j. x j \<le> p" | |
| 53688 | 2845 | proof | 
| 53186 | 2846 | fix j | 
| 53688 | 2847 | show "x j \<le> p" | 
| 2848 | proof (cases "x = a'") | |
| 53186 | 2849 | case False | 
| 53688 | 2850 | then show ?thesis | 
| 2851 | using x ksimplexD(4)[OF assms(1)] by auto | |
| 53186 | 2852 | next | 
| 2853 | case True | |
| 53688 | 2854 | show ?thesis | 
| 2855 | unfolding True | |
| 2856 | proof (cases "j = l") | |
| 2857 | case False | |
| 2858 | then show "a' j \<le>p" | |
| 2859 | unfolding True a'_def | |
| 2860 | using `u\<in>s` ksimplexD(4)[OF assms(1)] | |
| 2861 | by auto | |
| 2862 | next | |
| 2863 | case True | |
| 2864 | have *: "a l = u l" "v l = a l + 1" | |
| 2865 | using k(2)[of l] l(2)[of l] `k \<noteq> l` | |
| 2866 | by auto | |
| 2867 | have "u l + 1 \<le> p" | |
| 2868 | unfolding *[symmetric] | |
| 2869 | using ksimplexD(4)[OF assms(1)] | |
| 2870 | using `v \<in> s` | |
| 2871 | by auto | |
| 2872 | then show "a' j \<le>p" | |
| 2873 | unfolding a'_def True | |
| 2874 | by auto | |
| 2875 | qed | |
| 53186 | 2876 | qed | 
| 2877 | qed | |
| 2878 |       show "\<forall>j. j \<notin> {1..n} \<longrightarrow> x j = p"
 | |
| 53688 | 2879 | proof (rule, rule) | 
| 53186 | 2880 | fix j :: nat | 
| 2881 |         assume j: "j \<notin> {1..n}"
 | |
| 53688 | 2882 | show "x j = p" | 
| 2883 | proof (cases "x = a'") | |
| 53186 | 2884 | case False | 
| 53688 | 2885 | then show ?thesis | 
| 2886 | using j x ksimplexD(5)[OF assms(1)] by auto | |
| 53186 | 2887 | next | 
| 2888 | case True | |
| 53688 | 2889 | show ?thesis | 
| 53186 | 2890 | unfolding True a'_def | 
| 2891 | using j l(1) | |
| 2892 | using ksimplexD(5)[OF assms(1),rule_format,OF `u\<in>s` j] | |
| 2893 | by auto | |
| 53688 | 2894 | qed | 
| 53186 | 2895 | qed | 
| 2896 | fix y | |
| 53688 | 2897 |       assume y: "y \<in> insert a' (s - {a})"
 | 
| 53186 | 2898 | show "kle n x y \<or> kle n y x" | 
| 2899 | proof (cases "y = a'") | |
| 2900 | case True | |
| 2901 | show ?thesis | |
| 2902 | unfolding True | |
| 2903 | apply (cases "x = a'") | |
| 2904 | defer | |
| 2905 | apply (rule lem5) | |
| 2906 | using x | |
| 2907 | apply auto | |
| 2908 | done | |
| 2909 | next | |
| 2910 | case False | |
| 2911 | show ?thesis | |
| 2912 | proof (cases "x = a'") | |
| 2913 | case True | |
| 2914 | show ?thesis | |
| 2915 | unfolding True | |
| 2916 | using lem5[of y] using y by auto | |
| 2917 | next | |
| 2918 | case False | |
| 53674 | 2919 | then show ?thesis | 
| 53186 | 2920 | apply (rule_tac ksimplexD(6)[OF assms(1),rule_format]) | 
| 53688 | 2921 | using x y `y \<noteq> a'` | 
| 53186 | 2922 | apply auto | 
| 2923 | done | |
| 2924 | qed | |
| 2925 | qed | |
| 2926 | qed | |
| 53674 | 2927 |     then have "insert a' (s - {a}) \<in> ?A"
 | 
| 53186 | 2928 | unfolding mem_Collect_eq | 
| 2929 | apply - | |
| 53688 | 2930 | apply rule | 
| 2931 | apply assumption | |
| 53186 | 2932 | apply (rule_tac x = "a'" in bexI) | 
| 2933 | using aa' `a' \<notin> s` | |
| 2934 | apply auto | |
| 2935 | done | |
| 2936 | moreover | |
| 53688 | 2937 | have "s \<in> ?A" | 
| 2938 | using assms(1,2) by auto | |
| 2939 |     ultimately have  "?A \<supseteq> {s, insert a' (s - {a})}"
 | |
| 2940 | by auto | |
| 53186 | 2941 | moreover | 
| 2942 |     have "?A \<subseteq> {s, insert a' (s - {a})}"
 | |
| 53688 | 2943 | apply rule | 
| 2944 | unfolding mem_Collect_eq | |
| 53186 | 2945 | proof (erule conjE) | 
| 2946 | fix s' | |
| 2947 |       assume as: "ksimplex p n s'" and "\<exists>b\<in>s'. s' - {b} = s - {a}"
 | |
| 2948 | from this(2) guess a'' .. note a'' = this | |
| 53688 | 2949 | have "u \<noteq> v" | 
| 2950 | unfolding fun_eq_iff unfolding l(2) k(2) by auto | |
| 2951 | then have uv': "\<not> kle n v u" | |
| 2952 | using uv using kle_antisym by auto | |
| 2953 | have "u \<noteq> a" "v \<noteq> a" | |
| 2954 | unfolding fun_eq_iff k(2) l(2) by auto | |
| 2955 | then have uvs': "u \<in> s'" "v \<in> s'" | |
| 2956 | using `u \<in> s` `v \<in> s` using a'' by auto | |
| 53186 | 2957 | have lem6: "a \<in> s' \<or> a' \<in> s'" | 
| 2958 | proof (cases "\<forall>x\<in>s'. kle n x u \<or> kle n v x") | |
| 2959 | case False | |
| 2960 | then guess w unfolding ball_simps .. note w = this | |
| 53674 | 2961 | then have "kle n u w" "kle n w v" | 
| 53186 | 2962 | using ksimplexD(6)[OF as] uvs' by auto | 
| 53674 | 2963 | then have "w = a' \<or> w = a" | 
| 53186 | 2964 | using uxv[of w] uvs' w by auto | 
| 53688 | 2965 | then show ?thesis | 
| 2966 | using w by auto | |
| 53186 | 2967 | next | 
| 2968 | case True | |
| 2969 | have "\<not> (\<forall>x\<in>s'. kle n x u)" | |
| 2970 | unfolding ball_simps | |
| 2971 | apply (rule_tac x=v in bexI) | |
| 2972 | using uv `u \<noteq> v` | |
| 2973 | unfolding kle_antisym [of n u v,symmetric] | |
| 53688 | 2974 | using `v \<in> s'` | 
| 53186 | 2975 | apply auto | 
| 2976 | done | |
| 53674 | 2977 |         then have "\<exists>y\<in>s'. \<exists>k\<in>{1..n}. \<forall>j. y j = (if j = k then u j + 1 else u j)"
 | 
| 53688 | 2978 | using ksimplex_successor[OF as `u\<in>s'`] | 
| 2979 | by blast | |
| 53186 | 2980 | then guess w .. note w = this | 
| 2981 | from this(2) guess kk .. note kk = this[rule_format] | |
| 2982 | have "\<not> kle n w u" | |
| 2983 | apply - | |
| 53688 | 2984 | apply rule | 
| 2985 | apply (drule kle_imp_pointwise) | |
| 53186 | 2986 | apply (erule_tac x = kk in allE) | 
| 2987 | unfolding kk | |
| 2988 | apply auto | |
| 2989 | done | |
| 53674 | 2990 | then have *: "kle n v w" | 
| 53688 | 2991 | using True[rule_format,OF w(1)] | 
| 2992 | by auto | |
| 53674 | 2993 | then have False | 
| 53688 | 2994 | proof (cases "kk = l") | 
| 2995 | case False | |
| 53674 | 2996 | then show False using *[THEN kle_imp_pointwise, unfolded l(2) kk k(2)] | 
| 53186 | 2997 | apply (erule_tac x=l in allE) | 
| 2998 | using `k \<noteq> l` | |
| 2999 | apply auto | |
| 3000 | done | |
| 3001 | next | |
| 53688 | 3002 | case True | 
| 53674 | 3003 | then have "kk \<noteq> k" using `k \<noteq> l` by auto | 
| 53688 | 3004 | then show False | 
| 3005 | using *[THEN kle_imp_pointwise, unfolded l(2) kk k(2)] | |
| 53186 | 3006 | apply (erule_tac x=k in allE) | 
| 3007 | using `k \<noteq> l` | |
| 3008 | apply auto | |
| 3009 | done | |
| 3010 | qed | |
| 53688 | 3011 | then show ?thesis | 
| 3012 | by auto | |
| 53186 | 3013 | qed | 
| 53688 | 3014 |       then show "s' \<in> {s, insert a' (s - {a})}"
 | 
| 3015 | proof (cases "a \<in> s'") | |
| 53186 | 3016 | case True | 
| 53674 | 3017 | then have "s' = s" | 
| 53186 | 3018 | apply - | 
| 3019 | apply (rule lem1[OF a''(2)]) | |
| 3020 | using a'' `a \<in> s` | |
| 3021 | apply auto | |
| 3022 | done | |
| 53688 | 3023 | then show ?thesis | 
| 3024 | by auto | |
| 53186 | 3025 | next | 
| 53688 | 3026 | case False | 
| 3027 | then have "a' \<in> s'" | |
| 3028 | using lem6 by auto | |
| 53674 | 3029 |         then have "s' = insert a' (s - {a})"
 | 
| 53186 | 3030 | apply - | 
| 3031 | apply (rule lem1[of _ a'' _ a']) | |
| 53688 | 3032 | unfolding a''(2)[symmetric] | 
| 3033 | using a'' and `a' \<notin> s` | |
| 3034 | by auto | |
| 3035 | then show ?thesis | |
| 3036 | by auto | |
| 53186 | 3037 | qed | 
| 3038 | qed | |
| 53688 | 3039 |     ultimately have *: "?A = {s, insert a' (s - {a})}"
 | 
| 3040 | by blast | |
| 3041 |     have "s \<noteq> insert a' (s - {a})"
 | |
| 3042 | using `a'\<notin>s` by auto | |
| 3043 | then have ?thesis | |
| 3044 | unfolding * by auto | |
| 53186 | 3045 | } | 
| 53688 | 3046 | ultimately show ?thesis | 
| 3047 | by auto | |
| 53186 | 3048 | qed | 
| 3049 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3050 | |
| 53688 | 3051 | text {* Hence another step towards concreteness. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3052 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3053 | lemma kuhn_simplex_lemma: | 
| 53688 | 3054 |   assumes "\<forall>s. ksimplex p (n + 1) s \<longrightarrow> rl ` s \<subseteq> {0..n+1}"
 | 
| 3055 |     and "odd (card {f. \<exists>s a. ksimplex p (n + 1) s \<and> a \<in> s \<and> (f = s - {a}) \<and>
 | |
| 3056 |       rl ` f = {0 .. n} \<and> ((\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = p))})"
 | |
| 3057 |   shows "odd (card {s \<in> {s. ksimplex p (n + 1) s}. rl ` s = {0..n+1}})"
 | |
| 53186 | 3058 | proof - | 
| 3059 | have *: "\<And>x y. x = y \<Longrightarrow> odd (card x) \<Longrightarrow> odd (card y)" | |
| 3060 | by auto | |
| 53688 | 3061 |   have *: "odd (card {f \<in> {f. \<exists>s\<in>{s. ksimplex p (n + 1) s}. (\<exists>a\<in>s. f = s - {a})}.
 | 
| 3062 |     rl ` f = {0..n} \<and> ((\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = p))})"
 | |
| 53186 | 3063 | apply (rule *[OF _ assms(2)]) | 
| 3064 | apply auto | |
| 3065 | done | |
| 3066 | show ?thesis | |
| 3067 | apply (rule kuhn_complete_lemma[OF finite_simplices]) | |
| 3068 | prefer 6 | |
| 3069 | apply (rule *) | |
| 53688 | 3070 | apply rule | 
| 3071 | apply rule | |
| 3072 | apply rule | |
| 53186 | 3073 | apply (subst ksimplex_def) | 
| 3074 | defer | |
| 53688 | 3075 | apply rule | 
| 3076 | apply (rule assms(1)[rule_format]) | |
| 53186 | 3077 | unfolding mem_Collect_eq | 
| 3078 | apply assumption | |
| 3079 | apply default+ | |
| 3080 | unfolding mem_Collect_eq | |
| 3081 | apply (erule disjE bexE)+ | |
| 3082 | defer | |
| 3083 | apply (erule disjE bexE)+ | |
| 3084 | defer | |
| 3085 | apply default+ | |
| 3086 | unfolding mem_Collect_eq | |
| 3087 | apply (erule disjE bexE)+ | |
| 3088 | unfolding mem_Collect_eq | |
| 3089 | proof - | |
| 3090 | fix f s a | |
| 53688 | 3091 |     assume as: "ksimplex p (n + 1) s" "a \<in> s" "f = s - {a}"
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3092 |     let ?S = "{s. ksimplex p (n + 1) s \<and> (\<exists>a\<in>s. f = s - {a})}"
 | 
| 53186 | 3093 |     have S: "?S = {s'. ksimplex p (n + 1) s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})}"
 | 
| 3094 | unfolding as by blast | |
| 3095 |     {
 | |
| 3096 | fix j | |
| 3097 |       assume j: "j \<in> {1..n + 1}" "\<forall>x\<in>f. x j = 0"
 | |
| 53674 | 3098 |       then show "card {s. ksimplex p (n + 1) s \<and> (\<exists>a\<in>s. f = s - {a})} = 1"
 | 
| 53186 | 3099 | unfolding S | 
| 3100 | apply - | |
| 3101 | apply (rule ksimplex_replace_0) | |
| 3102 | apply (rule as)+ | |
| 3103 | unfolding as | |
| 3104 | apply auto | |
| 3105 | done | |
| 3106 | } | |
| 3107 |     {
 | |
| 3108 | fix j | |
| 3109 |       assume j: "j \<in> {1..n + 1}" "\<forall>x\<in>f. x j = p"
 | |
| 53674 | 3110 |       then show "card {s. ksimplex p (n + 1) s \<and> (\<exists>a\<in>s. f = s - {a})} = 1"
 | 
| 53186 | 3111 | unfolding S | 
| 3112 | apply - | |
| 3113 | apply (rule ksimplex_replace_1) | |
| 3114 | apply (rule as)+ | |
| 3115 | unfolding as | |
| 3116 | apply auto | |
| 3117 | done | |
| 3118 | } | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3119 |     show "\<not> ((\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = p)) \<Longrightarrow> card ?S = 2"
 | 
| 53186 | 3120 | unfolding S | 
| 3121 | apply (rule ksimplex_replace_2) | |
| 3122 | apply (rule as)+ | |
| 3123 | unfolding as | |
| 3124 | apply auto | |
| 3125 | done | |
| 3126 | qed auto | |
| 3127 | qed | |
| 3128 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3129 | |
| 53688 | 3130 | subsection {* Reduced labelling *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3131 | |
| 53688 | 3132 | definition "reduced label (n::nat) (x::nat \<Rightarrow> nat) = | 
| 3133 | (SOME k. k \<le> n \<and> (\<forall>i. 1 \<le> i \<and> i < k + 1 \<longrightarrow> label x i = 0) \<and> | |
| 53186 | 3134 | (k = n \<or> label x (k + 1) \<noteq> (0::nat)))" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3135 | |
| 53186 | 3136 | lemma reduced_labelling: | 
| 3137 | shows "reduced label n x \<le> n" (is ?t1) | |
| 53688 | 3138 | and "\<forall>i. 1 \<le> i \<and> i < reduced label n x + 1 \<longrightarrow> label x i = 0" (is ?t2) | 
| 3139 | and "reduced label n x = n \<or> label x (reduced label n x + 1) \<noteq> 0" (is ?t3) | |
| 53186 | 3140 | proof - | 
| 3141 | have num_WOP: "\<And>P k. P (k::nat) \<Longrightarrow> \<exists>n. P n \<and> (\<forall>m<n. \<not> P m)" | |
| 3142 | apply (drule ex_has_least_nat[where m="\<lambda>x. x"]) | |
| 53688 | 3143 | apply (erule exE) | 
| 3144 | apply (rule_tac x=x in exI) | |
| 53186 | 3145 | apply auto | 
| 3146 | done | |
| 53688 | 3147 | have *: "n \<le> n \<and> (label x (n + 1) \<noteq> 0 \<or> n = n)" | 
| 3148 | by auto | |
| 53186 | 3149 | then guess N | 
| 3150 | apply (drule_tac num_WOP[of "\<lambda>j. j\<le>n \<and> (label x (j+1) \<noteq> 0 \<or> n = j)"]) | |
| 3151 | apply (erule exE) | |
| 3152 | done | |
| 3153 | note N = this | |
| 3154 | have N': "N \<le> n" | |
| 3155 | "\<forall>i. 1 \<le> i \<and> i < N + 1 \<longrightarrow> label x i = 0" "N = n \<or> label x (N + 1) \<noteq> 0" | |
| 3156 | defer | |
| 3157 | proof (rule, rule) | |
| 3158 | fix i | |
| 53688 | 3159 | assume i: "1 \<le> i \<and> i < N + 1" | 
| 53674 | 3160 | then show "label x i = 0" | 
| 53186 | 3161 | using N[THEN conjunct2,THEN spec[where x="i - 1"]] | 
| 53688 | 3162 | using N | 
| 3163 | by auto | |
| 53186 | 3164 | qed (insert N, auto) | 
| 3165 | show ?t1 ?t2 ?t3 | |
| 3166 | unfolding reduced_def | |
| 3167 | apply (rule_tac[!] someI2_ex) | |
| 3168 | using N' | |
| 3169 | apply (auto intro!: exI[where x=N]) | |
| 3170 | done | |
| 3171 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3172 | |
| 53186 | 3173 | lemma reduced_labelling_unique: | 
| 3174 | fixes x :: "nat \<Rightarrow> nat" | |
| 3175 | assumes "r \<le> n" | |
| 53688 | 3176 | and "\<forall>i. 1 \<le> i \<and> i < r + 1 \<longrightarrow> label x i = 0" | 
| 3177 | and "r = n \<or> label x (r + 1) \<noteq> 0" | |
| 53186 | 3178 | shows "reduced label n x = r" | 
| 3179 | apply (rule le_antisym) | |
| 3180 | apply (rule_tac[!] ccontr) | |
| 3181 | unfolding not_le | |
| 3182 | using reduced_labelling[of label n x] | |
| 3183 | using assms | |
| 3184 | apply auto | |
| 3185 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3186 | |
| 53186 | 3187 | lemma reduced_labelling_zero: | 
| 53688 | 3188 |   assumes "j \<in> {1..n}"
 | 
| 3189 | and "label x j = 0" | |
| 53186 | 3190 | shows "reduced label n x \<noteq> j - 1" | 
| 3191 | using reduced_labelling[of label n x] | |
| 53688 | 3192 | using assms | 
| 3193 | by fastforce | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3194 | |
| 53186 | 3195 | lemma reduced_labelling_nonzero: | 
| 53688 | 3196 |   assumes "j\<in>{1..n}"
 | 
| 3197 | and "label x j \<noteq> 0" | |
| 53186 | 3198 | shows "reduced label n x < j" | 
| 3199 | using assms and reduced_labelling[of label n x] | |
| 3200 | apply (erule_tac x=j in allE) | |
| 3201 | apply auto | |
| 3202 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3203 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3204 | lemma reduced_labelling_Suc: | 
| 53186 | 3205 | assumes "reduced lab (n + 1) x \<noteq> n + 1" | 
| 3206 | shows "reduced lab (n + 1) x = reduced lab n x" | |
| 3207 | apply (subst eq_commute) | |
| 3208 | apply (rule reduced_labelling_unique) | |
| 3209 | using reduced_labelling[of lab "n+1" x] and assms | |
| 3210 | apply auto | |
| 3211 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3212 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3213 | lemma complete_face_top: | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3214 |   assumes "\<forall>x\<in>f. \<forall>j\<in>{1..n+1}. x j = 0 \<longrightarrow> lab x j = 0"
 | 
| 53688 | 3215 |     and "\<forall>x\<in>f. \<forall>j\<in>{1..n+1}. x j = p \<longrightarrow> lab x j = 1"
 | 
| 3216 |   shows "reduced lab (n + 1) ` f = {0..n} \<and>
 | |
| 3217 |       ((\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = p)) \<longleftrightarrow>
 | |
| 3218 |     reduced lab (n + 1) ` f = {0..n} \<and> (\<forall>x\<in>f. x (n + 1) = p)"
 | |
| 3219 | (is "?l = ?r") | |
| 53186 | 3220 | proof | 
| 3221 | assume ?l (is "?as \<and> (?a \<or> ?b)") | |
| 53674 | 3222 | then show ?r | 
| 53186 | 3223 | apply - | 
| 53688 | 3224 | apply rule | 
| 3225 | apply (erule conjE) | |
| 3226 | apply assumption | |
| 53186 | 3227 | proof (cases ?a) | 
| 3228 | case True | |
| 3229 | then guess j .. note j = this | |
| 3230 |     {
 | |
| 3231 | fix x | |
| 3232 | assume x: "x \<in> f" | |
| 3233 | have "reduced lab (n + 1) x \<noteq> j - 1" | |
| 3234 | using j | |
| 3235 | apply - | |
| 3236 | apply (rule reduced_labelling_zero) | |
| 3237 | defer | |
| 3238 | apply (rule assms(1)[rule_format]) | |
| 3239 | using x | |
| 3240 | apply auto | |
| 3241 | done | |
| 3242 | } | |
| 53688 | 3243 |     moreover have "j - 1 \<in> {0..n}"
 | 
| 3244 | using j by auto | |
| 53185 | 3245 | then guess y unfolding `?l`[THEN conjunct1,symmetric] and image_iff .. note y = this | 
| 53688 | 3246 | ultimately have False | 
| 3247 | by auto | |
| 3248 | then show "\<forall>x\<in>f. x (n + 1) = p" | |
| 3249 | by auto | |
| 53186 | 3250 | next | 
| 3251 | case False | |
| 53688 | 3252 | then have ?b using `?l` | 
| 3253 | by blast | |
| 53186 | 3254 | then guess j .. note j = this | 
| 3255 |     {
 | |
| 3256 | fix x | |
| 3257 | assume x: "x \<in> f" | |
| 3258 | have "reduced lab (n + 1) x < j" | |
| 3259 | using j | |
| 3260 | apply - | |
| 3261 | apply (rule reduced_labelling_nonzero) | |
| 3262 | using assms(2)[rule_format,of x j] and x | |
| 3263 | apply auto | |
| 3264 | done | |
| 3265 | } note * = this | |
| 3266 | have "j = n + 1" | |
| 3267 | proof (rule ccontr) | |
| 53688 | 3268 | assume "\<not> ?thesis" | 
| 3269 | then have "j < n + 1" | |
| 3270 | using j by auto | |
| 53186 | 3271 | moreover | 
| 53688 | 3272 |       have "n \<in> {0..n}"
 | 
| 3273 | by auto | |
| 53186 | 3274 | then guess y unfolding `?l`[THEN conjunct1,symmetric] image_iff .. | 
| 53688 | 3275 | ultimately show False | 
| 3276 | using *[of y] by auto | |
| 53186 | 3277 | qed | 
| 53688 | 3278 | then show "\<forall>x\<in>f. x (n + 1) = p" | 
| 3279 | using j by auto | |
| 53186 | 3280 | qed | 
| 53688 | 3281 | 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 | 3282 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3283 | |
| 53688 | 3284 | text {* Hence we get just about the nice induction. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3285 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3286 | lemma kuhn_induction: | 
| 53688 | 3287 | assumes "0 < p" | 
| 3288 |     and "\<forall>x. \<forall>j\<in>{1..n+1}. (\<forall>j. x j \<le> p) \<and> x j = 0 \<longrightarrow> lab x j = 0"
 | |
| 3289 |     and "\<forall>x. \<forall>j\<in>{1..n+1}. (\<forall>j. x j \<le> p) \<and> x j = p \<longrightarrow> lab x j = 1"
 | |
| 3290 |     and "odd (card {f. ksimplex p n f \<and> reduced lab n ` f = {0..n}})"
 | |
| 3291 |   shows "odd (card {s. ksimplex p (n + 1) s \<and> reduced lab (n + 1) `  s = {0..n+1}})"
 | |
| 53248 | 3292 | proof - | 
| 3293 | have *: "\<And>s t. odd (card s) \<Longrightarrow> s = t \<Longrightarrow> odd (card t)" | |
| 53688 | 3294 |     "\<And>s f. (\<And>x. f x \<le> n + 1) \<Longrightarrow> f ` s \<subseteq> {0..n+1}"
 | 
| 3295 | by auto | |
| 53248 | 3296 | show ?thesis | 
| 3297 | apply (rule kuhn_simplex_lemma[unfolded mem_Collect_eq]) | |
| 53688 | 3298 | apply rule | 
| 3299 | apply rule | |
| 3300 | apply (rule *) | |
| 3301 | apply (rule reduced_labelling) | |
| 53248 | 3302 | apply (rule *(1)[OF assms(4)]) | 
| 3303 | apply (rule set_eqI) | |
| 3304 | unfolding mem_Collect_eq | |
| 53688 | 3305 | apply rule | 
| 3306 | apply (erule conjE) | |
| 53248 | 3307 | defer | 
| 3308 | apply rule | |
| 3309 | proof - | |
| 3310 | fix f | |
| 3311 |     assume as: "ksimplex p n f" "reduced lab n ` f = {0..n}"
 | |
| 3312 |     have *: "\<forall>x\<in>f. \<forall>j\<in>{1..n + 1}. x j = 0 \<longrightarrow> lab x j = 0"
 | |
| 3313 |       "\<forall>x\<in>f. \<forall>j\<in>{1..n + 1}. x j = p \<longrightarrow> lab x j = 1"
 | |
| 53688 | 3314 | using assms(2-3) | 
| 3315 | using as(1)[unfolded ksimplex_def] | |
| 3316 | by auto | |
| 53248 | 3317 | have allp: "\<forall>x\<in>f. x (n + 1) = p" | 
| 53688 | 3318 | using assms(2) | 
| 3319 | using as(1)[unfolded ksimplex_def] | |
| 3320 | by auto | |
| 53248 | 3321 |     {
 | 
| 3322 | fix x | |
| 3323 | assume "x \<in> f" | |
| 53674 | 3324 | then have "reduced lab (n + 1) x < n + 1" | 
| 53248 | 3325 | apply - | 
| 3326 | apply (rule reduced_labelling_nonzero) | |
| 53688 | 3327 | defer | 
| 3328 | using assms(3) | |
| 3329 | using as(1)[unfolded ksimplex_def] | |
| 53248 | 3330 | apply auto | 
| 3331 | done | |
| 53674 | 3332 | then have "reduced lab (n + 1) x = reduced lab n x" | 
| 53248 | 3333 | apply - | 
| 3334 | apply (rule reduced_labelling_Suc) | |
| 3335 | using reduced_labelling(1) | |
| 3336 | apply auto | |
| 3337 | done | |
| 3338 | } | |
| 53674 | 3339 |     then have "reduced lab (n + 1) ` f = {0..n}"
 | 
| 53248 | 3340 | unfolding as(2)[symmetric] | 
| 3341 | apply - | |
| 3342 | apply (rule set_eqI) | |
| 3343 | unfolding image_iff | |
| 3344 | apply auto | |
| 3345 | done | |
| 3346 | moreover guess s using as(1)[unfolded simplex_top_face[OF assms(1) allp,symmetric]] .. | |
| 3347 | then guess a .. | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3348 | ultimately show "\<exists>s a. ksimplex p (n + 1) s \<and> | 
| 53846 | 3349 |         a \<in> s \<and> f = s - {a} \<and>
 | 
| 3350 |         reduced lab (n + 1) ` f = {0..n} \<and>
 | |
| 3351 |         ((\<exists>j\<in>{1..n + 1}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{1..n + 1}. \<forall>x\<in>f. x j = p))"
 | |
| 53248 | 3352 | apply (rule_tac x = s in exI) | 
| 3353 | apply (rule_tac x = a in exI) | |
| 3354 | unfolding complete_face_top[OF *] | |
| 3355 | using allp as(1) | |
| 3356 | apply auto | |
| 3357 | done | |
| 3358 | next | |
| 3359 | fix f | |
| 3360 | assume as: "\<exists>s a. ksimplex p (n + 1) s \<and> | |
| 53846 | 3361 |       a \<in> s \<and> f = s - {a} \<and> reduced lab (n + 1) ` f = {0..n} \<and>
 | 
| 3362 |       ((\<exists>j\<in>{1..n + 1}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{1..n + 1}. \<forall>x\<in>f. x j = p))"
 | |
| 53248 | 3363 | then guess s .. | 
| 53688 | 3364 | then guess a by (elim exE conjE) note sa = this | 
| 53248 | 3365 |     {
 | 
| 3366 | fix x | |
| 3367 | assume "x \<in> f" | |
| 53674 | 3368 | then have "reduced lab (n + 1) x \<in> reduced lab (n + 1) ` f" | 
| 53248 | 3369 | by auto | 
| 53674 | 3370 | then have "reduced lab (n + 1) x < n + 1" | 
| 53248 | 3371 | using sa(4) by auto | 
| 53674 | 3372 | then have "reduced lab (n + 1) x = reduced lab n x" | 
| 53248 | 3373 | apply - | 
| 3374 | apply (rule reduced_labelling_Suc) | |
| 3375 | using reduced_labelling(1) | |
| 3376 | apply auto | |
| 3377 | done | |
| 3378 | } | |
| 53674 | 3379 |     then show part1: "reduced lab n ` f = {0..n}"
 | 
| 53248 | 3380 | unfolding sa(4)[symmetric] | 
| 3381 | apply - | |
| 3382 | apply (rule set_eqI) | |
| 3383 | unfolding image_iff | |
| 3384 | apply auto | |
| 3385 | done | |
| 3386 | have *: "\<forall>x\<in>f. x (n + 1) = p" | |
| 3387 |     proof (cases "\<exists>j\<in>{1..n + 1}. \<forall>x\<in>f. x j = 0")
 | |
| 3388 | case True | |
| 3389 | then guess j .. | |
| 53674 | 3390 | then have "\<And>x. x \<in> f \<Longrightarrow> reduced lab (n + 1) x \<noteq> j - 1" | 
| 53248 | 3391 | apply - | 
| 3392 | apply (rule reduced_labelling_zero) | |
| 3393 | apply assumption | |
| 3394 | apply (rule assms(2)[rule_format]) | |
| 3395 | using sa(1)[unfolded ksimplex_def] | |
| 3396 | unfolding sa | |
| 3397 | apply auto | |
| 3398 | done | |
| 3399 | moreover | |
| 53688 | 3400 |       have "j - 1 \<in> {0..n}"
 | 
| 3401 |         using `j\<in>{1..n+1}` by auto
 | |
| 53248 | 3402 | ultimately have False | 
| 3403 | unfolding sa(4)[symmetric] | |
| 3404 | unfolding image_iff | |
| 3405 | by fastforce | |
| 53688 | 3406 | then show ?thesis | 
| 3407 | by auto | |
| 53248 | 3408 | next | 
| 3409 | case False | |
| 53674 | 3410 |       then have "\<exists>j\<in>{1..n + 1}. \<forall>x\<in>f. x j = p"
 | 
| 53248 | 3411 | using sa(5) by fastforce then guess j .. note j=this | 
| 53674 | 3412 | then show ?thesis | 
| 53248 | 3413 | proof (cases "j = n + 1") | 
| 53688 | 3414 | case False | 
| 3415 |         then have *: "j \<in> {1..n}"
 | |
| 53248 | 3416 | using j by auto | 
| 53674 | 3417 | then have "\<And>x. x \<in> f \<Longrightarrow> reduced lab n x < j" | 
| 53248 | 3418 | apply (rule reduced_labelling_nonzero) | 
| 3419 | proof - | |
| 3420 | fix x | |
| 3421 | assume "x \<in> f" | |
| 53674 | 3422 | then have "lab x j = 1" | 
| 53248 | 3423 | apply - | 
| 3424 | apply (rule assms(3)[rule_format,OF j(1)]) | |
| 3425 | using sa(1)[unfolded ksimplex_def] | |
| 3426 | using j | |
| 3427 | unfolding sa | |
| 3428 | apply auto | |
| 3429 | done | |
| 53688 | 3430 | then show "lab x j \<noteq> 0" | 
| 3431 | by auto | |
| 53248 | 3432 | qed | 
| 53688 | 3433 |         moreover have "j \<in> {0..n}"
 | 
| 3434 | using * by auto | |
| 53248 | 3435 | ultimately have False | 
| 3436 | unfolding part1[symmetric] | |
| 3437 | using * unfolding image_iff | |
| 3438 | by auto | |
| 53688 | 3439 | then show ?thesis | 
| 3440 | by auto | |
| 53248 | 3441 | qed auto | 
| 3442 | qed | |
| 53674 | 3443 | then show "ksimplex p n f" | 
| 53688 | 3444 | using as | 
| 3445 | unfolding simplex_top_face[OF assms(1) *,symmetric] | |
| 3446 | by auto | |
| 53248 | 3447 | qed | 
| 3448 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3449 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3450 | lemma kuhn_induction_Suc: | 
| 53688 | 3451 | assumes "0 < p" | 
| 3452 |     and "\<forall>x. \<forall>j\<in>{1..Suc n}. (\<forall>j. x j \<le> p) \<and> x j = 0 \<longrightarrow> lab x j = 0"
 | |
| 3453 |     and "\<forall>x. \<forall>j\<in>{1..Suc n}. (\<forall>j. x j \<le> p) \<and> x j = p \<longrightarrow> lab x j = 1"
 | |
| 3454 |     and "odd (card {f. ksimplex p n f \<and> reduced lab n ` f = {0..n}})"
 | |
| 53846 | 3455 |   shows "odd (card {s. ksimplex p (Suc n) s \<and> reduced lab (Suc n) ` s = {0..Suc n}})"
 | 
| 53688 | 3456 | using assms | 
| 3457 | unfolding Suc_eq_plus1 | |
| 3458 | by (rule kuhn_induction) | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3459 | |
| 53248 | 3460 | |
| 53688 | 3461 | text {* And so we get the final combinatorial result. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3462 | |
| 53688 | 3463 | lemma ksimplex_0: "ksimplex p 0 s \<longleftrightarrow> s = {(\<lambda>x. p)}"
 | 
| 3464 | (is "?l = ?r") | |
| 53248 | 3465 | proof | 
| 3466 | assume l: ?l | |
| 3467 | guess a using ksimplexD(3)[OF l, unfolded add_0] unfolding card_1_exists .. note a = this | |
| 3468 | have "a = (\<lambda>x. p)" | |
| 53688 | 3469 | using ksimplexD(5)[OF l, rule_format, OF a(1)] | 
| 3470 | by rule auto | |
| 3471 | then show ?r | |
| 3472 | using a by auto | |
| 53248 | 3473 | next | 
| 3474 | assume r: ?r | |
| 53688 | 3475 | show ?l | 
| 3476 | unfolding r ksimplex_eq by auto | |
| 53248 | 3477 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3478 | |
| 53248 | 3479 | lemma reduce_labelling_zero[simp]: "reduced lab 0 x = 0" | 
| 3480 | 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 | 3481 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3482 | lemma kuhn_combinatorial: | 
| 53688 | 3483 | assumes "0 < p" | 
| 3484 | and "\<forall>x j. (\<forall>j. x j \<le> p) \<and> 1 \<le> j \<and> j \<le> n \<and> x j = 0 \<longrightarrow> lab x j = 0" | |
| 3485 | and "\<forall>x j. (\<forall>j. x j \<le> p) \<and> 1 \<le> j \<and> j \<le> n \<and> x j = p \<longrightarrow> lab x j = 1" | |
| 3486 |   shows "odd (card {s. ksimplex p n s \<and> reduced lab n ` s = {0..n}})"
 | |
| 53248 | 3487 | using assms | 
| 3488 | proof (induct n) | |
| 53688 | 3489 |   let ?M = "\<lambda>n. {s. ksimplex p n s \<and> reduced lab n ` s = {0..n}}"
 | 
| 53248 | 3490 |   {
 | 
| 3491 | case 0 | |
| 53688 | 3492 |     have *: "?M 0 = {{\<lambda>x. p}}"
 | 
| 53248 | 3493 | unfolding ksimplex_0 by auto | 
| 53688 | 3494 | show ?case | 
| 3495 | unfolding * by auto | |
| 53248 | 3496 | next | 
| 3497 | case (Suc n) | |
| 3498 | have "odd (card (?M n))" | |
| 3499 | apply (rule Suc(1)[OF Suc(2)]) | |
| 3500 | using Suc(3-) | |
| 3501 | apply auto | |
| 3502 | done | |
| 53674 | 3503 | then show ?case | 
| 53248 | 3504 | apply - | 
| 3505 | apply (rule kuhn_induction_Suc) | |
| 3506 | using Suc(2-) | |
| 3507 | apply auto | |
| 3508 | done | |
| 3509 | } | |
| 3510 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3511 | |
| 53248 | 3512 | lemma kuhn_lemma: | 
| 53688 | 3513 | fixes n p :: nat | 
| 3514 | assumes "0 < p" | |
| 3515 | and "0 < n" | |
| 3516 |     and "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. label x i = (0::nat) \<or> label x i = 1)"
 | |
| 3517 |     and "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. x i = 0 \<longrightarrow> label x i = 0)"
 | |
| 3518 |     and "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. x i = p \<longrightarrow> label x i = 1)"
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3519 |   obtains q where "\<forall>i\<in>{1..n}. q i < p"
 | 
| 53688 | 3520 |     and "\<forall>i\<in>{1..n}. \<exists>r s.
 | 
| 3521 |       (\<forall>j\<in>{1..n}. q j \<le> r j \<and> r j \<le> q j + 1) \<and>
 | |
| 3522 |       (\<forall>j\<in>{1..n}. q j \<le> s j \<and> s j \<le> q j + 1) \<and>
 | |
| 3523 | label r i \<noteq> label s i" | |
| 53248 | 3524 | proof - | 
| 3525 |   let ?A = "{s. ksimplex p n s \<and> reduced label n ` s = {0..n}}"
 | |
| 53688 | 3526 | have "n \<noteq> 0" | 
| 3527 | using assms by auto | |
| 3528 | have conjD: "\<And>P Q. P \<and> Q \<Longrightarrow> P" "\<And>P Q. P \<and> Q \<Longrightarrow> Q" | |
| 53248 | 3529 | by auto | 
| 3530 | have "odd (card ?A)" | |
| 3531 | apply (rule kuhn_combinatorial[of p n label]) | |
| 3532 | using assms | |
| 3533 | apply auto | |
| 3534 | done | |
| 53674 | 3535 | then have "card ?A \<noteq> 0" | 
| 53248 | 3536 | apply - | 
| 3537 | apply (rule ccontr) | |
| 3538 | apply auto | |
| 3539 | done | |
| 53688 | 3540 |   then have "?A \<noteq> {}"
 | 
| 3541 | unfolding card_eq_0_iff by auto | |
| 53248 | 3542 | then obtain s where "s \<in> ?A" | 
| 3543 | by auto note s=conjD[OF this[unfolded mem_Collect_eq]] | |
| 53688 | 3544 | guess a b by (rule ksimplex_extrema_strong[OF s(1) `n \<noteq> 0`]) note ab = this | 
| 53248 | 3545 | show ?thesis | 
| 3546 | apply (rule that[of a]) | |
| 3547 | apply (rule_tac[!] ballI) | |
| 3548 | proof - | |
| 3549 | fix i | |
| 53688 | 3550 |     assume "i \<in> {1..n}"
 | 
| 53674 | 3551 | then have "a i + 1 \<le> p" | 
| 53248 | 3552 | apply - | 
| 3553 | apply (rule order_trans[of _ "b i"]) | |
| 3554 | apply (subst ab(5)[THEN spec[where x=i]]) | |
| 3555 | using s(1)[unfolded ksimplex_def] | |
| 3556 | defer | |
| 3557 | apply - | |
| 3558 | apply (erule conjE)+ | |
| 3559 | apply (drule_tac bspec[OF _ ab(2)])+ | |
| 3560 | apply auto | |
| 3561 | done | |
| 53688 | 3562 | then show "a i < p" | 
| 3563 | by auto | |
| 53248 | 3564 | next | 
| 3565 | case goal2 | |
| 53688 | 3566 | then have "i \<in> reduced label n ` s" | 
| 3567 | using s by auto | |
| 53248 | 3568 | then guess u unfolding image_iff .. note u = this | 
| 3569 | from goal2 have "i - 1 \<in> reduced label n ` s" | |
| 3570 | using s by auto | |
| 3571 | then guess v unfolding image_iff .. note v = this | |
| 3572 | show ?case | |
| 3573 | apply (rule_tac x = u in exI) | |
| 3574 | apply (rule_tac x = v in exI) | |
| 3575 | apply (rule conjI) | |
| 3576 | defer | |
| 3577 | apply (rule conjI) | |
| 3578 | defer 2 | |
| 3579 | apply (rule_tac[1-2] ballI) | |
| 3580 | proof - | |
| 3581 | show "label u i \<noteq> label v i" | |
| 3582 | using reduced_labelling [of label n u] reduced_labelling [of label n v] | |
| 3583 | unfolding u(2)[symmetric] v(2)[symmetric] | |
| 3584 | using goal2 | |
| 3585 | apply auto | |
| 3586 | done | |
| 3587 | fix j | |
| 3588 |       assume j: "j \<in> {1..n}"
 | |
| 53688 | 3589 | show "a j \<le> u j \<and> u j \<le> a j + 1" and "a j \<le> v j \<and> v j \<le> a j + 1" | 
| 53248 | 3590 | using conjD[OF ab(4)[rule_format, OF u(1)]] | 
| 3591 | and conjD[OF ab(4)[rule_format, OF v(1)]] | |
| 3592 | apply - | |
| 3593 | apply (drule_tac[!] kle_imp_pointwise)+ | |
| 3594 | apply (erule_tac[!] x=j in allE)+ | |
| 3595 | unfolding ab(5)[rule_format] | |
| 3596 | using j | |
| 3597 | apply auto | |
| 3598 | done | |
| 3599 | qed | |
| 3600 | qed | |
| 3601 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3602 | |
| 53185 | 3603 | |
| 53688 | 3604 | subsection {* The main result for the unit cube *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3605 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3606 | lemma kuhn_labelling_lemma': | 
| 53688 | 3607 | assumes "(\<forall>x::nat\<Rightarrow>real. P x \<longrightarrow> P (f x))" | 
| 3608 | 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 | 3609 | shows "\<exists>l. (\<forall>x i. l x i \<le> (1::nat)) \<and> | 
| 53688 | 3610 | (\<forall>x i. P x \<and> Q i \<and> x i = 0 \<longrightarrow> l x i = 0) \<and> | 
| 3611 | (\<forall>x i. P x \<and> Q i \<and> x i = 1 \<longrightarrow> l x i = 1) \<and> | |
| 3612 | (\<forall>x i. P x \<and> Q i \<and> l x i = 0 \<longrightarrow> x i \<le> f x i) \<and> | |
| 3613 | (\<forall>x i. P x \<and> Q i \<and> l x i = 1 \<longrightarrow> f x i \<le> x i)" | |
| 53185 | 3614 | proof - | 
| 53688 | 3615 | have and_forall_thm: "\<And>P Q. (\<forall>x. P x) \<and> (\<forall>x. Q x) \<longleftrightarrow> (\<forall>x. P x \<and> Q x)" | 
| 3616 | by auto | |
| 3617 | 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 | 3618 | by auto | 
| 3619 | show ?thesis | |
| 3620 | unfolding and_forall_thm | |
| 3621 | apply (subst choice_iff[symmetric])+ | |
| 53688 | 3622 | apply rule | 
| 3623 | apply rule | |
| 3624 | proof - | |
| 53185 | 3625 | case goal1 | 
| 53688 | 3626 | let ?R = "\<lambda>y::nat. | 
| 3627 | (P x \<and> Q xa \<and> x xa = 0 \<longrightarrow> y = 0) \<and> | |
| 53185 | 3628 | (P x \<and> Q xa \<and> x xa = 1 \<longrightarrow> y = 1) \<and> | 
| 3629 | (P x \<and> Q xa \<and> y = 0 \<longrightarrow> x xa \<le> (f x) xa) \<and> | |
| 3630 | (P x \<and> Q xa \<and> y = 1 \<longrightarrow> (f x) xa \<le> x xa)" | |
| 3631 |     {
 | |
| 53688 | 3632 | assume "P x" and "Q xa" | 
| 3633 | then have "0 \<le> f x xa \<and> f x xa \<le> 1" | |
| 53185 | 3634 | using assms(2)[rule_format,of "f x" xa] | 
| 3635 | apply (drule_tac assms(1)[rule_format]) | |
| 3636 | apply auto | |
| 3637 | done | |
| 3638 | } | |
| 53688 | 3639 | then have "?R 0 \<or> ?R 1" | 
| 3640 | by auto | |
| 3641 | then show ?case | |
| 3642 | by auto | |
| 53185 | 3643 | qed | 
| 3644 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3645 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3646 | lemma brouwer_cube: | 
| 53185 | 3647 | fixes f :: "'a::ordered_euclidean_space \<Rightarrow> 'a::ordered_euclidean_space" | 
| 53674 | 3648 |   assumes "continuous_on {0..(\<Sum>Basis)} f"
 | 
| 3649 |     and "f ` {0..(\<Sum>Basis)} \<subseteq> {0..(\<Sum>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 | 3650 |   shows "\<exists>x\<in>{0..(\<Sum>Basis)}. f x = x"
 | 
| 53185 | 3651 | proof (rule ccontr) | 
| 3652 |   def n \<equiv> "DIM('a)"
 | |
| 3653 | have n: "1 \<le> n" "0 < n" "n \<noteq> 0" | |
| 3654 | unfolding n_def by (auto simp add: Suc_le_eq DIM_positive) | |
| 53674 | 3655 | assume "\<not> ?thesis" | 
| 3656 |   then have *: "\<not> (\<exists>x\<in>{0..\<Sum>Basis}. f x - x = 0)"
 | |
| 3657 | by auto | |
| 53185 | 3658 | guess d | 
| 3659 | apply (rule brouwer_compactness_lemma[OF compact_interval _ *]) | |
| 3660 | apply (rule continuous_on_intros assms)+ | |
| 3661 | done | |
| 3662 | note d = this [rule_format] | |
| 3663 |   have *: "\<forall>x. x \<in> {0..\<Sum>Basis} \<longrightarrow> f x \<in> {0..\<Sum>Basis}"  "\<forall>x. x \<in> {0..(\<Sum>Basis)::'a} \<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 | 3664 | (\<forall>i\<in>Basis. True \<longrightarrow> 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)" | 
| 53185 | 3665 | using assms(2)[unfolded image_subset_iff Ball_def] | 
| 3666 | unfolding mem_interval by auto | |
| 3667 | guess label using kuhn_labelling_lemma[OF *] by (elim exE conjE) | |
| 3668 | note label = this [rule_format] | |
| 53674 | 3669 |   have lem1: "\<forall>x\<in>{0..\<Sum>Basis}.\<forall>y\<in>{0..\<Sum>Basis}.\<forall>i\<in>Basis. label x i \<noteq> label y i \<longrightarrow>
 | 
| 3670 | abs (f x \<bullet> i - x \<bullet> i) \<le> norm (f y - f x) + norm (y - x)" | |
| 53185 | 3671 | proof safe | 
| 3672 | fix x y :: 'a | |
| 53688 | 3673 |     assume x: "x \<in> {0..\<Sum>Basis}"
 | 
| 3674 |     assume y: "y \<in> {0..\<Sum>Basis}"
 | |
| 53185 | 3675 | fix i | 
| 3676 | assume i: "label x i \<noteq> label y i" "i \<in> Basis" | |
| 3677 | have *: "\<And>x y fx fy :: real. x \<le> fx \<and> fy \<le> y \<or> fx \<le> x \<and> y \<le> fy \<Longrightarrow> | |
| 3678 | abs (fx - x) \<le> abs (fy - fx) + abs (y - x)" by auto | |
| 53688 | 3679 | have "\<bar>(f x - x) \<bullet> i\<bar> \<le> abs ((f y - f x)\<bullet>i) + abs ((y - x)\<bullet>i)" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3680 | unfolding inner_simps | 
| 53185 | 3681 | apply (rule *) | 
| 3682 | apply (cases "label x i = 0") | |
| 53688 | 3683 | apply (rule disjI1) | 
| 3684 | apply rule | |
| 53185 | 3685 | prefer 3 | 
| 53688 | 3686 | apply (rule disjI2) | 
| 3687 | apply rule | |
| 3688 | proof - | |
| 53185 | 3689 | assume lx: "label x i = 0" | 
| 53674 | 3690 | then have ly: "label y i = 1" | 
| 53688 | 3691 | using i label(1)[of i y] | 
| 3692 | by auto | |
| 53185 | 3693 | show "x \<bullet> i \<le> f x \<bullet> i" | 
| 3694 | apply (rule label(4)[rule_format]) | |
| 53674 | 3695 | using x y lx i(2) | 
| 53252 | 3696 | apply auto | 
| 53185 | 3697 | done | 
| 3698 | show "f y \<bullet> i \<le> y \<bullet> i" | |
| 3699 | apply (rule label(5)[rule_format]) | |
| 53674 | 3700 | using x y ly i(2) | 
| 53252 | 3701 | apply auto | 
| 53185 | 3702 | done | 
| 3703 | next | |
| 3704 | assume "label x i \<noteq> 0" | |
| 53688 | 3705 | then have l: "label x i = 1" "label y i = 0" | 
| 3706 | using i label(1)[of i x] label(1)[of i y] | |
| 3707 | by auto | |
| 53185 | 3708 | show "f x \<bullet> i \<le> x \<bullet> i" | 
| 3709 | apply (rule label(5)[rule_format]) | |
| 53674 | 3710 | using x y l i(2) | 
| 53252 | 3711 | apply auto | 
| 53185 | 3712 | done | 
| 3713 | show "y \<bullet> i \<le> f y \<bullet> i" | |
| 3714 | apply (rule label(4)[rule_format]) | |
| 53674 | 3715 | using x y l i(2) | 
| 53252 | 3716 | apply auto | 
| 53185 | 3717 | done | 
| 3718 | qed | |
| 3719 | also have "\<dots> \<le> norm (f y - f x) + norm (y - x)" | |
| 3720 | apply (rule add_mono) | |
| 3721 | apply (rule Basis_le_norm[OF i(2)])+ | |
| 3722 | done | |
| 3723 | finally show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)" | |
| 3724 | 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 | 3725 | qed | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3726 |   have "\<exists>e>0. \<forall>x\<in>{0..\<Sum>Basis}. \<forall>y\<in>{0..\<Sum>Basis}. \<forall>z\<in>{0..\<Sum>Basis}. \<forall>i\<in>Basis.
 | 
| 53688 | 3727 | norm (x - z) < e \<and> norm (y - z) < e \<and> label x i \<noteq> label y i \<longrightarrow> | 
| 3728 | abs ((f(z) - z)\<bullet>i) < d / (real n)" | |
| 53185 | 3729 | proof - | 
| 53688 | 3730 | have d': "d / real n / 8 > 0" | 
| 53185 | 3731 | apply (rule divide_pos_pos)+ | 
| 53688 | 3732 | using d(1) | 
| 3733 | unfolding n_def | |
| 53185 | 3734 | apply (auto simp: DIM_positive) | 
| 3735 | done | |
| 3736 |     have *: "uniformly_continuous_on {0..\<Sum>Basis} f"
 | |
| 3737 | by (rule compact_uniformly_continuous[OF assms(1) compact_interval]) | |
| 3738 | guess e using *[unfolded uniformly_continuous_on_def,rule_format,OF d'] by (elim exE conjE) | |
| 36587 | 3739 | note e=this[rule_format,unfolded dist_norm] | 
| 53185 | 3740 | show ?thesis | 
| 3741 | apply (rule_tac x="min (e/2) (d/real n/8)" in exI) | |
| 53248 | 3742 | apply safe | 
| 3743 | proof - | |
| 53185 | 3744 | show "0 < min (e / 2) (d / real n / 8)" | 
| 3745 | using d' e by auto | |
| 3746 | fix x y z i | |
| 53688 | 3747 | assume as: | 
| 3748 |         "x \<in> {0..\<Sum>Basis}" "y \<in> {0..\<Sum>Basis}" "z \<in> {0..\<Sum>Basis}"
 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36587diff
changeset | 3749 | "norm (x - z) < min (e / 2) (d / real n / 8)" | 
| 53688 | 3750 | "norm (y - z) < min (e / 2) (d / real n / 8)" | 
| 3751 | "label x i \<noteq> label y i" | |
| 3752 | assume i: "i \<in> Basis" | |
| 53185 | 3753 | have *: "\<And>z fz x fx n1 n2 n3 n4 d4 d :: real. abs(fx - x) \<le> n1 + n2 \<Longrightarrow> | 
| 53688 | 3754 | abs (fx - fz) \<le> n3 \<Longrightarrow> abs (x - z) \<le> n4 \<Longrightarrow> | 
| 53185 | 3755 | n1 < d4 \<Longrightarrow> n2 < 2 * d4 \<Longrightarrow> n3 < d4 \<Longrightarrow> n4 < d4 \<Longrightarrow> | 
| 53688 | 3756 | (8 * d4 = d) \<Longrightarrow> abs(fz - z) < d" | 
| 3757 | by auto | |
| 3758 | show "\<bar>(f z - z) \<bullet> i\<bar> < d / real n" | |
| 3759 | unfolding inner_simps | |
| 53185 | 3760 | proof (rule *) | 
| 3761 | show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y -f x) + norm (y - x)" | |
| 3762 | apply (rule lem1[rule_format]) | |
| 53688 | 3763 | using as i | 
| 3764 | apply auto | |
| 53185 | 3765 | 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 | 3766 | 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)" | 
| 53688 | 3767 | unfolding inner_diff_left[symmetric] | 
| 3768 | by (rule Basis_le_norm[OF i])+ | |
| 3769 | have tria: "norm (y - x) \<le> norm (y - z) + norm (x - z)" | |
| 53185 | 3770 | using dist_triangle[of y x z, unfolded dist_norm] | 
| 53688 | 3771 | unfolding norm_minus_commute | 
| 3772 | by auto | |
| 53185 | 3773 | also have "\<dots> < e / 2 + e / 2" | 
| 3774 | apply (rule add_strict_mono) | |
| 53252 | 3775 | using as(4,5) | 
| 3776 | apply auto | |
| 53185 | 3777 | done | 
| 3778 | finally show "norm (f y - f x) < d / real n / 8" | |
| 3779 | apply - | |
| 3780 | apply (rule e(2)) | |
| 53252 | 3781 | using as | 
| 3782 | apply auto | |
| 53185 | 3783 | done | 
| 3784 | have "norm (y - z) + norm (x - z) < d / real n / 8 + d / real n / 8" | |
| 3785 | apply (rule add_strict_mono) | |
| 53252 | 3786 | using as | 
| 3787 | apply auto | |
| 53185 | 3788 | done | 
| 53688 | 3789 | then show "norm (y - x) < 2 * (d / real n / 8)" | 
| 3790 | using tria | |
| 3791 | by auto | |
| 53185 | 3792 | show "norm (f x - f z) < d / real n / 8" | 
| 3793 | apply (rule e(2)) | |
| 53252 | 3794 | using as e(1) | 
| 3795 | apply auto | |
| 53185 | 3796 | done | 
| 3797 | qed (insert as, auto) | |
| 3798 | qed | |
| 3799 | qed | |
| 3800 | then guess e by (elim exE conjE) note e=this[rule_format] | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3801 | guess p using real_arch_simple[of "1 + real n / e"] .. note p=this | 
| 53185 | 3802 | have "1 + real n / e > 0" | 
| 3803 | apply (rule add_pos_pos) | |
| 3804 | defer | |
| 3805 | apply (rule divide_pos_pos) | |
| 53252 | 3806 | using e(1) n | 
| 3807 | apply auto | |
| 53185 | 3808 | done | 
| 53688 | 3809 | then have "p > 0" | 
| 3810 | 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 | 3811 | |
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3812 |   obtain b :: "nat \<Rightarrow> 'a" where b: "bij_betw b {1..n} Basis"
 | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3813 | by atomize_elim (auto simp: n_def intro!: finite_same_card_bij) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3814 |   def b' \<equiv> "inv_into {1..n} b"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3815 |   then have b': "bij_betw b' Basis {1..n}"
 | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3816 | using bij_betw_inv_into[OF b] by auto | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3817 |   then have b'_Basis: "\<And>i. i \<in> Basis \<Longrightarrow> b' i \<in> {Suc 0 .. n}"
 | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3818 | 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 | 3819 | have bb'[simp]:"\<And>i. i \<in> Basis \<Longrightarrow> b (b' i) = i" | 
| 53688 | 3820 | unfolding b'_def | 
| 3821 | using b | |
| 3822 | by (auto simp: f_inv_into_f bij_betw_def) | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3823 |   have b'b[simp]:"\<And>i. i \<in> {1..n} \<Longrightarrow> b' (b i) = i"
 | 
| 53688 | 3824 | unfolding b'_def | 
| 3825 | using b | |
| 3826 | by (auto simp: inv_into_f_eq bij_betw_def) | |
| 3827 | have *: "\<And>x :: nat. x = 0 \<or> x = 1 \<longleftrightarrow> x \<le> 1" | |
| 3828 | by auto | |
| 53185 | 3829 |   have b'': "\<And>j. j \<in> {Suc 0..n} \<Longrightarrow> b j \<in> Basis"
 | 
| 3830 | using b unfolding bij_betw_def by auto | |
| 3831 |   have q1: "0 < p" "0 < n"  "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> 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 | 3832 |     (\<forall>i\<in>{1..n}. (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0 \<or>
 | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3833 | (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)" | 
| 53688 | 3834 | unfolding * | 
| 3835 | using `p > 0` `n > 0` | |
| 3836 | using label(1)[OF b''] | |
| 3837 | by auto | |
| 53185 | 3838 |   have q2: "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..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 | 3839 | (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0)" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3840 |     "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. x i = p \<longrightarrow>
 | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3841 | (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)" | 
| 53688 | 3842 | apply rule | 
| 3843 | apply rule | |
| 3844 | apply rule | |
| 3845 | apply rule | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3846 | defer | 
| 53688 | 3847 | apply rule | 
| 3848 | apply rule | |
| 3849 | apply rule | |
| 3850 | apply rule | |
| 3851 | proof - | |
| 53185 | 3852 | fix x i | 
| 3853 |     assume as: "\<forall>i\<in>{1..n}. x i \<le> p" "i \<in> {1..n}"
 | |
| 3854 |     {
 | |
| 3855 | assume "x i = p \<or> x i = 0" | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3856 |       have "(\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<in> {0::'a..\<Sum>Basis}"
 | 
| 53688 | 3857 | unfolding mem_interval | 
| 3858 | using as b'_Basis | |
| 53185 | 3859 | by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) | 
| 3860 | } | |
| 3861 | note cube = this | |
| 3862 |     {
 | |
| 3863 | assume "x i = p" | |
| 53674 | 3864 | then show "(label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1" | 
| 53688 | 3865 | unfolding o_def | 
| 3866 | using cube as `p > 0` | |
| 53185 | 3867 | by (intro label(3)) (auto simp add: b'') | 
| 3868 | } | |
| 3869 |     {
 | |
| 3870 | assume "x i = 0" | |
| 53674 | 3871 | then show "(label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0" | 
| 53688 | 3872 | unfolding o_def using cube as `p > 0` | 
| 53185 | 3873 | by (intro label(2)) (auto simp add: b'') | 
| 3874 | } | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36587diff
changeset | 3875 | qed | 
| 53185 | 3876 | guess q by (rule kuhn_lemma[OF q1 q2]) note q = this | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3877 | def z \<equiv> "(\<Sum>i\<in>Basis. (real (q (b' i)) / real p) *\<^sub>R i)::'a" | 
| 53688 | 3878 | have "\<exists>i\<in>Basis. d / real n \<le> abs ((f z - z)\<bullet>i)" | 
| 53185 | 3879 | 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 | 3880 |     have "\<forall>i\<in>Basis. q (b' i) \<in> {0..p}"
 | 
| 53688 | 3881 | using q(1) b' | 
| 3882 | by (auto intro: less_imp_le simp: bij_betw_def) | |
| 3883 |     then have "z \<in> {0..\<Sum>Basis}"
 | |
| 3884 | unfolding z_def mem_interval | |
| 3885 | using b'_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 | 3886 | by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) | 
| 53688 | 3887 | then have d_fz_z: "d \<le> norm (f z - z)" | 
| 3888 | by (rule d) | |
| 3889 | assume "\<not> ?thesis" | |
| 53674 | 3890 | then have as: "\<forall>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar> < d / real n" | 
| 53688 | 3891 | using `n > 0` | 
| 3892 | by (auto simp add: not_le 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 | 3893 | have "norm (f z - z) \<le> (\<Sum>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar>)" | 
| 53688 | 3894 | unfolding inner_diff_left[symmetric] | 
| 3895 | by (rule norm_le_l1) | |
| 53185 | 3896 | also have "\<dots> < (\<Sum>(i::'a) \<in> Basis. d / real n)" | 
| 3897 | apply (rule setsum_strict_mono) | |
| 53688 | 3898 | using as | 
| 3899 | apply auto | |
| 53185 | 3900 | done | 
| 3901 | also have "\<dots> = d" | |
| 53688 | 3902 | using DIM_positive[where 'a='a] | 
| 3903 | by (auto simp: real_eq_of_nat n_def) | |
| 3904 | finally show False | |
| 3905 | using d_fz_z by auto | |
| 53185 | 3906 | qed | 
| 3907 | then guess i .. note i = this | |
| 3908 |   have *: "b' i \<in> {1..n}"
 | |
| 53688 | 3909 | using i | 
| 3910 | using b'[unfolded bij_betw_def] | |
| 3911 | by auto | |
| 53185 | 3912 | guess r using q(2)[rule_format,OF *] .. | 
| 3913 | then guess s by (elim exE conjE) note rs = this[rule_format] | |
| 3914 |   have b'_im: "\<And>i. i \<in> Basis \<Longrightarrow>  b' i \<in> {1..n}"
 | |
| 3915 | using b' unfolding bij_betw_def 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 | 3916 | def r' \<equiv> "(\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i)::'a" | 
| 53185 | 3917 | have "\<And>i. i \<in> Basis \<Longrightarrow> r (b' i) \<le> p" | 
| 3918 | apply (rule order_trans) | |
| 3919 | apply (rule rs(1)[OF b'_im,THEN conjunct2]) | |
| 53252 | 3920 | using q(1)[rule_format,OF b'_im] | 
| 3921 | apply (auto simp add: Suc_le_eq) | |
| 53185 | 3922 | done | 
| 53674 | 3923 |   then have "r' \<in> {0..\<Sum>Basis}"
 | 
| 53688 | 3924 | unfolding r'_def mem_interval | 
| 3925 | using b'_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 | 3926 | by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3927 | def s' \<equiv> "(\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i)::'a" | 
| 53688 | 3928 | have "\<And>i. i \<in> Basis \<Longrightarrow> s (b' i) \<le> p" | 
| 53185 | 3929 | apply (rule order_trans) | 
| 3930 | apply (rule rs(2)[OF b'_im, THEN conjunct2]) | |
| 53252 | 3931 | using q(1)[rule_format,OF b'_im] | 
| 3932 | apply (auto simp add: Suc_le_eq) | |
| 53185 | 3933 | done | 
| 53674 | 3934 |   then have "s' \<in> {0..\<Sum>Basis}"
 | 
| 53688 | 3935 | unfolding s'_def mem_interval | 
| 3936 | using b'_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 | 3937 | by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) | 
| 53185 | 3938 |   have "z \<in> {0..\<Sum>Basis}"
 | 
| 53688 | 3939 | unfolding z_def mem_interval | 
| 3940 | using b'_Basis q(1)[rule_format,OF b'_im] `p > 0` | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3941 | by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le) | 
| 53688 | 3942 | have *: "\<And>x. 1 + real x = real (Suc x)" | 
| 3943 | by auto | |
| 3944 |   {
 | |
| 3945 | have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)" | |
| 53185 | 3946 | apply (rule setsum_mono) | 
| 53252 | 3947 | using rs(1)[OF b'_im] | 
| 3948 | apply (auto simp add:* field_simps) | |
| 53185 | 3949 | done | 
| 53688 | 3950 | also have "\<dots> < e * real p" | 
| 3951 | using p `e > 0` `p > 0` | |
| 53185 | 3952 | by (auto simp add: field_simps n_def real_of_nat_def) | 
| 3953 | finally have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) < e * real p" . | |
| 3954 | } | |
| 3955 | moreover | |
| 53688 | 3956 |   {
 | 
| 3957 | have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)" | |
| 53185 | 3958 | apply (rule setsum_mono) | 
| 53252 | 3959 | using rs(2)[OF b'_im] | 
| 3960 | apply (auto simp add:* field_simps) | |
| 53185 | 3961 | done | 
| 53688 | 3962 | also have "\<dots> < e * real p" | 
| 3963 | using p `e > 0` `p > 0` | |
| 53185 | 3964 | by (auto simp add: field_simps n_def real_of_nat_def) | 
| 3965 | finally have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) < e * real p" . | |
| 3966 | } | |
| 3967 | ultimately | |
| 53688 | 3968 | have "norm (r' - z) < e" and "norm (s' - z) < e" | 
| 53185 | 3969 | unfolding r'_def s'_def z_def | 
| 53688 | 3970 | using `p > 0` | 
| 53185 | 3971 | apply (rule_tac[!] le_less_trans[OF norm_le_l1]) | 
| 3972 | apply (auto simp add: field_simps setsum_divide_distrib[symmetric] inner_diff_left) | |
| 3973 | done | |
| 53674 | 3974 | then have "\<bar>(f z - z) \<bullet> i\<bar> < d / real n" | 
| 53688 | 3975 | using rs(3) i | 
| 3976 | unfolding r'_def[symmetric] s'_def[symmetric] o_def bb' | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 3977 |     by (intro e(2)[OF `r'\<in>{0..\<Sum>Basis}` `s'\<in>{0..\<Sum>Basis}` `z\<in>{0..\<Sum>Basis}`]) auto
 | 
| 53688 | 3978 | then show False | 
| 3979 | 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 | 3980 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3981 | |
| 53185 | 3982 | |
| 53688 | 3983 | subsection {* Retractions *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3984 | |
| 53688 | 3985 | 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 | 3986 | |
| 53185 | 3987 | definition retract_of (infixl "retract'_of" 12) | 
| 3988 | 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 | 3989 | |
| 53674 | 3990 | 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 | 3991 | 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 | 3992 | |
| 53688 | 3993 | subsection {* Preservation of fixpoints under (more general notion of) retraction *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 3994 | |
| 53185 | 3995 | lemma invertible_fixpoint_property: | 
| 53674 | 3996 | fixes s :: "'a::euclidean_space set" | 
| 3997 | and t :: "'b::euclidean_space set" | |
| 3998 | assumes "continuous_on t i" | |
| 3999 | and "i ` t \<subseteq> s" | |
| 53688 | 4000 | and "continuous_on s r" | 
| 4001 | and "r ` s \<subseteq> t" | |
| 53674 | 4002 | and "\<forall>y\<in>t. r (i y) = y" | 
| 4003 | and "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)" | |
| 4004 | and "continuous_on t g" | |
| 4005 | and "g ` t \<subseteq> t" | |
| 4006 | obtains y where "y \<in> t" and "g y = y" | |
| 53185 | 4007 | proof - | 
| 4008 | have "\<exists>x\<in>s. (i \<circ> g \<circ> r) x = x" | |
| 53688 | 4009 | apply (rule assms(6)[rule_format]) | 
| 4010 | apply rule | |
| 53185 | 4011 | apply (rule continuous_on_compose assms)+ | 
| 53688 | 4012 | apply ((rule continuous_on_subset)?, rule assms)+ | 
| 4013 | using assms(2,4,8) | |
| 4014 | unfolding image_compose | |
| 53185 | 4015 | apply auto | 
| 4016 | apply blast | |
| 4017 | done | |
| 4018 | then guess x .. note x = this | |
| 53674 | 4019 | then have *: "g (r x) \<in> t" | 
| 4020 | using assms(4,8) by auto | |
| 4021 | have "r ((i \<circ> g \<circ> r) x) = r x" | |
| 4022 | using x by auto | |
| 4023 | then show ?thesis | |
| 53185 | 4024 | apply (rule_tac that[of "r x"]) | 
| 53674 | 4025 | using x | 
| 4026 | unfolding o_def | |
| 4027 | unfolding assms(5)[rule_format,OF *] | |
| 4028 | using assms(4) | |
| 53185 | 4029 | apply auto | 
| 4030 | done | |
| 4031 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4032 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4033 | lemma homeomorphic_fixpoint_property: | 
| 53674 | 4034 | fixes s :: "'a::euclidean_space set" | 
| 4035 | and t :: "'b::euclidean_space set" | |
| 53185 | 4036 | 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 | 4037 | shows "(\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)) \<longleftrightarrow> | 
| 53248 | 4038 | (\<forall>g. continuous_on t g \<and> g ` t \<subseteq> t \<longrightarrow> (\<exists>y\<in>t. g y = y))" | 
| 53185 | 4039 | proof - | 
| 4040 | guess r using assms[unfolded homeomorphic_def homeomorphism_def] .. | |
| 4041 | then guess i .. | |
| 53674 | 4042 | then show ?thesis | 
| 53185 | 4043 | apply - | 
| 4044 | apply rule | |
| 4045 | apply (rule_tac[!] allI impI)+ | |
| 4046 | apply (rule_tac g=g in invertible_fixpoint_property[of t i s r]) | |
| 4047 | prefer 10 | |
| 4048 | apply (rule_tac g=f in invertible_fixpoint_property[of s r t i]) | |
| 4049 | apply auto | |
| 4050 | done | |
| 4051 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4052 | |
| 53185 | 4053 | lemma retract_fixpoint_property: | 
| 53688 | 4054 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 53674 | 4055 | and s :: "'a set" | 
| 53185 | 4056 | assumes "t retract_of s" | 
| 53674 | 4057 | and "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)" | 
| 4058 | and "continuous_on t g" | |
| 4059 | and "g ` t \<subseteq> t" | |
| 4060 | obtains y where "y \<in> t" and "g y = y" | |
| 53185 | 4061 | proof - | 
| 4062 | guess h using assms(1) unfolding retract_of_def .. | |
| 53674 | 4063 | then show ?thesis | 
| 53185 | 4064 | unfolding retraction_def | 
| 4065 | apply - | |
| 4066 | apply (rule invertible_fixpoint_property[OF continuous_on_id _ _ _ _ assms(2), of t h g]) | |
| 4067 | prefer 7 | |
| 53248 | 4068 | apply (rule_tac y = y in that) | 
| 4069 | using assms | |
| 4070 | apply auto | |
| 53185 | 4071 | done | 
| 4072 | qed | |
| 4073 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4074 | |
| 53688 | 4075 | subsection {* The Brouwer theorem for any set with nonempty interior *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4076 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4077 | lemma brouwer_weak: | 
| 53248 | 4078 | fixes f :: "'a::ordered_euclidean_space \<Rightarrow> 'a::ordered_euclidean_space" | 
| 53674 | 4079 | assumes "compact s" | 
| 4080 | and "convex s" | |
| 4081 |     and "interior s \<noteq> {}"
 | |
| 4082 | and "continuous_on s f" | |
| 4083 | and "f ` s \<subseteq> s" | |
| 4084 | obtains x where "x \<in> s" and "f x = x" | |
| 53185 | 4085 | proof - | 
| 4086 |   have *: "interior {0::'a..\<Sum>Basis} \<noteq> {}"
 | |
| 53674 | 4087 | unfolding interior_closed_interval interval_eq_empty | 
| 4088 | by auto | |
| 53185 | 4089 |   have *: "{0::'a..\<Sum>Basis} homeomorphic s"
 | 
| 4090 | using homeomorphic_convex_compact[OF convex_interval(1) compact_interval * assms(2,1,3)] . | |
| 4091 |   have "\<forall>f. continuous_on {0::'a..\<Sum>Basis} f \<and> f ` {0::'a..\<Sum>Basis} \<subseteq> {0::'a..\<Sum>Basis} \<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 | 4092 |     (\<exists>x\<in>{0::'a..\<Sum>Basis}. 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 | 4093 | using brouwer_cube by auto | 
| 53674 | 4094 | then show ?thesis | 
| 53185 | 4095 | unfolding homeomorphic_fixpoint_property[OF *] | 
| 4096 | apply (erule_tac x=f in allE) | |
| 4097 | apply (erule impE) | |
| 4098 | defer | |
| 4099 | apply (erule bexE) | |
| 4100 | apply (rule_tac x=y in that) | |
| 53252 | 4101 | using assms | 
| 4102 | apply auto | |
| 53185 | 4103 | done | 
| 4104 | qed | |
| 4105 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4106 | |
| 53688 | 4107 | text {* And in particular for a closed ball. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4108 | |
| 53185 | 4109 | lemma brouwer_ball: | 
| 4110 | fixes f :: "'a::ordered_euclidean_space \<Rightarrow> 'a" | |
| 53674 | 4111 | assumes "e > 0" | 
| 4112 | and "continuous_on (cball a e) f" | |
| 53688 | 4113 | and "f ` cball a e \<subseteq> cball a e" | 
| 53674 | 4114 | obtains x where "x \<in> cball a e" and "f x = x" | 
| 53185 | 4115 | using brouwer_weak[OF compact_cball convex_cball, of a e f] | 
| 4116 | 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 | 4117 | 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 | 4118 | |
| 53185 | 4119 | text {*Still more general form; could derive this directly without using the
 | 
| 36334 | 4120 |   rather involved @{text "HOMEOMORPHIC_CONVEX_COMPACT"} theorem, just using
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4121 | a scaling and translation to put the set inside the unit cube. *} | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4122 | |
| 53248 | 4123 | lemma brouwer: | 
| 53674 | 4124 | fixes f :: "'a::ordered_euclidean_space \<Rightarrow> 'a" | 
| 4125 | assumes "compact s" | |
| 4126 | and "convex s" | |
| 4127 |     and "s \<noteq> {}"
 | |
| 4128 | and "continuous_on s f" | |
| 4129 | and "f ` s \<subseteq> s" | |
| 4130 | obtains x where "x \<in> s" and "f x = x" | |
| 53185 | 4131 | proof - | 
| 4132 | have "\<exists>e>0. s \<subseteq> cball 0 e" | |
| 53688 | 4133 | using compact_imp_bounded[OF assms(1)] | 
| 4134 | unfolding bounded_pos | |
| 53674 | 4135 | apply (erule_tac exE) | 
| 4136 | apply (rule_tac x=b in exI) | |
| 53185 | 4137 | apply (auto simp add: dist_norm) | 
| 4138 | done | |
| 4139 | then guess e by (elim exE conjE) | |
| 4140 | note e = this | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4141 | have "\<exists>x\<in> cball 0 e. (f \<circ> closest_point s) x = x" | 
| 53185 | 4142 | apply (rule_tac brouwer_ball[OF e(1), of 0 "f \<circ> closest_point s"]) | 
| 4143 | apply (rule continuous_on_compose ) | |
| 4144 | apply (rule continuous_on_closest_point[OF assms(2) compact_imp_closed[OF assms(1)] assms(3)]) | |
| 4145 | apply (rule continuous_on_subset[OF assms(4)]) | |
| 4146 | apply (insert closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)]) | |
| 4147 | defer | |
| 4148 | using assms(5)[unfolded subset_eq] | |
| 4149 | using e(2)[unfolded subset_eq mem_cball] | |
| 4150 | apply (auto simp add: dist_norm) | |
| 4151 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4152 | then guess x .. note x=this | 
| 53185 | 4153 | have *: "closest_point s x = x" | 
| 4154 | apply (rule closest_point_self) | |
| 4155 | apply (rule assms(5)[unfolded subset_eq,THEN bspec[where x="x"], unfolded image_iff]) | |
| 4156 | apply (rule_tac x="closest_point s x" in bexI) | |
| 4157 | using x | |
| 4158 | unfolding o_def | |
| 4159 | using closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3), of x] | |
| 4160 | apply auto | |
| 4161 | done | |
| 4162 | show thesis | |
| 4163 | apply (rule_tac x="closest_point s x" in that) | |
| 4164 | unfolding x(2)[unfolded o_def] | |
| 4165 | apply (rule closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)]) | |
| 53674 | 4166 | using * | 
| 4167 | apply auto | |
| 4168 | done | |
| 53185 | 4169 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4170 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36587diff
changeset | 4171 | text {*So we get the no-retraction theorem. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4172 | |
| 53185 | 4173 | lemma no_retraction_cball: | 
| 53674 | 4174 | fixes a :: "'a::ordered_euclidean_space" | 
| 4175 | assumes "e > 0" | |
| 4176 | shows "\<not> (frontier (cball a e) retract_of (cball a e))" | |
| 53185 | 4177 | proof | 
| 4178 | case goal1 | |
| 53674 | 4179 | have *: "\<And>xa. a - (2 *\<^sub>R a - xa) = - (a - xa)" | 
| 53185 | 4180 | using scaleR_left_distrib[of 1 1 a] by auto | 
| 4181 | guess x | |
| 4182 | apply (rule retract_fixpoint_property[OF goal1, of "\<lambda>x. scaleR 2 a - x"]) | |
| 53674 | 4183 | apply rule | 
| 4184 | apply rule | |
| 4185 | apply (erule conjE) | |
| 53185 | 4186 | apply (rule brouwer_ball[OF assms]) | 
| 4187 | apply assumption+ | |
| 4188 | apply (rule_tac x=x in bexI) | |
| 4189 | apply assumption+ | |
| 4190 | apply (rule continuous_on_intros)+ | |
| 4191 | unfolding frontier_cball subset_eq Ball_def image_iff | |
| 53674 | 4192 | apply rule | 
| 4193 | apply rule | |
| 4194 | apply (erule bexE) | |
| 53185 | 4195 | unfolding dist_norm | 
| 4196 | apply (simp add: * norm_minus_commute) | |
| 4197 | done | |
| 4198 | note x = this | |
| 53674 | 4199 | then have "scaleR 2 a = scaleR 1 x + scaleR 1 x" | 
| 53248 | 4200 | by (auto simp add: algebra_simps) | 
| 53674 | 4201 | then have "a = x" | 
| 53688 | 4202 | unfolding scaleR_left_distrib[symmetric] | 
| 4203 | by auto | |
| 53674 | 4204 | then show False | 
| 4205 | using x assms by auto | |
| 53185 | 4206 | qed | 
| 4207 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4208 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4209 | subsection {*Bijections between intervals. *}
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4210 | |
| 53248 | 4211 | definition interval_bij :: "'a \<times> 'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<Rightarrow> 'a::ordered_euclidean_space" | 
| 4212 | where "interval_bij = | |
| 4213 | (\<lambda>(a, b) (u, v) x. (\<Sum>i\<in>Basis. (u\<bullet>i + (x\<bullet>i - a\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (v\<bullet>i - u\<bullet>i)) *\<^sub>R i))" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4214 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4215 | lemma interval_bij_affine: | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4216 | "interval_bij (a,b) (u,v) = (\<lambda>x. (\<Sum>i\<in>Basis. ((v\<bullet>i - u\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (x\<bullet>i)) *\<^sub>R i) + | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4217 | (\<Sum>i\<in>Basis. (u\<bullet>i - (v\<bullet>i - u\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (a\<bullet>i)) *\<^sub>R i))" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4218 | by (auto simp: setsum_addf[symmetric] scaleR_add_left[symmetric] interval_bij_def fun_eq_iff | 
| 53248 | 4219 | field_simps inner_simps add_divide_distrib[symmetric] intro!: setsum_cong) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4220 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4221 | lemma continuous_interval_bij: | 
| 53674 | 4222 | fixes a b :: "'a::ordered_euclidean_space" | 
| 4223 | shows "continuous (at x) (interval_bij (a, b) (u, v))" | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4224 | by (auto simp add: divide_inverse interval_bij_def intro!: continuous_setsum continuous_intros) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4225 | |
| 53674 | 4226 | lemma continuous_on_interval_bij: "continuous_on s (interval_bij (a, b) (u, v))" | 
| 53185 | 4227 | apply(rule continuous_at_imp_continuous_on) | 
| 4228 | apply (rule, rule continuous_interval_bij) | |
| 4229 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4230 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4231 | lemma in_interval_interval_bij: | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4232 | fixes a b u v x :: "'a::ordered_euclidean_space" | 
| 53674 | 4233 |   assumes "x \<in> {a..b}"
 | 
| 4234 |     and "{u..v} \<noteq> {}"
 | |
| 53688 | 4235 |   shows "interval_bij (a, b) (u, v) x \<in> {u..v}"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4236 | apply (simp only: interval_bij_def split_conv mem_interval inner_setsum_left_Basis cong: ball_cong) | 
| 53248 | 4237 | apply safe | 
| 4238 | proof - | |
| 53185 | 4239 | fix i :: 'a | 
| 4240 | assume i: "i \<in> Basis" | |
| 53674 | 4241 |   have "{a..b} \<noteq> {}"
 | 
| 4242 | using assms 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 | 4243 | with i have *: "a\<bullet>i \<le> b\<bullet>i" "u\<bullet>i \<le> v\<bullet>i" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4244 | using assms(2) by (auto simp add: interval_eq_empty not_less) | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4245 | have x: "a\<bullet>i\<le>x\<bullet>i" "x\<bullet>i\<le>b\<bullet>i" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4246 | using assms(1)[unfolded mem_interval] using i by auto | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4247 | have "0 \<le> (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i)" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4248 | using * x by (auto intro!: mult_nonneg_nonneg divide_nonneg_nonneg) | 
| 53674 | 4249 | then show "u \<bullet> i \<le> u \<bullet> i + (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i)" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4250 | using * by auto | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4251 | have "((x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (v \<bullet> i - u \<bullet> i) \<le> 1 * (v \<bullet> i - u \<bullet> i)" | 
| 53185 | 4252 | apply (rule mult_right_mono) | 
| 4253 | unfolding divide_le_eq_1 | |
| 53252 | 4254 | using * x | 
| 4255 | apply auto | |
| 53185 | 4256 | done | 
| 53674 | 4257 | then show "u \<bullet> i + (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i) \<le> v \<bullet> i" | 
| 53185 | 4258 | using * by auto | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36587diff
changeset | 4259 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4260 | |
| 53185 | 4261 | lemma interval_bij_bij: | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 4262 | "\<forall>(i::'a::ordered_euclidean_space)\<in>Basis. a\<bullet>i < b\<bullet>i \<and> u\<bullet>i < v\<bullet>i \<Longrightarrow> | 
| 53846 | 4263 | interval_bij (a, b) (u, v) (interval_bij (u, v) (a, b) x) = 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 | 4264 | by (auto simp: interval_bij_def euclidean_eq_iff[where 'a='a]) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4265 | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 4266 | end |