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