| author | haftmann | 
| Wed, 24 Feb 2010 14:19:52 +0100 | |
| changeset 35366 | 6d474096698c | 
| parent 34964 | 4e8be3c04d37 | 
| child 35729 | 3cd1e4b65111 | 
| permissions | -rw-r--r-- | 
| 
33741
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
2  | 
(* ========================================================================= *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
3  | 
(* 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
 | 
4  | 
(* *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
5  | 
(* At the moment this is just Brouwer's fixpoint theorem. The proof is from *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
6  | 
(* Kuhn: "some combinatorial lemmas in topology", IBM J. v4. (1960) p. 518 *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
7  | 
(* See "http://www.research.ibm.com/journal/rd/045/ibmrd0405K.pdf". *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
8  | 
(* *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
9  | 
(* The script below is quite messy, but at least we avoid formalizing any *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
10  | 
(* topological machinery; we don't even use barycentric subdivision; this is *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
11  | 
(* the big advantage of Kuhn's proof over the usual Sperner's lemma one. *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
12  | 
(* *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
13  | 
(* (c) Copyright, John Harrison 1998-2008 *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
14  | 
(* ========================================================================= *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
15  | 
|
| 
33759
 
b369324fc244
Added the contributions of Robert Himmelmann to CONTRIBUTIONS and NEWS
 
hoelzl 
parents: 
33758 
diff
changeset
 | 
16  | 
(* Author: John Harrison  | 
| 
 
b369324fc244
Added the contributions of Robert Himmelmann to CONTRIBUTIONS and NEWS
 
hoelzl 
parents: 
33758 
diff
changeset
 | 
17  | 
Translation from HOL light: Robert Himmelmann, TU Muenchen *)  | 
| 
33741
 
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  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
22  | 
imports Convex_Euclidean_Space  | 
| 
 
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  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
25  | 
declare norm_scaleR[simp]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
26  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
27  | 
lemma brouwer_compactness_lemma:  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
28  | 
assumes "compact s" "continuous_on s f" "\<not> (\<exists>x\<in>s. (f x = (0::real^'n)))"  | 
| 
33741
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
29  | 
  obtains d where "0 < d" "\<forall>x\<in>s. d \<le> norm(f x)" proof(cases "s={}") case False
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
30  | 
have "continuous_on s (norm \<circ> f)" by(rule continuous_on_intros continuous_on_norm assms(2))+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
31  | 
then obtain x where x:"x\<in>s" "\<forall>y\<in>s. (norm \<circ> f) x \<le> (norm \<circ> f) y"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
32  | 
using continuous_attains_inf[OF assms(1), of "norm \<circ> f"] and False unfolding o_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
 | 
33  | 
have "(norm \<circ> f) x > 0" using assms(3) and x(1) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
34  | 
thus ?thesis apply(rule that) using x(2) unfolding o_def by auto qed(rule that[of 1], auto)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
35  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
36  | 
lemma kuhn_labelling_lemma:  | 
| 34289 | 37  | 
assumes "(\<forall>x::real^_. P x \<longrightarrow> P (f x))" "\<forall>x. P x \<longrightarrow> (\<forall>i. 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
 | 
38  | 
shows "\<exists>l. (\<forall>x i. l x i \<le> (1::nat)) \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
39  | 
(\<forall>x i. P x \<and> Q i \<and> (x$i = 0) \<longrightarrow> (l x i = 0)) \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
40  | 
(\<forall>x i. P x \<and> Q i \<and> (x$i = 1) \<longrightarrow> (l x i = 1)) \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
41  | 
(\<forall>x i. P x \<and> Q i \<and> (l x i = 0) \<longrightarrow> x$i \<le> f(x)$i) \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
42  | 
(\<forall>x i. P x \<and> Q i \<and> (l x i = 1) \<longrightarrow> f(x)$i \<le> x$i)" proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
43  | 
have and_forall_thm:"\<And>P Q. (\<forall>x. P x) \<and> (\<forall>x. Q x) \<longleftrightarrow> (\<forall>x. P x \<and> Q x)" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
44  | 
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)" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
45  | 
show ?thesis unfolding and_forall_thm apply(subst choice_iff[THEN sym])+ proof(rule,rule) case goal1  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
46  | 
let ?R = "\<lambda>y. (P x \<and> Q xa \<and> x $ xa = 0 \<longrightarrow> y = (0::nat)) \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
47  | 
(P x \<and> Q xa \<and> x $ xa = 1 \<longrightarrow> y = 1) \<and> (P x \<and> Q xa \<and> y = 0 \<longrightarrow> x $ xa \<le> f x $ xa) \<and> (P x \<and> Q xa \<and> y = 1 \<longrightarrow> f x $ xa \<le> x $ xa)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
48  | 
    { assume "P x" "Q xa" hence "0 \<le> f x $ xa \<and> f x $ xa \<le> 1" using assms(2)[rule_format,of "f x" xa]
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
49  | 
apply(drule_tac assms(1)[rule_format]) by auto }  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
50  | 
hence "?R 0 \<or> ?R 1" by auto thus ?case by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
52  | 
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
 | 
53  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
54  | 
lemma setsum_Un_disjoint':assumes "finite A" "finite B" "A \<inter> B = {}" "A \<union> B = C"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
55  | 
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
 | 
56  | 
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
 | 
57  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
58  | 
lemma kuhn_counting_lemma: assumes "finite faces" "finite simplices"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
59  | 
  "\<forall>f\<in>faces. bnd f  \<longrightarrow> (card {s \<in> simplices. face f s} = 1)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
60  | 
  "\<forall>f\<in>faces. \<not> bnd f \<longrightarrow> (card {s \<in> simplices. face f s} = 2)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
61  | 
  "\<forall>s\<in>simplices. compo s  \<longrightarrow> (card {f \<in> faces. face f s \<and> compo' f} = 1)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
62  | 
  "\<forall>s\<in>simplices. \<not> compo s \<longrightarrow> (card {f \<in> faces. face f s \<and> compo' f} = 0) \<or>
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
63  | 
                             (card {f \<in> faces. face f s \<and> compo' f} = 2)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
64  | 
    "odd(card {f \<in> faces. compo' f \<and> bnd f})"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
65  | 
  shows "odd(card {s \<in> simplices. compo s})" proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
66  | 
  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} = {f\<in>faces. compo' f \<and> face f x}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
67  | 
    "\<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} = {}" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
68  | 
  hence lem1:"setsum (\<lambda>s. (card {f \<in> faces. face f s \<and> compo' f})) simplices =
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
69  | 
    setsum (\<lambda>s. card {f \<in> {f \<in> faces. compo' f \<and> bnd f}. face f s}) simplices +
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
70  | 
    setsum (\<lambda>s. card {f \<in> {f \<in> faces. compo' f \<and> \<not> (bnd f)}. face f s}) simplices"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
71  | 
unfolding setsum_addf[THEN sym] apply- apply(rule setsum_cong2)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
72  | 
using assms(1) by(auto simp add: card_Un_Int, auto simp add:conj_commute)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
73  | 
  have lem2:"setsum (\<lambda>j. card {f \<in> {f \<in> faces. compo' f \<and> bnd f}. face f j}) simplices = 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
74  | 
              1 * card {f \<in> faces. compo' f \<and> bnd f}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
75  | 
       "setsum (\<lambda>j. card {f \<in> {f \<in> faces. compo' f \<and> \<not> bnd f}. face f j}) simplices = 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
76  | 
              2 * card {f \<in> faces. compo' f \<and> \<not> bnd f}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
77  | 
apply(rule_tac[!] setsum_multicount) 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
 | 
78  | 
  have lem3:"setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f}) simplices =
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
79  | 
    setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f}) {s \<in> simplices.   compo s}+
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
80  | 
    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
 | 
81  | 
apply(rule setsum_Un_disjoint') using assms(2) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
82  | 
  have lem4:"setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f}) {s \<in> simplices. compo s}
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
83  | 
    = setsum (\<lambda>s. 1) {s \<in> simplices. compo s}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
84  | 
apply(rule setsum_cong2) using assms(5) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
85  | 
  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
 | 
86  | 
    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
 | 
87  | 
           {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
 | 
88  | 
    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
 | 
89  | 
           {s \<in> simplices. (\<not> compo s) \<and> (card {f \<in> faces. face f s \<and> compo' f} = 2)}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
90  | 
apply(rule setsum_Un_disjoint') using assms(2,6) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
91  | 
  have *:"int (\<Sum>s\<in>{s \<in> simplices. compo 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
 | 
92  | 
    int (card {f \<in> faces. compo' f \<and> bnd f} + 2 * card {f \<in> faces. compo' f \<and> \<not> bnd f}) - 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
93  | 
    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
 | 
94  | 
using lem1[unfolded lem3 lem2 lem5] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
95  | 
have even_minus_odd:"\<And>x y. even x \<Longrightarrow> odd (y::int) \<Longrightarrow> odd (x - y)" 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
 | 
96  | 
have odd_minus_even:"\<And>x y. odd x \<Longrightarrow> even (y::int) \<Longrightarrow> odd (x - y)" 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
 | 
97  | 
show ?thesis unfolding even_nat_def unfolding card_def and lem4[THEN sym] and *[unfolded card_def]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
98  | 
unfolding card_def[THEN sym] apply(rule odd_minus_even) unfolding zadd_int[THEN sym] apply(rule odd_plus_even)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
99  | 
apply(rule assms(7)[unfolded even_nat_def]) unfolding int_mult by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
100  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
101  | 
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
 | 
102  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
103  | 
lemma card_1_exists: "card s = 1 \<longleftrightarrow> (\<exists>!x. x \<in> s)" unfolding One_nat_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
104  | 
apply rule apply(drule card_eq_SucD) defer apply(erule ex1E) proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
105  | 
fix x assume as:"x \<in> s" "\<forall>y. y \<in> s \<longrightarrow> y = x"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
106  | 
  have *:"s = insert x {}" apply- apply(rule set_ext,rule) unfolding singleton_iff
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
107  | 
apply(rule as(2)[rule_format]) using as(1) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
108  | 
show "card s = Suc 0" unfolding * using card_insert by auto qed auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
109  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
110  | 
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)))" proof  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
111  | 
  assume "card s = 2" then obtain x y where obt:"s = {x, y}" "x\<noteq>y" unfolding numeral_2_eq_2 apply - apply(erule exE conjE|drule card_eq_SucD)+ by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
112  | 
show "\<exists>x\<in>s. \<exists>y\<in>s. x \<noteq> y \<and> (\<forall>z\<in>s. z = x \<or> z = y)" using obt by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
113  | 
assume "\<exists>x\<in>s. \<exists>y\<in>s. x \<noteq> y \<and> (\<forall>z\<in>s. z = x \<or> z = y)" then guess x .. from this(2) guess y ..  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
114  | 
  with `x\<in>s` have *:"s = {x, y}" "x\<noteq>y" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
115  | 
from this(2) show "card s = 2" unfolding * by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
116  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
117  | 
lemma image_lemma_0: assumes "card {a\<in>s. f ` (s - {a}) = t - {b}} = n"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
118  | 
  shows "card {s'. \<exists>a\<in>s. (s' = s - {a}) \<and> (f ` s' = t - {b})} = n" proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
119  | 
  have *:"{s'. \<exists>a\<in>s. (s' = s - {a}) \<and> (f ` s' = t - {b})} = (\<lambda>a. s - {a}) ` {a\<in>s. f ` (s - {a}) = t - {b}}" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
120  | 
show ?thesis unfolding * unfolding assms[THEN sym] apply(rule card_image) unfolding inj_on_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
121  | 
apply(rule,rule,rule) unfolding mem_Collect_eq by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
122  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
123  | 
lemma image_lemma_1: assumes "finite s" "finite t" "card s = card t" "f ` s = t" "b \<in> t"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
124  | 
  shows "card {s'. \<exists>a\<in>s. s' = s - {a} \<and>  f ` s' = t - {b}} = 1" proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
125  | 
obtain a where a:"b = f a" "a\<in>s" using assms(4-5) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
126  | 
have inj:"inj_on f s" apply(rule eq_card_imp_inj_on) using assms(1-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
 | 
127  | 
  have *:"{a \<in> s. f ` (s - {a}) = t - {b}} = {a}" apply(rule set_ext) unfolding singleton_iff
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
128  | 
apply(rule,rule inj[unfolded inj_on_def,rule_format]) unfolding a using a(2) and assms and inj[unfolded inj_on_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
 | 
129  | 
show ?thesis apply(rule image_lemma_0) unfolding * by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
130  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
131  | 
lemma image_lemma_2: assumes "finite s" "finite t" "card s = card t" "f ` s \<subseteq> t" "f ` s \<noteq> t" "b \<in> t"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
132  | 
  shows "(card {s'. \<exists>a\<in>s. (s' = s - {a}) \<and> f ` s' = t - {b}} = 0) \<or>
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
133  | 
         (card {s'. \<exists>a\<in>s. (s' = s - {a}) \<and> f ` s' = t - {b}} = 2)" proof(cases "{a\<in>s. f ` (s - {a}) = t - {b}} = {}")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
134  | 
case True thus ?thesis apply-apply(rule disjI1, rule image_lemma_0) using assms(1) by(auto simp add:card_0_eq)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
135  | 
next let ?M = "{a\<in>s. f ` (s - {a}) = t - {b}}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
136  | 
  case False then obtain a where "a\<in>?M" by auto hence a:"a\<in>s" "f ` (s - {a}) = t - {b}" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
137  | 
  have "f a \<in> t - {b}" using a and 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
 | 
138  | 
  hence "\<exists>c \<in> s - {a}. f a = f c" unfolding image_iff[symmetric] and a by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
139  | 
then obtain c where c:"c \<in> s" "a \<noteq> c" "f a = f c" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
140  | 
  hence *:"f ` (s - {c}) = f ` (s - {a})" apply-apply(rule set_ext,rule) proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
141  | 
    fix x assume "x \<in> f ` (s - {a})" then obtain y where y:"f y = x" "y\<in>s- {a}" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
142  | 
    thus "x \<in> f ` (s - {c})" unfolding image_iff apply(rule_tac x="if y = c then a else y" in bexI) using c a by auto qed auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
143  | 
have "c\<in>?M" unfolding mem_Collect_eq and * using a and c(1) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
144  | 
show ?thesis apply(rule disjI2, rule image_lemma_0) unfolding card_2_exists  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
145  | 
apply(rule bexI[OF _ `a\<in>?M`], rule bexI[OF _ `c\<in>?M`],rule,rule `a\<noteq>c`) proof(rule,unfold mem_Collect_eq,erule conjE)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
146  | 
    fix z assume as:"z \<in> s" "f ` (s - {z}) = t - {b}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
147  | 
    have inj:"inj_on f (s - {z})" apply(rule eq_card_imp_inj_on) unfolding as using as(1) and 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
 | 
148  | 
show "z = a \<or> z = c" proof(rule ccontr)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
149  | 
assume "\<not> (z = a \<or> z = c)" thus False using inj[unfolded inj_on_def,THEN bspec[where x=a],THEN bspec[where x=c]]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
150  | 
using `a\<in>s` `c\<in>s` `f a = f c` `a\<noteq>c` by auto qed qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
151  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
152  | 
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
 | 
153  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
154  | 
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
 | 
155  | 
assumes "finite simplices"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
156  | 
  "\<forall>f s. face f s \<longleftrightarrow> (\<exists>a\<in>s. f = s - {a})" "\<forall>s\<in>simplices. card s = n + 2" "\<forall>s\<in>simplices. (rl ` s) \<subseteq> {0..n+1}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
157  | 
  "\<forall>f\<in> {f. \<exists>s\<in>simplices. face f s}. bnd f  \<longrightarrow> (card {s\<in>simplices. face f s} = 1)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
158  | 
  "\<forall>f\<in> {f. \<exists>s\<in>simplices. face f s}. \<not>bnd f \<longrightarrow> (card {s\<in>simplices. face f s} = 2)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
159  | 
  "odd(card {f\<in>{f. \<exists>s\<in>simplices. face f s}. rl ` f = {0..n} \<and> bnd f})"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
160  | 
  shows "odd (card {s\<in>simplices. (rl ` s = {0..n+1})})" 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
161  | 
apply(rule kuhn_counting_lemma) defer apply(rule assms)+ prefer 3 apply(rule assms) proof(rule_tac[1-2] ballI impI)+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
162  | 
  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})})" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
163  | 
have **: "\<forall>s\<in>simplices. card s = n + 2 \<and> finite s" using assms(3) 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
 | 
164  | 
  show "finite {f. \<exists>s\<in>simplices. face f s}" unfolding assms(2)[rule_format] and *
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
165  | 
apply(rule finite_UN_I[OF assms(1)]) using ** by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
166  | 
  have *:"\<And> P f s. s\<in>simplices \<Longrightarrow> (f \<in> {f. \<exists>s\<in>simplices. \<exists>a\<in>s. f = s - {a}}) \<and>
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
167  | 
    (\<exists>a\<in>s. (f = s - {a})) \<and> P f \<longleftrightarrow> (\<exists>a\<in>s. (f = s - {a}) \<and> P f)" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
168  | 
  fix s assume s:"s\<in>simplices" let ?S = "{f \<in> {f. \<exists>s\<in>simplices. face f s}. face f s \<and> rl ` f = {0..n}}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
169  | 
    have "{0..n + 1} - {n + 1} = {0..n}" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
170  | 
    hence S:"?S = {s'. \<exists>a\<in>s. s' = s - {a} \<and> rl ` s' = {0..n + 1} - {n + 1}}" apply- apply(rule set_ext)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
171  | 
      unfolding assms(2)[rule_format] mem_Collect_eq and *[OF s, unfolded mem_Collect_eq, where P="\<lambda>x. rl ` x = {0..n}"] by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
172  | 
    show "rl ` s = {0..n+1} \<Longrightarrow> card ?S = 1" "rl ` s \<noteq> {0..n+1} \<Longrightarrow> card ?S = 0 \<or> card ?S = 2" unfolding S
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
173  | 
apply(rule_tac[!] image_lemma_1 image_lemma_2) using ** assms(4) and s by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
174  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
175  | 
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
 | 
176  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
177  | 
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)))"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
178  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
179  | 
lemma kle_refl[intro]: "kle n x x" unfolding kle_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
 | 
180  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
181  | 
lemma kle_antisym: "kle n x y \<and> kle n y x \<longleftrightarrow> (x = y)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
182  | 
unfolding kle_def apply rule apply(rule ext) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
183  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
184  | 
lemma pointwise_minimal_pointwise_maximal: fixes s::"(nat\<Rightarrow>nat) set"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
185  | 
  assumes  "finite s" "s \<noteq> {}" "\<forall>x\<in>s. \<forall>y\<in>s. (\<forall>j. x j \<le> y j) \<or> (\<forall>j. y j \<le> x j)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
186  | 
shows "\<exists>a\<in>s. \<forall>x\<in>s. \<forall>j. a j \<le> x j" "\<exists>a\<in>s. \<forall>x\<in>s. \<forall>j. x j \<le> a j"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
187  | 
using assms unfolding atomize_conj apply- proof(induct s rule:finite_induct)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
188  | 
fix x and F::"(nat\<Rightarrow>nat) set" assume as:"finite F" "x \<notin> F"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
189  | 
    "\<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
 | 
190  | 
        \<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
 | 
191  | 
"\<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)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
192  | 
  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)" proof(cases "F={}")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
193  | 
case True thus ?thesis apply-apply(rule,rule_tac[!] x=x in bexI) by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
194  | 
case False obtain a b where a:"a\<in>insert x F" "\<forall>x\<in>F. \<forall>j. a j \<le> x j" and  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
195  | 
b:"b\<in>insert x F" "\<forall>x\<in>F. \<forall>j. x j \<le> b j" using as(3)[OF False] using as(5) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
196  | 
have "\<exists>a\<in>insert x F. \<forall>x\<in>insert x F. \<forall>j. a j \<le> x j"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
197  | 
using as(5)[rule_format,OF a(1) insertI1] apply- proof(erule disjE)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
198  | 
assume "\<forall>j. a j \<le> x j" thus ?thesis apply(rule_tac x=a in bexI) using a by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
199  | 
assume "\<forall>j. x j \<le> a j" thus ?thesis apply(rule_tac x=x in bexI) apply(rule,rule) using a apply -  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
200  | 
apply(erule_tac x=xa in ballE) apply(erule_tac x=j in allE)+ by auto qed moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
201  | 
have "\<exists>b\<in>insert x F. \<forall>x\<in>insert x F. \<forall>j. x j \<le> b j"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
202  | 
using as(5)[rule_format,OF b(1) insertI1] apply- proof(erule disjE)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
203  | 
assume "\<forall>j. x j \<le> b j" thus ?thesis apply(rule_tac x=b in bexI) using b by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
204  | 
assume "\<forall>j. b j \<le> x j" thus ?thesis apply(rule_tac x=x in bexI) apply(rule,rule) using b apply -  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
205  | 
apply(erule_tac x=xa in ballE) apply(erule_tac x=j in allE)+ by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
206  | 
ultimately show ?thesis by auto qed qed(auto)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
207  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
208  | 
lemma kle_imp_pointwise: "kle n x y \<Longrightarrow> (\<forall>j. x j \<le> y j)" unfolding kle_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
 | 
209  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
210  | 
lemma pointwise_antisym: fixes x::"nat \<Rightarrow> nat"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
211  | 
shows "(\<forall>j. x j \<le> y j) \<and> (\<forall>j. y j \<le> x j) \<longleftrightarrow> (x = y)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
212  | 
apply(rule, rule ext,erule conjE) apply(erule_tac x=xa in allE)+ by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
213  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
214  | 
lemma kle_trans: assumes "kle n x y" "kle n y z" "kle n x z \<or> kle n z x" shows "kle n x z"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
215  | 
using assms apply- apply(erule disjE) apply assumption proof- case goal1  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
216  | 
hence "x=z" apply- apply(rule ext) apply(drule kle_imp_pointwise)+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
217  | 
apply(erule_tac x=xa in allE)+ by auto thus ?case by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
218  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
219  | 
lemma kle_strict: assumes "kle n x y" "x \<noteq> y"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
220  | 
shows "\<forall>j. x j \<le> y j" "\<exists>k. 1 \<le> k \<and> k \<le> n \<and> x(k) < y(k)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
221  | 
apply(rule kle_imp_pointwise[OF assms(1)]) proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
222  | 
guess k using assms(1)[unfolded kle_def] .. note k = this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
223  | 
  show "\<exists>k. 1 \<le> k \<and> k \<le> n \<and> x(k) < y(k)" proof(cases "k={}")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
224  | 
case True hence "x=y" apply-apply(rule ext) using k by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
225  | 
thus ?thesis using assms(2) by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
226  | 
case False hence "(SOME k'. k' \<in> k) \<in> k" apply-apply(rule someI_ex) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
227  | 
thus ?thesis apply(rule_tac x="SOME k'. k' \<in> k" in exI) using k by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
228  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
229  | 
lemma kle_minimal: assumes "finite s" "s \<noteq> {}" "\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
230  | 
shows "\<exists>a\<in>s. \<forall>x\<in>s. kle n a x" proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
231  | 
have "\<exists>a\<in>s. \<forall>x\<in>s. \<forall>j. a j \<le> x j" apply(rule pointwise_minimal_pointwise_maximal(1)[OF assms(1-2)])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
232  | 
apply(rule,rule) apply(drule_tac assms(3)[rule_format],assumption) using kle_imp_pointwise by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
233  | 
then guess a .. note a=this show ?thesis apply(rule_tac x=a in bexI) proof fix x assume "x\<in>s"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
234  | 
show "kle n a x" using assms(3)[rule_format,OF a(1) `x\<in>s`] apply- proof(erule disjE)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
235  | 
assume "kle n x a" hence "x = a" apply- unfolding pointwise_antisym[THEN sym]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
236  | 
apply(drule kle_imp_pointwise) using a(2)[rule_format,OF `x\<in>s`] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
237  | 
thus ?thesis using kle_refl by auto qed qed(insert a, auto) qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
238  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
239  | 
lemma kle_maximal: assumes "finite s" "s \<noteq> {}" "\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
240  | 
shows "\<exists>a\<in>s. \<forall>x\<in>s. kle n x a" proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
241  | 
have "\<exists>a\<in>s. \<forall>x\<in>s. \<forall>j. a j \<ge> x j" apply(rule pointwise_minimal_pointwise_maximal(2)[OF assms(1-2)])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
242  | 
apply(rule,rule) apply(drule_tac assms(3)[rule_format],assumption) using kle_imp_pointwise by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
243  | 
then guess a .. note a=this show ?thesis apply(rule_tac x=a in bexI) proof fix x assume "x\<in>s"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
244  | 
show "kle n x a" using assms(3)[rule_format,OF a(1) `x\<in>s`] apply- proof(erule disjE)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
245  | 
assume "kle n a x" hence "x = a" apply- unfolding pointwise_antisym[THEN sym]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
246  | 
apply(drule kle_imp_pointwise) using a(2)[rule_format,OF `x\<in>s`] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
247  | 
thus ?thesis using kle_refl by auto qed qed(insert a, auto) qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
248  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
249  | 
lemma kle_strict_set: assumes "kle n x y" "x \<noteq> y"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
250  | 
  shows "1 \<le> card {k\<in>{1..n}. x k < y k}" proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
251  | 
guess i using kle_strict(2)[OF assms] ..  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
252  | 
  hence "card {i} \<le> card {k\<in>{1..n}. x k < y k}" apply- apply(rule card_mono) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
253  | 
thus ?thesis by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
254  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
255  | 
lemma kle_range_combine:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
256  | 
assumes "kle n x y" "kle n y z" "kle n x z \<or> kle n z x"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
257  | 
  "m1 \<le> card {k\<in>{1..n}. x k < y k}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
258  | 
  "m2 \<le> card {k\<in>{1..n}. y 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
 | 
259  | 
  shows "kle n x z \<and> m1 + m2 \<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
 | 
260  | 
apply(rule,rule kle_trans[OF assms(1-3)]) proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
261  | 
have "\<And>j. x j < y j \<Longrightarrow> x j < z j" apply(rule less_le_trans) using kle_imp_pointwise[OF assms(2)] by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
262  | 
have "\<And>j. y j < z j \<Longrightarrow> x j < z j" apply(rule le_less_trans) using kle_imp_pointwise[OF assms(1)] by auto ultimately  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
263  | 
  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}" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
264  | 
  have **:"{k \<in> {1..n}. x k < y k} \<inter> {k \<in> {1..n}. y k < z k} = {}" unfolding disjoint_iff_not_equal
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
265  | 
apply(rule,rule,unfold mem_Collect_eq,rule ccontr) apply(erule conjE)+ proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
266  | 
    fix i j assume as:"i \<in> {1..n}" "x i < y i" "j \<in> {1..n}" "y j < z j" "\<not> i \<noteq> j"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
267  | 
guess kx using assms(1)[unfolded kle_def] .. note kx=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
268  | 
have "x i < y i" using as by auto hence "i \<in> kx" using as(1) kx apply(rule_tac ccontr) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
269  | 
hence "x i + 1 = y i" using kx by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
270  | 
guess ky using assms(2)[unfolded kle_def] .. note ky=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
271  | 
have "y i < z i" using as by auto hence "i \<in> ky" using as(1) ky apply(rule_tac ccontr) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
272  | 
hence "y i + 1 = z i" using ky by auto ultimately  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
273  | 
have "z i = x i + 2" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
274  | 
thus False using assms(3) unfolding kle_def by(auto simp add: split_if_eq1) qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
275  | 
  have fin:"\<And>P. finite {x\<in>{1..n::nat}. P x}" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
276  | 
  have "m1 + m2 \<le> card {k\<in>{1..n}. x k < y k} + card {k\<in>{1..n}. y k < z k}" using assms(4-5) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
277  | 
  also have "\<dots> \<le>  card {k\<in>{1..n}. x k < z k}" unfolding card_Un_Int[OF fin fin] unfolding * ** by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
278  | 
  finally show " m1 + m2 \<le> card {k \<in> {1..n}. x k < z k}" by auto qed
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
279  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
280  | 
lemma kle_range_combine_l:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
281  | 
  assumes "kle n x y" "kle n y z" "kle n x z \<or> kle n z x" "m \<le> card {k\<in>{1..n}. y(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
 | 
282  | 
  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
 | 
283  | 
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
 | 
284  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
285  | 
lemma kle_range_combine_r:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
286  | 
  assumes "kle n x y" "kle n y z" "kle n x z \<or> kle n z x" "m \<le> card {k\<in>{1..n}. x k < y k}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
287  | 
  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
 | 
288  | 
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
 | 
289  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
290  | 
lemma kle_range_induct:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
291  | 
assumes "card s = Suc m" "\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
292  | 
  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}" proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
293  | 
have "finite s" "s\<noteq>{}" using assms(1) 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
 | 
294  | 
thus ?thesis using assms apply- proof(induct m arbitrary: s)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
295  | 
case 0 thus ?case using kle_refl by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
296  | 
case (Suc m) then obtain a where a:"a\<in>s" "\<forall>x\<in>s. kle n a x" using kle_minimal[of s n] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
297  | 
  show ?case proof(cases "s \<subseteq> {a}") case False
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
298  | 
    hence "card (s - {a}) = Suc m" "s - {a} \<noteq> {}" using card_Diff_singleton[OF _ a(1)] Suc(4) `finite s` by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
299  | 
    then obtain x b where xb:"x\<in>s - {a}" "b\<in>s - {a}" "kle n x b" "m \<le> card {k \<in> {1..n}. x k < b k}" 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
300  | 
      using Suc(1)[of "s - {a}"] using Suc(5) `finite s` by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
301  | 
    have "1 \<le> card {k \<in> {1..n}. a k < x k}" "m \<le> card {k \<in> {1..n}. x k < b k}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
302  | 
apply(rule kle_strict_set) apply(rule a(2)[rule_format]) using a and xb by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
303  | 
thus ?thesis apply(rule_tac x=a in bexI, rule_tac x=b in bexI)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
304  | 
using kle_range_combine[OF a(2)[rule_format] xb(3) Suc(5)[rule_format], of 1 "m"] using a(1) xb(1-2) by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
305  | 
    case True hence "s = {a}" using Suc(3) by auto hence "card s = 1" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
306  | 
hence False using Suc(4) `finite s` by auto thus ?thesis by auto qed qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
307  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
308  | 
lemma kle_Suc: "kle n x y \<Longrightarrow> kle (n + 1) x y"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
309  | 
unfolding kle_def apply(erule exE) apply(rule_tac x=k in exI) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
310  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
311  | 
lemma kle_trans_1: assumes "kle n x y" shows "x j \<le> y j" "y j \<le> x j + 1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
312  | 
using assms[unfolded kle_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
 | 
313  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
314  | 
lemma kle_trans_2: assumes "kle n a b" "kle n b c" "\<forall>j. c j \<le> a j + 1" shows "kle n a c" proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
315  | 
guess kk1 using assms(1)[unfolded kle_def] .. note kk1=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
316  | 
guess kk2 using assms(2)[unfolded kle_def] .. note kk2=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
317  | 
show ?thesis unfolding kle_def apply(rule_tac x="kk1 \<union> kk2" in exI) apply(rule) defer proof  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
318  | 
fix i show "c i = a i + (if i \<in> kk1 \<union> kk2 then 1 else 0)" proof(cases "i\<in>kk1 \<union> kk2")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
319  | 
case True hence "c i \<ge> a i + (if i \<in> kk1 \<union> kk2 then 1 else 0)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
320  | 
unfolding kk1[THEN conjunct2,rule_format,of i] kk2[THEN conjunct2,rule_format,of i] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
321  | 
moreover have "c i \<le> a i + (if i \<in> kk1 \<union> kk2 then 1 else 0)" using True assms(3) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
322  | 
ultimately show ?thesis by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
323  | 
case False thus ?thesis using kk1 kk2 by auto qed qed(insert kk1 kk2, auto) qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
324  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
325  | 
lemma kle_between_r: assumes "kle n a b" "kle n b c" "kle n a x" "kle n c x" shows "kle n b x"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
326  | 
apply(rule kle_trans_2[OF assms(2,4)]) proof have *:"\<And>c b x::nat. x \<le> c + 1 \<Longrightarrow> c \<le> b \<Longrightarrow> x \<le> b + 1" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
327  | 
fix j show "x j \<le> b j + 1" apply(rule *)using kle_trans_1[OF assms(1),of j] kle_trans_1[OF assms(3), of j] by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
328  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
329  | 
lemma kle_between_l: assumes "kle n a b" "kle n b c" "kle n x a" "kle n x c" shows "kle n x b"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
330  | 
apply(rule kle_trans_2[OF assms(3,1)]) proof have *:"\<And>c b x::nat. c \<le> x + 1 \<Longrightarrow> b \<le> c \<Longrightarrow> b \<le> x + 1" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
331  | 
fix j show "b j \<le> x j + 1" apply(rule *) using kle_trans_1[OF assms(2),of j] kle_trans_1[OF assms(4), of j] by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
332  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
333  | 
lemma kle_adjacent:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
334  | 
assumes "\<forall>j. b j = (if j = k then a(j) + 1 else a j)" "kle n a x" "kle n x b"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
335  | 
shows "(x = a) \<or> (x = b)" proof(cases "x k = a k")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
336  | 
case True show ?thesis apply(rule disjI1,rule ext) proof- fix j  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
337  | 
have "x j \<le> a j" using kle_imp_pointwise[OF assms(3),THEN spec[where x=j]]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
338  | 
unfolding assms(1)[rule_format] apply-apply(cases "j=k") using True by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
339  | 
thus "x j = a j" using kle_imp_pointwise[OF assms(2),THEN spec[where x=j]] by auto qed next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
340  | 
case False show ?thesis apply(rule disjI2,rule ext) proof- fix j  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
341  | 
have "x j \<ge> b j" using kle_imp_pointwise[OF assms(2),THEN spec[where x=j]]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
342  | 
unfolding assms(1)[rule_format] apply-apply(cases "j=k") using False by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
343  | 
thus "x j = b j" using kle_imp_pointwise[OF assms(3),THEN spec[where x=j]] by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
344  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
345  | 
subsection {* kuhn's notion of a simplex (a reformulation to avoid so much indexing). *}
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
346  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
347  | 
definition "ksimplex p n (s::(nat \<Rightarrow> nat) set) \<longleftrightarrow>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
348  | 
(card s = n + 1 \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
349  | 
(\<forall>x\<in>s. \<forall>j. x(j) \<le> p) \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
350  | 
        (\<forall>x\<in>s. \<forall>j. j\<notin>{1..n} \<longrightarrow> (x j = p)) \<and>
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
351  | 
(\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x))"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
352  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
353  | 
lemma ksimplexI:"card s = n + 1 \<Longrightarrow>  \<forall>x\<in>s. \<forall>j. x j \<le> p \<Longrightarrow> \<forall>x\<in>s. \<forall>j. j \<notin> {1..?n} \<longrightarrow> x j = ?p \<Longrightarrow> \<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x \<Longrightarrow> ksimplex p n s"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
354  | 
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
 | 
355  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
356  | 
lemma ksimplex_eq: "ksimplex p n (s::(nat \<Rightarrow> nat) set) \<longleftrightarrow>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
357  | 
(card s = n + 1 \<and> finite s \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
358  | 
(\<forall>x\<in>s. \<forall>j. x(j) \<le> p) \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
359  | 
        (\<forall>x\<in>s. \<forall>j. j\<notin>{1..n} \<longrightarrow> (x j = p)) \<and>
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
360  | 
(\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x))"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
361  | 
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
 | 
362  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
363  | 
lemma ksimplex_extrema: assumes "ksimplex p n s" obtains a b where "a \<in> s" "b \<in> s"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
364  | 
  "\<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))" proof(cases "n=0")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
365  | 
  case True obtain x where *:"s = {x}" using assms[unfolded ksimplex_eq True,THEN conjunct1]
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
366  | 
unfolding add_0_left card_1_exists by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
367  | 
show ?thesis apply(rule that[of x x]) unfolding * True by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
368  | 
note assm = assms[unfolded ksimplex_eq]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
369  | 
  case False have "s\<noteq>{}" 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
 | 
370  | 
  obtain a where a:"a\<in>s" "\<forall>x\<in>s. kle n a x" using `s\<noteq>{}` assm using kle_minimal[of s n] by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
371  | 
  obtain b where b:"b\<in>s" "\<forall>x\<in>s. kle n x b" using `s\<noteq>{}` assm using kle_maximal[of s n] by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
372  | 
  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}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
373  | 
using kle_range_induct[of s n 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
 | 
374  | 
  have "kle n c b \<and> n \<le> card {k \<in> {1..n}. c k < b k}" apply(rule kle_range_combine_r[where y=d]) using c_d a b by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
375  | 
  hence "kle n a b \<and> n \<le> card {k\<in>{1..n}. a(k) < b(k)}" apply-apply(rule kle_range_combine_l[where y=c]) using a `c\<in>s` `b\<in>s` by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
376  | 
  moreover have "card {1..n} \<ge> card {k\<in>{1..n}. a(k) < b(k)}" apply(rule card_mono) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
377  | 
  ultimately have *:"{k\<in>{1 .. n}. a k < b k} = {1..n}" apply- apply(rule card_subset_eq) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
378  | 
show ?thesis apply(rule that[OF a(1) b(1)]) defer apply(subst *[THEN sym]) unfolding mem_Collect_eq proof  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
379  | 
guess k using a(2)[rule_format,OF b(1),unfolded kle_def] .. note k=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
380  | 
    fix i show "b i = (if i \<in> {1..n} \<and> a i < b i then a i + 1 else a i)" proof(cases "i \<in> {1..n}")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
381  | 
case True thus ?thesis unfolding k[THEN conjunct2,rule_format] by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
382  | 
case False have "a i = p" using assm and False `a\<in>s` by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
383  | 
moreover have "b i = p" using assm and False `b\<in>s` by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
384  | 
ultimately show ?thesis by auto qed qed(insert a(2) b(2) assm,auto) qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
385  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
386  | 
lemma ksimplex_extrema_strong:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
387  | 
assumes "ksimplex p n s" "n \<noteq> 0" obtains a b where "a \<in> s" "b \<in> s" "a \<noteq> b"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
388  | 
  "\<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))" proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
389  | 
  obtain a b where ab:"a \<in> s" "b \<in> s" "\<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))" 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
390  | 
apply(rule ksimplex_extrema[OF assms(1)]) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
391  | 
have "a \<noteq> b" apply(rule ccontr) unfolding not_not apply(drule cong[of _ _ 1 1]) using ab(4) assms(2) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
392  | 
thus ?thesis apply(rule_tac that[of a b]) using ab by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
393  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
394  | 
lemma ksimplexD:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
395  | 
assumes "ksimplex p n s"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
396  | 
  shows "card s = n + 1" "finite s" "card s = n + 1" "\<forall>x\<in>s. \<forall>j. x j \<le> p" "\<forall>x\<in>s. \<forall>j. j \<notin> {1..?n} \<longrightarrow> x j = p"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
397  | 
"\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x" using assms unfolding ksimplex_eq by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
398  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
399  | 
lemma ksimplex_successor:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
400  | 
assumes "ksimplex p n s" "a \<in> s"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
401  | 
  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)))"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
402  | 
proof(cases "\<forall>x\<in>s. kle n x a") case True thus ?thesis by auto next note assm = ksimplexD[OF assms(1)]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
403  | 
  case False 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"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
404  | 
    using kle_minimal[of "{x\<in>s. \<not> kle n x a}" n] and 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
 | 
405  | 
  hence  **:"1 \<le> card {k\<in>{1..n}. a k < b k}" apply- apply(rule kle_strict_set) using assm(6) and `a\<in>s` by(auto simp add:kle_refl)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
406  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
407  | 
  let ?kle1 = "{x \<in> s. \<not> kle n x a}" have "card ?kle1 > 0" apply(rule ccontr) using assm(2) and False by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
408  | 
hence sizekle1: "card ?kle1 = Suc (card ?kle1 - 1)" using assm(2) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
409  | 
  obtain c d where c_d: "c \<in> s" "\<not> kle n c a" "d \<in> s" "\<not> kle n d a" "kle n c d" "card ?kle1 - 1 \<le> card {k \<in> {1..n}. c k < d k}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
410  | 
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
 | 
411  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
412  | 
  let ?kle2 = "{x \<in> s. kle n x a}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
413  | 
have "card ?kle2 > 0" apply(rule ccontr) using assm(6)[rule_format,of a a] and `a\<in>s` and assm(2) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
414  | 
hence sizekle2:"card ?kle2 = Suc (card ?kle2 - 1)" using assm(2) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
415  | 
  obtain e f where e_f: "e \<in> s" "kle n e a" "f \<in> s" "kle n f a" "kle n e f" "card ?kle2 - 1 \<le> card {k \<in> {1..n}. e k < f k}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
416  | 
using kle_range_induct[OF sizekle2, 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
 | 
417  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
418  | 
  have "card {k\<in>{1..n}. a k < b k} = 1" proof(rule ccontr) case goal1
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
419  | 
    hence as:"card {k\<in>{1..n}. a k < b k} \<ge> 2" using ** by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
420  | 
    have *:"finite ?kle2" "finite ?kle1" "?kle2 \<union> ?kle1 = s" "?kle2 \<inter> ?kle1 = {}" using assm(2) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
421  | 
have "(card ?kle2 - 1) + 2 + (card ?kle1 - 1) = card ?kle2 + card ?kle1" using sizekle1 sizekle2 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
422  | 
also have "\<dots> = n + 1" unfolding card_Un_Int[OF *(1-2)] *(3-) using assm(3) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
423  | 
finally have n:"(card ?kle2 - 1) + (2 + (card ?kle1 - 1)) = n + 1" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
424  | 
    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}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
425  | 
apply(rule kle_range_combine_r[where y=f]) using e_f using `a\<in>s` assm(6) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
426  | 
    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}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
427  | 
apply(rule kle_range_combine_l[where y=c]) using c_d using assm(6) and b by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
428  | 
    hence "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}" apply-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
429  | 
apply(rule kle_range_combine[where y=b]) using as and b assm(6) `a\<in>s` `d\<in>s` apply- by blast+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
430  | 
    ultimately have "kle n e d \<and> (card ?kle2 - 1) + (2 + (card ?kle1 - 1)) \<le> card {k\<in>{1..n}. e k < d k}" apply-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
431  | 
apply(rule kle_range_combine[where y=a]) using assm(6)[rule_format,OF `e\<in>s` `d\<in>s`] apply - by blast+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
432  | 
    moreover have "card {k \<in> {1..n}. e k < d k} \<le> card {1..n}" apply(rule card_mono) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
433  | 
ultimately show False unfolding n by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
434  | 
then guess k unfolding card_1_exists .. note k=this[unfolded mem_Collect_eq]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
435  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
436  | 
show ?thesis apply(rule disjI2) apply(rule_tac x=b in bexI,rule_tac x=k in bexI) proof  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
437  | 
fix j::nat have "kle n a b" using b and assm(6)[rule_format, OF `a\<in>s` `b\<in>s`] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
438  | 
then guess kk unfolding kle_def .. note kk_raw=this note kk=this[THEN conjunct2,rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
439  | 
have kkk:"k\<in>kk" apply(rule ccontr) using k(1) unfolding kk by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
440  | 
show "b j = (if j = k then a j + 1 else a j)" proof(cases "j\<in>kk")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
441  | 
case True hence "j=k" apply-apply(rule k(2)[rule_format]) using kk_raw kkk by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
442  | 
thus ?thesis unfolding kk using kkk by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
443  | 
case False hence "j\<noteq>k" using k(2)[rule_format, of j k] using kk_raw kkk by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
444  | 
thus ?thesis unfolding kk using kkk using False by auto qed qed(insert k(1) `b\<in>s`, auto) qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
445  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
446  | 
lemma ksimplex_predecessor:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
447  | 
assumes "ksimplex p n s" "a \<in> s"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
448  | 
  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)))"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
449  | 
proof(cases "\<forall>x\<in>s. kle n a x") case True thus ?thesis by auto next note assm = ksimplexD[OF assms(1)]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
450  | 
  case False 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" 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
451  | 
    using kle_maximal[of "{x\<in>s. \<not> kle n a x}" n] and 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
 | 
452  | 
  hence  **:"1 \<le> card {k\<in>{1..n}. a k > b k}" apply- apply(rule kle_strict_set) using assm(6) and `a\<in>s` by(auto simp add:kle_refl)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
453  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
454  | 
  let ?kle1 = "{x \<in> s. \<not> kle n a x}" have "card ?kle1 > 0" apply(rule ccontr) using assm(2) and False by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
455  | 
hence sizekle1:"card ?kle1 = Suc (card ?kle1 - 1)" using assm(2) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
456  | 
  obtain c d where c_d: "c \<in> s" "\<not> kle n a c" "d \<in> s" "\<not> kle n a d" "kle n d c" "card ?kle1 - 1 \<le> card {k \<in> {1..n}. c k > d k}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
457  | 
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
 | 
458  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
459  | 
  let ?kle2 = "{x \<in> s. kle n a x}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
460  | 
have "card ?kle2 > 0" apply(rule ccontr) using assm(6)[rule_format,of a a] and `a\<in>s` and assm(2) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
461  | 
hence sizekle2:"card ?kle2 = Suc (card ?kle2 - 1)" using assm(2) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
462  | 
  obtain e f where e_f: "e \<in> s" "kle n a e" "f \<in> s" "kle n a f" "kle n f e" "card ?kle2 - 1 \<le> card {k \<in> {1..n}. e k > f k}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
463  | 
using kle_range_induct[OF sizekle2, 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
 | 
464  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
465  | 
  have "card {k\<in>{1..n}. a k > b k} = 1" proof(rule ccontr) case goal1
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
466  | 
    hence as:"card {k\<in>{1..n}. a k > b k} \<ge> 2" using ** by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
467  | 
    have *:"finite ?kle2" "finite ?kle1" "?kle2 \<union> ?kle1 = s" "?kle2 \<inter> ?kle1 = {}" using assm(2) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
468  | 
have "(card ?kle2 - 1) + 2 + (card ?kle1 - 1) = card ?kle2 + card ?kle1" using sizekle1 sizekle2 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
469  | 
also have "\<dots> = n + 1" unfolding card_Un_Int[OF *(1-2)] *(3-) using assm(3) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
470  | 
finally have n:"(card ?kle1 - 1) + 2 + (card ?kle2 - 1) = n + 1" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
471  | 
    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}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
472  | 
apply(rule kle_range_combine_l[where y=f]) using e_f using `a\<in>s` assm(6) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
473  | 
    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}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
474  | 
apply(rule kle_range_combine_r[where y=c]) using c_d using assm(6) and b by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
475  | 
    hence "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}" apply-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
476  | 
apply(rule kle_range_combine[where y=b]) using as and b assm(6) `a\<in>s` `d\<in>s` by blast+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
477  | 
    ultimately have "kle n d e \<and> (card ?kle1 - 1 + 2) + (card ?kle2 - 1) \<le> card {k\<in>{1..n}. e k > d k}" apply-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
478  | 
apply(rule kle_range_combine[where y=a]) using assm(6)[rule_format,OF `e\<in>s` `d\<in>s`] apply - by blast+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
479  | 
    moreover have "card {k \<in> {1..n}. e k > d k} \<le> card {1..n}" apply(rule card_mono) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
480  | 
ultimately show False unfolding n by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
481  | 
then guess k unfolding card_1_exists .. note k=this[unfolded mem_Collect_eq]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
482  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
483  | 
show ?thesis apply(rule disjI2) apply(rule_tac x=b in bexI,rule_tac x=k in bexI) proof  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
484  | 
fix j::nat have "kle n b a" using b and assm(6)[rule_format, OF `a\<in>s` `b\<in>s`] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
485  | 
then guess kk unfolding kle_def .. note kk_raw=this note kk=this[THEN conjunct2,rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
486  | 
have kkk:"k\<in>kk" apply(rule ccontr) using k(1) unfolding kk by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
487  | 
show "a j = (if j = k then b j + 1 else b j)" proof(cases "j\<in>kk")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
488  | 
case True hence "j=k" apply-apply(rule k(2)[rule_format]) using kk_raw kkk by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
489  | 
thus ?thesis unfolding kk using kkk by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
490  | 
case False hence "j\<noteq>k" using k(2)[rule_format, of j k] using kk_raw kkk by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
491  | 
thus ?thesis unfolding kk using kkk using False by auto qed qed(insert k(1) `b\<in>s`, auto) qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
492  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
493  | 
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
 | 
494  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
495  | 
lemma card_funspace': assumes "finite s" "finite t" "card s = m" "card t = n"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
496  | 
  shows "card {f. (\<forall>x\<in>s. f x \<in> t) \<and> (\<forall>x\<in>UNIV - s. f x = d)} = n ^ m" (is "card (?M s) = _")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
497  | 
using assms apply - proof(induct m arbitrary: s)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
498  | 
  have *:"{f. \<forall>x. f x = d} = {\<lambda>x. d}" apply(rule set_ext,rule)unfolding mem_Collect_eq apply(rule,rule ext) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
499  | 
case 0 thus ?case by(auto simp add: *) next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
500  | 
case (Suc m) guess a using card_eq_SucD[OF Suc(4)] .. then guess s0  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
501  | 
apply(erule_tac exE) apply(erule conjE)+ . note as0 = this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
502  | 
have **:"card s0 = m" using as0 using Suc(2) Suc(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
 | 
503  | 
  let ?l = "(\<lambda>(b,g) x. if x = a then b else g x)" have *:"?M (insert a s0) = ?l ` {(b,g). b\<in>t \<and> g\<in>?M s0}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
504  | 
apply(rule set_ext,rule) unfolding mem_Collect_eq image_iff apply(erule conjE)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
505  | 
apply(rule_tac x="(x a, \<lambda>y. if y\<in>s0 then x y else d)" in bexI) apply(rule ext) prefer 3 apply rule defer  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
506  | 
apply(erule bexE,rule) unfolding mem_Collect_eq apply(erule splitE)+ apply(erule conjE)+ proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
507  | 
fix x xa xb xc y assume as:"x = (\<lambda>(b, g) x. if x = a then b else g x) xa" "xb \<in> UNIV - insert a s0" "xa = (xc, y)" "xc \<in> t"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
508  | 
"\<forall>x\<in>s0. y x \<in> t" "\<forall>x\<in>UNIV - s0. y x = d" thus "x xb = d" unfolding as by auto qed auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
509  | 
  have inj:"inj_on ?l {(b,g). b\<in>t \<and> g\<in>?M s0}" unfolding inj_on_def apply(rule,rule,rule) unfolding mem_Collect_eq apply(erule splitE conjE)+ proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
510  | 
case goal1 note as = this(1,4-)[unfolded goal1 split_conv]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
511  | 
have "xa = xb" using as(1)[THEN cong[of _ _ a]] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
512  | 
moreover have "ya = yb" proof(rule ext) fix x show "ya x = yb x" proof(cases "x = a")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
513  | 
case False thus ?thesis using as(1)[THEN cong[of _ _ x x]] by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
514  | 
case True thus ?thesis using as(5,7) using as0(2) by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
515  | 
ultimately show ?case unfolding goal1 by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
516  | 
have "finite s0" using `finite s` unfolding as0 by simp  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
517  | 
show ?case unfolding as0 * card_image[OF inj] using assms  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
518  | 
unfolding SetCompr_Sigma_eq apply-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
519  | 
unfolding card_cartesian_product  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
520  | 
using Suc(1)[OF `finite s0` `finite t` ** `card t = n`] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
521  | 
qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
522  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
523  | 
lemma card_funspace: assumes "finite s" "finite t"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
524  | 
  shows "card {f. (\<forall>x\<in>s. f x \<in> t) \<and> (\<forall>x\<in>UNIV - s. f x = d)} = (card t) ^ (card s)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
525  | 
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
 | 
526  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
527  | 
lemma finite_funspace: assumes "finite s" "finite t"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
528  | 
  shows "finite {f. (\<forall>x\<in>s. f x \<in> t) \<and> (\<forall>x\<in>UNIV - s. f x = d)}" (is "finite ?S")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
529  | 
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
 | 
530  | 
case True  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
531  | 
have "card ?S = (card t) ^ (card s)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
532  | 
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
 | 
533  | 
thus ?thesis using True 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
 | 
534  | 
next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
535  | 
  case False hence "t = {}" using `finite t` by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
536  | 
show ?thesis  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
537  | 
  proof (cases "s = {}")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
538  | 
    have *:"{f. \<forall>x. f x = d} = {\<lambda>x. d}" by (auto intro: ext)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
539  | 
    case True thus ?thesis using `t = {}` by (auto simp: *)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
540  | 
next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
541  | 
    case False thus ?thesis using `t = {}` by simp
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
542  | 
qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
543  | 
qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
544  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
545  | 
lemma finite_simplices: "finite {s. ksimplex p n s}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
546  | 
  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)}}"])
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
547  | 
unfolding ksimplex_def defer apply(rule finite_Collect_subsets) apply(rule finite_funspace) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
548  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
549  | 
lemma simplex_top_face: assumes "0<p" "\<forall>x\<in>f. x (n + 1) = p"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
550  | 
  shows "(\<exists>s a. ksimplex p (n + 1) s \<and> a \<in> s \<and> (f = s - {a})) \<longleftrightarrow> ksimplex p n f" (is "?ls = ?rs") proof
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
551  | 
assume ?ls then guess s .. then guess a apply-apply(erule exE,(erule conjE)+) . note sa=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
552  | 
show ?rs unfolding ksimplex_def sa(3) apply(rule) defer apply rule defer apply(rule,rule,rule,rule) defer apply(rule,rule) proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
553  | 
    fix x y assume as:"x \<in>s - {a}" "y \<in>s - {a}" have xyp:"x (n + 1) = y (n + 1)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
554  | 
using as(1)[unfolded sa(3)[THEN sym], THEN assms(2)[rule_format]]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
555  | 
using as(2)[unfolded sa(3)[THEN sym], THEN assms(2)[rule_format]] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
556  | 
show "kle n x y \<or> kle n y x" proof(cases "kle (n + 1) x y")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
557  | 
case True then guess k unfolding kle_def .. note k=this hence *:"n+1 \<notin> k" using xyp by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
558  | 
      have "\<not> (\<exists>x\<in>k. x\<notin>{1..n})" apply(rule ccontr) unfolding not_not apply(erule bexE) proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
559  | 
	fix x assume as:"x \<in> k" "x \<notin> {1..n}" have "x \<noteq> n+1" using as and * by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
560  | 
thus False using as and k[THEN conjunct1,unfolded subset_eq] by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
561  | 
thus ?thesis apply-apply(rule disjI1) unfolding kle_def using k apply(rule_tac x=k in exI) by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
562  | 
case False hence "kle (n + 1) y x" using ksimplexD(6)[OF sa(1),rule_format, of x y] using as by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
563  | 
then guess k unfolding kle_def .. note k=this hence *:"n+1 \<notin> k" using xyp by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
564  | 
      hence "\<not> (\<exists>x\<in>k. x\<notin>{1..n})" apply-apply(rule ccontr) unfolding not_not apply(erule bexE) proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
565  | 
	fix x assume as:"x \<in> k" "x \<notin> {1..n}" have "x \<noteq> n+1" using as and * by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
566  | 
thus False using as and k[THEN conjunct1,unfolded subset_eq] by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
567  | 
thus ?thesis apply-apply(rule disjI2) unfolding kle_def using k apply(rule_tac x=k in exI) by auto qed next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
568  | 
    fix x j assume as:"x\<in>s - {a}" "j\<notin>{1..n}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
569  | 
thus "x j = p" using as(1)[unfolded sa(3)[THEN sym], THEN assms(2)[rule_format]]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
570  | 
apply(cases "j = n+1") using sa(1)[unfolded ksimplex_def] by auto qed(insert sa ksimplexD[OF sa(1)], auto) next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
571  | 
assume ?rs note rs=ksimplexD[OF this] guess a b apply(rule ksimplex_extrema[OF `?rs`]) . note ab = this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
572  | 
def c \<equiv> "\<lambda>i. if i = (n + 1) then p - 1 else a i"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
573  | 
have "c\<notin>f" apply(rule ccontr) unfolding not_not apply(drule assms(2)[rule_format]) unfolding c_def using assms(1) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
574  | 
thus ?ls apply(rule_tac x="insert c f" in exI,rule_tac x=c in exI) unfolding ksimplex_def conj_assoc  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
575  | 
apply(rule conjI) defer apply(rule conjI) defer apply(rule conjI) defer apply(rule conjI) defer  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
576  | 
proof(rule_tac[3-5] ballI allI)+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
577  | 
fix x j assume x:"x \<in> insert c f" thus "x j \<le> p" proof (cases "x=c")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
578  | 
case True show ?thesis unfolding True c_def apply(cases "j=n+1") using ab(1) and rs(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
 | 
579  | 
qed(insert x rs(4), auto simp add:c_def)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
580  | 
    show "j \<notin> {1..n + 1} \<longrightarrow> x j = p" apply(cases "x=c") using x ab(1) rs(5) unfolding c_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
 | 
581  | 
    { fix z assume z:"z \<in> insert c f" hence "kle (n + 1) c z" apply(cases "z = c") (*defer apply(rule kle_Suc)*) proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
582  | 
case False hence "z\<in>f" using z by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
583  | 
then guess k apply(drule_tac ab(3)[THEN bspec[where x=z], THEN conjunct1]) unfolding kle_def apply(erule exE) .  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
584  | 
thus "kle (n + 1) c z" unfolding kle_def apply(rule_tac x="insert (n + 1) k" in exI) unfolding c_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
585  | 
using ab using rs(5)[rule_format,OF ab(1),of "n + 1"] assms(1) by auto qed auto } note * = this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
586  | 
fix y assume y:"y \<in> insert c f" show "kle (n + 1) x y \<or> kle (n + 1) y x" proof(cases "x = c \<or> y = c")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
587  | 
case False hence **:"x\<in>f" "y\<in>f" using x y by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
588  | 
show ?thesis using rs(6)[rule_format,OF **] by(auto dest: kle_Suc) qed(insert * x y, auto)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
589  | 
qed(insert rs, auto) qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
590  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
591  | 
lemma ksimplex_fix_plane:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
592  | 
  assumes "a \<in> s" "j\<in>{1..n::nat}" "\<forall>x\<in>s - {a}. x j = q" "a0 \<in> s" "a1 \<in> s"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
593  | 
  "\<forall>i. a1 i = ((if i\<in>{1..n} then a0 i + 1 else a0 i)::nat)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
594  | 
shows "(a = a0) \<or> (a = a1)" proof- 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" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
595  | 
show ?thesis apply(rule ccontr) using *[OF assms(3), of a0 a1] unfolding assms(6)[THEN spec[where x=j]]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
596  | 
using assms(1-2,4-5) by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
597  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
598  | 
lemma ksimplex_fix_plane_0:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
599  | 
  assumes "a \<in> s" "j\<in>{1..n::nat}" "\<forall>x\<in>s - {a}. x j = 0" "a0 \<in> s" "a1 \<in> s"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
600  | 
  "\<forall>i. a1 i = ((if i\<in>{1..n} then a0 i + 1 else a0 i)::nat)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
601  | 
shows "a = a1" apply(rule ccontr) using ksimplex_fix_plane[OF assms]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
602  | 
using assms(3)[THEN bspec[where x=a1]] using assms(2,5)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
603  | 
unfolding assms(6)[THEN spec[where x=j]] by simp  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
604  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
605  | 
lemma ksimplex_fix_plane_p:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
606  | 
  assumes "ksimplex p n s" "a \<in> s" "j\<in>{1..n}" "\<forall>x\<in>s - {a}. x j = p" "a0 \<in> s" "a1 \<in> s"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
607  | 
  "\<forall>i. a1 i = (if i\<in>{1..n} then a0 i + 1 else a0 i)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
608  | 
shows "a = a0" proof(rule ccontr) note s = ksimplexD[OF assms(1),rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
609  | 
  assume as:"a \<noteq> a0" hence *:"a0 \<in> s - {a}" using assms(5) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
610  | 
hence "a1 = a" using ksimplex_fix_plane[OF assms(2-)] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
611  | 
thus False using as using assms(3,5) and assms(7)[rule_format,of j]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
612  | 
unfolding assms(4)[rule_format,OF *] using s(4)[OF assms(6), of j] by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
613  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
614  | 
lemma ksimplex_replace_0:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
615  | 
  assumes "ksimplex p n s" "a \<in> s" "n \<noteq> 0" "j\<in>{1..n}" "\<forall>x\<in>s - {a}. x j = 0"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
616  | 
  shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 1" proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
617  | 
  have *:"\<And>s' a a'. s' - {a'} = s - {a} \<Longrightarrow> a' = a \<Longrightarrow> a' \<in> s' \<Longrightarrow> a \<in> s \<Longrightarrow> (s' = s)" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
618  | 
  have **:"\<And>s' a'. ksimplex p n s' \<Longrightarrow> a' \<in> s' \<Longrightarrow> s' - {a'} = s - {a} \<Longrightarrow> s' = s" proof- case goal1
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
619  | 
guess a0 a1 apply(rule ksimplex_extrema_strong[OF assms(1,3)]) . note exta = this[rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
620  | 
have a:"a = a1" apply(rule ksimplex_fix_plane_0[OF assms(2,4-5)]) using exta(1-2,5) by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
621  | 
guess b0 b1 apply(rule ksimplex_extrema_strong[OF goal1(1) assms(3)]) . note extb = this[rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
622  | 
have a':"a' = b1" apply(rule ksimplex_fix_plane_0[OF goal1(2) assms(4), of b0]) unfolding goal1(3) using assms extb goal1 by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
623  | 
have "b0 = a0" unfolding kle_antisym[THEN sym, of b0 a0 n] using exta extb using goal1(3) unfolding a a' by blast  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
624  | 
hence "b1 = a1" apply-apply(rule ext) unfolding exta(5) extb(5) by auto ultimately  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
625  | 
show "s' = s" apply-apply(rule *[of _ a1 b1]) using exta(1-2) extb(1-2) goal1 by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
626  | 
show ?thesis unfolding card_1_exists apply-apply(rule ex1I[of _ s])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
627  | 
unfolding mem_Collect_eq defer apply(erule conjE bexE)+ apply(rule_tac a'=b in **) using assms(1,2) by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
628  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
629  | 
lemma ksimplex_replace_1:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
630  | 
  assumes "ksimplex p n s" "a \<in> s" "n \<noteq> 0" "j\<in>{1..n}" "\<forall>x\<in>s - {a}. x j = p"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
631  | 
  shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 1" proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
632  | 
  have lem:"\<And>a a' s'. s' - {a'} = s - {a} \<Longrightarrow> a' = a \<Longrightarrow> a' \<in> s' \<Longrightarrow> a \<in> s \<Longrightarrow> s' = s" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
633  | 
  have lem:"\<And>s' a'. ksimplex p n s' \<Longrightarrow> a'\<in>s' \<Longrightarrow> s' - {a'} = s - {a} \<Longrightarrow> s' = s" proof- case goal1
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
634  | 
guess a0 a1 apply(rule ksimplex_extrema_strong[OF assms(1,3)]) . note exta = this[rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
635  | 
have a:"a = a0" apply(rule ksimplex_fix_plane_p[OF assms(1-2,4-5) exta(1,2)]) unfolding exta by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
636  | 
guess b0 b1 apply(rule ksimplex_extrema_strong[OF goal1(1) assms(3)]) . note extb = this[rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
637  | 
have a':"a' = b0" apply(rule ksimplex_fix_plane_p[OF goal1(1-2) assms(4), of _ b1]) unfolding goal1 extb using extb(1,2) assms(5) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
638  | 
moreover have *:"b1 = a1" unfolding kle_antisym[THEN sym, of b1 a1 n] using exta extb using goal1(3) unfolding a a' by blast moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
639  | 
have "a0 = b0" apply(rule ext) proof- case goal1 show "a0 x = b0 x"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
640  | 
	using *[THEN cong, of x x] unfolding exta extb apply-apply(cases "x\<in>{1..n}") by auto qed
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
641  | 
ultimately show "s' = s" apply-apply(rule lem[OF goal1(3) _ goal1(2) assms(2)]) by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
642  | 
show ?thesis unfolding card_1_exists apply(rule ex1I[of _ s]) unfolding mem_Collect_eq apply(rule,rule assms(1))  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
643  | 
apply(rule_tac x=a in bexI) prefer 3 apply(erule conjE bexE)+ apply(rule_tac a'=b in lem) using assms(1-2) by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
644  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
645  | 
lemma ksimplex_replace_2:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
646  | 
  assumes "ksimplex p n s" "a \<in> s" "n \<noteq> 0" "~(\<exists>j\<in>{1..n}. \<forall>x\<in>s - {a}. x j = 0)" "~(\<exists>j\<in>{1..n}. \<forall>x\<in>s - {a}. x j = p)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
647  | 
  shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 2" (is "card ?A = 2")  proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
648  | 
  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" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
649  | 
  have lem2:"\<And>a b. a\<in>s \<Longrightarrow> b\<noteq>a \<Longrightarrow> s \<noteq> insert b (s - {a})" proof case goal1
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
650  | 
    hence "a\<in>insert b (s - {a})" by auto hence "a\<in> s - {a}" unfolding insert_iff using goal1 by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
651  | 
thus False by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
652  | 
guess a0 a1 apply(rule ksimplex_extrema_strong[OF assms(1,3)]) . note a0a1=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
653  | 
  { assume "a=a0"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
654  | 
have *:"\<And>P Q. (P \<or> Q) \<Longrightarrow> \<not> P \<Longrightarrow> Q" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
655  | 
have "\<exists>x\<in>s. \<not> kle n x a0" apply(rule_tac x=a1 in bexI) proof assume as:"kle n a1 a0"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
656  | 
show False using kle_imp_pointwise[OF as,THEN spec[where x=1]] unfolding a0a1(5)[THEN spec[where x=1]]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
657  | 
using assms(3) by auto qed(insert a0a1,auto)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
658  | 
    hence "\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. y j = (if j = k then a0 j + 1 else a0 j)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
659  | 
apply(rule_tac *[OF ksimplex_successor[OF assms(1-2),unfolded `a=a0`]]) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
660  | 
then guess a2 .. from this(2) guess k .. note k=this note a2=`a2\<in>s`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
661  | 
def a3 \<equiv> "\<lambda>j. if j = k then a1 j + 1 else a1 j"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
662  | 
have "a3 \<notin> s" proof assume "a3\<in>s" hence "kle n a3 a1" using a0a1(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
 | 
663  | 
thus False apply(drule_tac kle_imp_pointwise) unfolding a3_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
664  | 
apply(erule_tac x=k in allE) by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
665  | 
hence "a3 \<noteq> a0" "a3 \<noteq> a1" using a0a1 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
666  | 
have "a2 \<noteq> a0" using k(2)[THEN spec[where x=k]] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
667  | 
    have lem3:"\<And>x. x\<in>(s - {a0}) \<Longrightarrow> kle n a2 x" proof(rule ccontr) case goal1 hence as:"x\<in>s" "x\<noteq>a0" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
668  | 
have "kle n a2 x \<or> kle n x a2" using ksimplexD(6)[OF assms(1)] and as `a2\<in>s` by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
669  | 
have "kle n a0 x" using a0a1(4) as by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
670  | 
ultimately have "x = a0 \<or> x = a2" apply-apply(rule kle_adjacent[OF k(2)]) using goal1(2) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
671  | 
hence "x = a2" using as by auto thus False using goal1(2) using kle_refl by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
672  | 
    let ?s = "insert a3 (s - {a0})" have "ksimplex p n ?s" apply(rule ksimplexI) proof(rule_tac[2-] ballI,rule_tac[4] ballI)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
673  | 
show "card ?s = n + 1" using ksimplexD(2-3)[OF assms(1)]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
674  | 
using `a3\<noteq>a0` `a3\<notin>s` `a0\<in>s` by(auto simp add:card_insert_if)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
675  | 
      fix x assume x:"x \<in> insert a3 (s - {a0})"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
676  | 
show "\<forall>j. x j \<le> p" proof(rule,cases "x = a3")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
677  | 
fix j case False thus "x j\<le>p" using x ksimplexD(4)[OF assms(1)] by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
678  | 
fix j case True show "x j\<le>p" unfolding True proof(cases "j=k")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
679  | 
case False thus "a3 j \<le>p" unfolding True a3_def using `a1\<in>s` ksimplexD(4)[OF assms(1)] by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
680  | 
guess a4 using assms(5)[unfolded bex_simps ball_simps,rule_format,OF k(1)] .. note a4=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
681  | 
have "a2 k \<le> a4 k" using lem3[OF a4(1)[unfolded `a=a0`],THEN kle_imp_pointwise] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
682  | 
also have "\<dots> < p" using ksimplexD(4)[OF assms(1),rule_format,of a4 k] using a4 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
683  | 
finally have *:"a0 k + 1 < p" unfolding k(2)[rule_format] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
684  | 
case True thus "a3 j \<le>p" unfolding a3_def unfolding a0a1(5)[rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
685  | 
using k(1) k(2)assms(5) using * by simp qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
686  | 
      show "\<forall>j. j \<notin> {1..n} \<longrightarrow> x j = p" proof(rule,rule,cases "x=a3") fix j::nat assume j:"j\<notin>{1..n}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
687  | 
	{ case False thus "x j = p" using j x ksimplexD(5)[OF assms(1)] by auto }
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
688  | 
case True show "x j = p" unfolding True a3_def using j k(1)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
689  | 
using ksimplexD(5)[OF assms(1),rule_format,OF `a1\<in>s` j] by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
690  | 
      fix y assume y:"y\<in>insert a3 (s - {a0})"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
691  | 
have lem4:"\<And>x. x\<in>s \<Longrightarrow> x\<noteq>a0 \<Longrightarrow> kle n x a3" proof- case goal1  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
692  | 
guess kk using a0a1(4)[rule_format,OF `x\<in>s`,THEN conjunct2,unfolded kle_def]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
693  | 
apply-apply(erule exE,erule conjE) . note kk=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
694  | 
have "k\<notin>kk" proof assume "k\<in>kk"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
695  | 
hence "a1 k = x k + 1" using kk by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
696  | 
hence "a0 k = x k" unfolding a0a1(5)[rule_format] using k(1) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
697  | 
hence "a2 k = x k + 1" unfolding k(2)[rule_format] by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
698  | 
have "a2 k \<le> x k" using lem3[of x,THEN kle_imp_pointwise] goal1 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
699  | 
ultimately show False by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
700  | 
thus ?case unfolding kle_def apply(rule_tac x="insert k kk" in exI) using kk(1)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
701  | 
unfolding a3_def kle_def kk(2)[rule_format] using k(1) by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
702  | 
show "kle n x y \<or> kle n y x" proof(cases "y=a3")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
703  | 
case True show ?thesis unfolding True apply(cases "x=a3") defer apply(rule disjI1,rule lem4)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
704  | 
using x by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
705  | 
case False show ?thesis proof(cases "x=a3") case True show ?thesis unfolding True  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
706  | 
apply(rule disjI2,rule lem4) using y False by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
707  | 
case False thus ?thesis apply(rule_tac ksimplexD(6)[OF assms(1),rule_format])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
708  | 
using x y `y\<noteq>a3` by auto qed qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
709  | 
    hence "insert a3 (s - {a0}) \<in> ?A" unfolding mem_Collect_eq apply-apply(rule,assumption)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
710  | 
apply(rule_tac x="a3" in bexI) unfolding `a=a0` using `a3\<notin>s` by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
711  | 
    have "s \<in> ?A" using assms(1,2) by auto ultimately have  "?A \<supseteq> {s, insert a3 (s - {a0})}" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
712  | 
    moreover have "?A \<subseteq> {s, insert a3 (s - {a0})}" apply(rule) unfolding mem_Collect_eq proof(erule conjE)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
713  | 
      fix s' assume as:"ksimplex p n s'" and "\<exists>b\<in>s'. s' - {b} = s - {a}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
714  | 
from this(2) guess a' .. note a'=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
715  | 
guess a_min a_max apply(rule ksimplex_extrema_strong[OF as assms(3)]) . note min_max=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
716  | 
      have *:"\<forall>x\<in>s' - {a'}. x k = a2 k" unfolding a' proof fix x assume x:"x\<in>s-{a}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
717  | 
hence "kle n a2 x" apply-apply(rule lem3) using `a=a0` by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
718  | 
hence "a2 k \<le> x k" apply(drule_tac kle_imp_pointwise) by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
719  | 
have "x k \<le> a2 k" unfolding k(2)[rule_format] using a0a1(4)[rule_format,of x,THEN conjunct1]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
720  | 
unfolding kle_def using x by auto ultimately show "x k = a2 k" by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
721  | 
have **:"a'=a_min \<or> a'=a_max" apply(rule ksimplex_fix_plane[OF a'(1) k(1) *]) using min_max by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
722  | 
      show "s' \<in> {s, insert a3 (s - {a0})}" proof(cases "a'=a_min")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
723  | 
case True have "a_max = a1" unfolding kle_antisym[THEN sym,of a_max a1 n] apply(rule)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
724  | 
apply(rule a0a1(4)[rule_format,THEN conjunct2]) defer proof(rule min_max(4)[rule_format,THEN conjunct2])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
725  | 
show "a1\<in>s'" using a' unfolding `a=a0` using a0a1 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
726  | 
show "a_max \<in> s" proof(rule ccontr) assume "a_max\<notin>s"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
727  | 
hence "a_max = a'" using a' min_max by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
728  | 
thus False unfolding True using min_max by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
729  | 
hence "\<forall>i. a_max i = a1 i" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
730  | 
hence "a' = a" unfolding True `a=a0` apply-apply(subst expand_fun_eq,rule)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
731  | 
apply(erule_tac x=x in allE) unfolding a0a1(5)[rule_format] min_max(5)[rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
732  | 
	proof- case goal1 thus ?case apply(cases "x\<in>{1..n}") by auto qed
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
733  | 
hence "s' = s" apply-apply(rule lem1[OF a'(2)]) using `a\<in>s` `a'\<in>s'` by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
734  | 
thus ?thesis by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
735  | 
case False hence as:"a' = a_max" using ** by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
736  | 
have "a_min = a2" unfolding kle_antisym[THEN sym, of _ _ n] apply rule  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
737  | 
apply(rule min_max(4)[rule_format,THEN conjunct1]) defer proof(rule lem3)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
738  | 
	  show "a_min \<in> s - {a0}" unfolding a'(2)[THEN sym,unfolded `a=a0`] 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
739  | 
unfolding as using min_max(1-3) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
740  | 
have "a2 \<noteq> a" unfolding `a=a0` using k(2)[rule_format,of k] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
741  | 
	  hence "a2 \<in> s - {a}" using a2 by auto thus "a2 \<in> s'" unfolding a'(2)[THEN sym] by auto qed
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
742  | 
hence "\<forall>i. a_min i = a2 i" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
743  | 
hence "a' = a3" unfolding as `a=a0` apply-apply(subst expand_fun_eq,rule)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
744  | 
apply(erule_tac x=x in allE) unfolding a0a1(5)[rule_format] min_max(5)[rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
745  | 
unfolding a3_def k(2)[rule_format] unfolding a0a1(5)[rule_format] proof- case goal1  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
746  | 
	  show ?case unfolding goal1 apply(cases "x\<in>{1..n}") defer apply(cases "x=k")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
747  | 
	    using `k\<in>{1..n}` by auto qed
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
748  | 
	hence "s' = insert a3 (s - {a0})" apply-apply(rule lem1) defer apply assumption
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
749  | 
apply(rule a'(1)) unfolding a' `a=a0` using `a3\<notin>s` 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  | 
thus ?thesis by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
751  | 
    ultimately have *:"?A = {s, insert a3 (s - {a0})}" by blast
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
752  | 
    have "s \<noteq> insert a3 (s - {a0})" using `a3\<notin>s` by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
753  | 
hence ?thesis unfolding * by auto } moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
754  | 
  { assume "a=a1"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
755  | 
have *:"\<And>P Q. (P \<or> Q) \<Longrightarrow> \<not> P \<Longrightarrow> Q" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
756  | 
have "\<exists>x\<in>s. \<not> kle n a1 x" apply(rule_tac x=a0 in bexI) proof assume as:"kle n a1 a0"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
757  | 
show False using kle_imp_pointwise[OF as,THEN spec[where x=1]] unfolding a0a1(5)[THEN spec[where x=1]]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
758  | 
using assms(3) by auto qed(insert a0a1,auto)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
759  | 
    hence "\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. a1 j = (if j = k then y j + 1 else y j)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
760  | 
apply(rule_tac *[OF ksimplex_predecessor[OF assms(1-2),unfolded `a=a1`]]) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
761  | 
then guess a2 .. from this(2) guess k .. note k=this note a2=`a2\<in>s`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
762  | 
def a3 \<equiv> "\<lambda>j. if j = k then a0 j - 1 else a0 j"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
763  | 
have "a2 \<noteq> a1" using k(2)[THEN spec[where x=k]] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
764  | 
    have lem3:"\<And>x. x\<in>(s - {a1}) \<Longrightarrow> kle n x a2" proof(rule ccontr) case goal1 hence as:"x\<in>s" "x\<noteq>a1" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
765  | 
have "kle n a2 x \<or> kle n x a2" using ksimplexD(6)[OF assms(1)] and as `a2\<in>s` by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
766  | 
have "kle n x a1" using a0a1(4) as by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
767  | 
ultimately have "x = a2 \<or> x = a1" apply-apply(rule kle_adjacent[OF k(2)]) using goal1(2) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
768  | 
hence "x = a2" using as by auto thus False using goal1(2) using kle_refl by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
769  | 
have "a0 k \<noteq> 0" proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
770  | 
      guess a4 using assms(4)[unfolded bex_simps ball_simps,rule_format,OF `k\<in>{1..n}`] .. note a4=this
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
771  | 
have "a4 k \<le> a2 k" using lem3[OF a4(1)[unfolded `a=a1`],THEN kle_imp_pointwise] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
772  | 
moreover have "a4 k > 0" using a4 by auto ultimately have "a2 k > 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
 | 
773  | 
hence "a1 k > 1" unfolding k(2)[rule_format] by simp  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
774  | 
thus ?thesis unfolding a0a1(5)[rule_format] using k(1) by simp qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
775  | 
hence lem4:"\<forall>j. a0 j = (if j=k then a3 j + 1 else a3 j)" unfolding a3_def by simp  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
776  | 
have "\<not> kle n a0 a3" apply(rule ccontr) unfolding not_not apply(drule kle_imp_pointwise)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
777  | 
unfolding lem4[rule_format] apply(erule_tac x=k in allE) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
778  | 
hence "a3 \<notin> s" using a0a1(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
 | 
779  | 
hence "a3 \<noteq> a1" "a3 \<noteq> a0" using a0a1 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
780  | 
    let ?s = "insert a3 (s - {a1})" have "ksimplex p n ?s" apply(rule ksimplexI) proof(rule_tac[2-] ballI,rule_tac[4] ballI)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
781  | 
show "card ?s = n+1" using ksimplexD(2-3)[OF assms(1)]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
782  | 
using `a3\<noteq>a0` `a3\<notin>s` `a1\<in>s` by(auto simp add:card_insert_if)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
783  | 
      fix x assume x:"x \<in> insert a3 (s - {a1})"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
784  | 
show "\<forall>j. x j \<le> p" proof(rule,cases "x = a3")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
785  | 
fix j case False thus "x j\<le>p" using x ksimplexD(4)[OF assms(1)] by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
786  | 
fix j case True show "x j\<le>p" unfolding True proof(cases "j=k")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
787  | 
case False thus "a3 j \<le>p" unfolding True a3_def using `a0\<in>s` ksimplexD(4)[OF assms(1)] by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
788  | 
guess a4 using assms(5)[unfolded bex_simps ball_simps,rule_format,OF k(1)] .. note a4=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
789  | 
case True have "a3 k \<le> a0 k" unfolding lem4[rule_format] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
790  | 
also have "\<dots> \<le> p" using ksimplexD(4)[OF assms(1),rule_format,of a0 k] a0a1 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
791  | 
finally show "a3 j \<le> p" unfolding True by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
792  | 
      show "\<forall>j. j \<notin> {1..n} \<longrightarrow> x j = p" proof(rule,rule,cases "x=a3") fix j::nat assume j:"j\<notin>{1..n}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
793  | 
	{ case False thus "x j = p" using j x ksimplexD(5)[OF assms(1)] by auto }
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
794  | 
case True show "x j = p" unfolding True a3_def using j k(1)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
795  | 
using ksimplexD(5)[OF assms(1),rule_format,OF `a0\<in>s` j] by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
796  | 
      fix y assume y:"y\<in>insert a3 (s - {a1})"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
797  | 
      have lem4:"\<And>x. x\<in>s \<Longrightarrow> x\<noteq>a1 \<Longrightarrow> kle n a3 x" proof- case goal1 hence *:"x\<in>s - {a1}" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
798  | 
have "kle n a3 a2" proof- have "kle n a0 a1" using a0a1 by auto then guess kk unfolding kle_def ..  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
799  | 
thus ?thesis unfolding kle_def apply(rule_tac x=kk in exI) unfolding lem4[rule_format] k(2)[rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
800  | 
apply(rule)defer proof(rule) case goal1 thus ?case apply-apply(erule conjE)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
801  | 
apply(erule_tac[!] x=j in allE) apply(cases "j\<in>kk") apply(case_tac[!] "j=k") by auto qed auto qed moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
802  | 
        have "kle n a3 a0" unfolding kle_def lem4[rule_format] apply(rule_tac x="{k}" in exI) using k(1) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
803  | 
ultimately show ?case apply-apply(rule kle_between_l[of _ a0 _ a2]) using lem3[OF *]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
804  | 
using a0a1(4)[rule_format,OF goal1(1)] by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
805  | 
show "kle n x y \<or> kle n y x" proof(cases "y=a3")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
806  | 
case True show ?thesis unfolding True apply(cases "x=a3") defer apply(rule disjI2,rule lem4)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
807  | 
using x by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
808  | 
case False show ?thesis proof(cases "x=a3") case True show ?thesis unfolding True  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
809  | 
apply(rule disjI1,rule lem4) using y False by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
810  | 
case False thus ?thesis apply(rule_tac ksimplexD(6)[OF assms(1),rule_format])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
811  | 
using x y `y\<noteq>a3` by auto qed qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
812  | 
    hence "insert a3 (s - {a1}) \<in> ?A" unfolding mem_Collect_eq apply-apply(rule,assumption)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
813  | 
apply(rule_tac x="a3" in bexI) unfolding `a=a1` using `a3\<notin>s` by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
814  | 
    have "s \<in> ?A" using assms(1,2) by auto ultimately have  "?A \<supseteq> {s, insert a3 (s - {a1})}" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
815  | 
    moreover have "?A \<subseteq> {s, insert a3 (s - {a1})}" apply(rule) unfolding mem_Collect_eq proof(erule conjE)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
816  | 
      fix s' assume as:"ksimplex p n s'" and "\<exists>b\<in>s'. s' - {b} = s - {a}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
817  | 
from this(2) guess a' .. note a'=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
818  | 
guess a_min a_max apply(rule ksimplex_extrema_strong[OF as assms(3)]) . note min_max=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
819  | 
      have *:"\<forall>x\<in>s' - {a'}. x k = a2 k" unfolding a' proof fix x assume x:"x\<in>s-{a}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
820  | 
hence "kle n x a2" apply-apply(rule lem3) using `a=a1` by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
821  | 
hence "x k \<le> a2 k" apply(drule_tac kle_imp_pointwise) by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
822  | 
	{ have "a2 k \<le> a0 k" using k(2)[rule_format,of k] unfolding a0a1(5)[rule_format] using k(1) by simp
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
823  | 
also have "\<dots> \<le> x k" using a0a1(4)[rule_format,of x,THEN conjunct1,THEN kle_imp_pointwise] x by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
824  | 
finally have "a2 k \<le> x k" . } ultimately show "x k = a2 k" by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
825  | 
have **:"a'=a_min \<or> a'=a_max" apply(rule ksimplex_fix_plane[OF a'(1) k(1) *]) using min_max by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
826  | 
have "a2 \<noteq> a1" proof assume as:"a2 = a1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
827  | 
show False using k(2) unfolding as apply(erule_tac x=k in allE) by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
828  | 
      hence a2':"a2 \<in> s' - {a'}" unfolding a' using a2 unfolding `a=a1` by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
829  | 
      show "s' \<in> {s, insert a3 (s - {a1})}" proof(cases "a'=a_min")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
830  | 
	case True have "a_max \<in> s - {a1}" using min_max unfolding a'(2)[unfolded `a=a1`,THEN sym] True by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
831  | 
hence "a_max = a2" unfolding kle_antisym[THEN sym,of a_max a2 n] apply-apply(rule)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
832  | 
apply(rule lem3,assumption) apply(rule min_max(4)[rule_format,THEN conjunct2]) using a2' by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
833  | 
hence a_max:"\<forall>i. a_max i = a2 i" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
834  | 
	have *:"\<forall>j. a2 j = (if j\<in>{1..n} then a3 j + 1 else a3 j)" 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
835  | 
using k(2) unfolding lem4[rule_format] a0a1(5)[rule_format] apply-apply(rule,erule_tac x=j in allE)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
836  | 
	proof- case goal1 thus ?case apply(cases "j\<in>{1..n}",case_tac[!] "j=k") by auto qed
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
837  | 
have "\<forall>i. a_min i = a3 i" using a_max apply-apply(rule,erule_tac x=i in allE)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
838  | 
unfolding min_max(5)[rule_format] *[rule_format] proof- case goal1  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
839  | 
	  thus ?case apply(cases "i\<in>{1..n}") by auto qed hence "a_min = a3" unfolding expand_fun_eq .
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
840  | 
	hence "s' = insert a3 (s - {a1})" using a' unfolding `a=a1` True by auto thus ?thesis by auto next
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
841  | 
case False hence as:"a'=a_max" using ** by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
842  | 
have "a_min = a0" unfolding kle_antisym[THEN sym,of _ _ n] apply(rule)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
843  | 
apply(rule min_max(4)[rule_format,THEN conjunct1]) defer apply(rule a0a1(4)[rule_format,THEN conjunct1]) proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
844  | 
	  have "a_min \<in> s - {a1}" using min_max(1,3) unfolding a'(2)[THEN sym,unfolded `a=a1`] as by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
845  | 
	  thus "a_min \<in> s" by auto have "a0 \<in> s - {a1}" using a0a1(1-3) by auto thus "a0 \<in> s'"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
846  | 
unfolding a'(2)[THEN sym,unfolded `a=a1`] by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
847  | 
hence "\<forall>i. a_max i = a1 i" unfolding a0a1(5)[rule_format] min_max(5)[rule_format] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
848  | 
hence "s' = s" apply-apply(rule lem1[OF a'(2)]) using `a\<in>s` `a'\<in>s'` unfolding as `a=a1` unfolding expand_fun_eq by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
849  | 
thus ?thesis by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
850  | 
    ultimately have *:"?A = {s, insert a3 (s - {a1})}" by blast
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
851  | 
    have "s \<noteq> insert a3 (s - {a1})" using `a3\<notin>s` by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
852  | 
hence ?thesis unfolding * by auto } moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
853  | 
  { assume as:"a\<noteq>a0" "a\<noteq>a1" have "\<not> (\<forall>x\<in>s. kle n a x)" proof case goal1
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
854  | 
have "a=a0" unfolding kle_antisym[THEN sym,of _ _ n] apply(rule)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
855  | 
using goal1 a0a1 assms(2) by auto thus False using as by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
856  | 
    hence "\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. a j = (if j = k then y j + 1 else y j)" using  ksimplex_predecessor[OF assms(1-2)] by blast
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
857  | 
then guess u .. from this(2) guess k .. note k = this[rule_format] note u = `u\<in>s`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
858  | 
have "\<not> (\<forall>x\<in>s. kle n x a)" proof case goal1  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
859  | 
have "a=a1" unfolding kle_antisym[THEN sym,of _ _ n] apply(rule)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
860  | 
using goal1 a0a1 assms(2) by auto thus False using as by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
861  | 
    hence "\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. y j = (if j = k then a j + 1 else a j)" using  ksimplex_successor[OF assms(1-2)] by blast
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
862  | 
then guess v .. from this(2) guess l .. note l = this[rule_format] note v = `v\<in>s`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
863  | 
def a' \<equiv> "\<lambda>j. if j = l then u j + 1 else u j"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
864  | 
have kl:"k \<noteq> l" proof assume "k=l" have *:"\<And>P. (if P then (1::nat) else 0) \<noteq> 2" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
865  | 
thus False using ksimplexD(6)[OF assms(1),rule_format,OF u v] unfolding kle_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
866  | 
unfolding l(2) k(2) `k=l` apply-apply(erule disjE)apply(erule_tac[!] exE conjE)+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
867  | 
apply(erule_tac[!] x=l in allE)+ by(auto simp add: *) qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
868  | 
hence aa':"a'\<noteq>a" apply-apply rule unfolding expand_fun_eq unfolding a'_def k(2)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
869  | 
apply(erule_tac x=l in allE) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
870  | 
have "a' \<notin> s" apply(rule) apply(drule ksimplexD(6)[OF assms(1),rule_format,OF `a\<in>s`]) proof(cases "kle n a a'")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
871  | 
case goal2 hence "kle n a' a" by auto thus False apply(drule_tac kle_imp_pointwise)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
872  | 
apply(erule_tac x=l in allE) unfolding a'_def k(2) using kl by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
873  | 
case True thus False apply(drule_tac kle_imp_pointwise)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
874  | 
apply(erule_tac x=k in allE) unfolding a'_def k(2) using kl by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
875  | 
have kle_uv:"kle n u a" "kle n u a'" "kle n a v" "kle n a' v" unfolding kle_def apply-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
876  | 
      apply(rule_tac[1] x="{k}" in exI,rule_tac[2] x="{l}" in exI)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
877  | 
      apply(rule_tac[3] x="{l}" in exI,rule_tac[4] x="{k}" in exI)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
878  | 
unfolding l(2) k(2) a'_def using l(1) k(1) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
879  | 
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)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
880  | 
proof- case goal1 thus ?case proof(cases "x k = u k", case_tac[!] "x l = u l")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
881  | 
assume as:"x l = u l" "x k = u k"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
882  | 
have "x = u" unfolding expand_fun_eq  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
883  | 
using goal1(2)[THEN kle_imp_pointwise,unfolded l(2)] unfolding k(2) apply-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
884  | 
using goal1(1)[THEN kle_imp_pointwise] apply-apply rule apply(erule_tac x=xa in allE)+ proof- case goal1  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
885  | 
thus ?case apply(cases "x=l") apply(case_tac[!] "x=k") using as by auto qed thus ?case by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
886  | 
assume as:"x l \<noteq> u l" "x k = u k"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
887  | 
have "x = a'" unfolding expand_fun_eq unfolding a'_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
888  | 
using goal1(2)[THEN kle_imp_pointwise] unfolding l(2) k(2) apply-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
889  | 
using goal1(1)[THEN kle_imp_pointwise] apply-apply rule apply(erule_tac x=xa in allE)+ proof- case goal1  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
890  | 
thus ?case apply(cases "x=l") apply(case_tac[!] "x=k") using as by auto qed thus ?case by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
891  | 
assume as:"x l = u l" "x k \<noteq> u k"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
892  | 
have "x = a" unfolding expand_fun_eq  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
893  | 
using goal1(2)[THEN kle_imp_pointwise] unfolding l(2) k(2) apply-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
894  | 
using goal1(1)[THEN kle_imp_pointwise] apply-apply rule apply(erule_tac x=xa in allE)+ proof- case goal1  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
895  | 
thus ?case apply(cases "x=l") apply(case_tac[!] "x=k") using as by auto qed thus ?case by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
896  | 
assume as:"x l \<noteq> u l" "x k \<noteq> u k"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
897  | 
have "x = v" unfolding expand_fun_eq  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
898  | 
using goal1(2)[THEN kle_imp_pointwise] unfolding l(2) k(2) apply-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
899  | 
using goal1(1)[THEN kle_imp_pointwise] apply-apply rule apply(erule_tac x=xa in allE)+ proof- case goal1  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
900  | 
thus ?case apply(cases "x=l") apply(case_tac[!] "x=k") using as `k\<noteq>l` by auto qed thus ?case by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
901  | 
have uv:"kle n u v" apply(rule kle_trans[OF kle_uv(1,3)]) using ksimplexD(6)[OF assms(1)] using u v by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
902  | 
have lem3:"\<And>x. x\<in>s \<Longrightarrow> kle n v x \<Longrightarrow> kle n a' x" apply(rule kle_between_r[of _ u _ v])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
903  | 
prefer 3 apply(rule kle_trans[OF uv]) defer apply(rule ksimplexD(6)[OF assms(1),rule_format])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
904  | 
using kle_uv `u\<in>s` by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
905  | 
have lem4:"\<And>x. x\<in>s \<Longrightarrow> kle n x u \<Longrightarrow> kle n x a'" apply(rule kle_between_l[of _ u _ v])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
906  | 
prefer 4 apply(rule kle_trans[OF _ uv]) defer apply(rule ksimplexD(6)[OF assms(1),rule_format])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
907  | 
using kle_uv `v\<in>s` by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
908  | 
have lem5:"\<And>x. x\<in>s \<Longrightarrow> x\<noteq>a \<Longrightarrow> kle n x a' \<or> kle n a' x" proof- case goal1 thus ?case  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
909  | 
proof(cases "kle n v x \<or> kle n x u") case True thus ?thesis using goal1 by(auto intro:lem3 lem4) next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
910  | 
case False hence *:"kle n u x" "kle n x v" using ksimplexD(6)[OF assms(1)] using goal1 `u\<in>s` `v\<in>s` by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
911  | 
show ?thesis using uxv[OF *] using kle_uv using goal1 by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
912  | 
    have "ksimplex p n (insert a' (s - {a}))" apply(rule ksimplexI) proof(rule_tac[2-] ballI,rule_tac[4] ballI)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
913  | 
      show "card (insert a' (s - {a})) = n + 1" using ksimplexD(2-3)[OF assms(1)]
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
914  | 
using `a'\<noteq>a` `a'\<notin>s` `a\<in>s` by(auto simp add:card_insert_if)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
915  | 
      fix x assume x:"x \<in> insert a' (s - {a})"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
916  | 
show "\<forall>j. x j \<le> p" proof(rule,cases "x = a'")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
917  | 
fix j case False thus "x j\<le>p" using x ksimplexD(4)[OF assms(1)] by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
918  | 
fix j case True show "x j\<le>p" unfolding True proof(cases "j=l")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
919  | 
case False thus "a' j \<le>p" unfolding True a'_def using `u\<in>s` ksimplexD(4)[OF assms(1)] by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
920  | 
case True have *:"a l = u l" "v l = a l + 1" using k(2)[of l] l(2)[of l] `k\<noteq>l` by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
921  | 
have "u l + 1 \<le> p" unfolding *[THEN sym] using ksimplexD(4)[OF assms(1)] using `v\<in>s` by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
922  | 
thus "a' j \<le>p" unfolding a'_def True by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
923  | 
      show "\<forall>j. j \<notin> {1..n} \<longrightarrow> x j = p" proof(rule,rule,cases "x=a'") fix j::nat assume j:"j\<notin>{1..n}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
924  | 
	{ case False thus "x j = p" using j x ksimplexD(5)[OF assms(1)] by auto }
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
925  | 
case True show "x j = p" unfolding True a'_def using j l(1)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
926  | 
using ksimplexD(5)[OF assms(1),rule_format,OF `u\<in>s` j] by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
927  | 
      fix y assume y:"y\<in>insert a' (s - {a})"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
928  | 
show "kle n x y \<or> kle n y x" proof(cases "y=a'")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
929  | 
case True show ?thesis unfolding True apply(cases "x=a'") defer apply(rule lem5) using x by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
930  | 
case False show ?thesis proof(cases "x=a'") case True show ?thesis unfolding True  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
931  | 
using lem5[of y] using y by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
932  | 
case False thus ?thesis apply(rule_tac ksimplexD(6)[OF assms(1),rule_format])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
933  | 
using x y `y\<noteq>a'` by auto qed qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
934  | 
    hence "insert a' (s - {a}) \<in> ?A" unfolding mem_Collect_eq apply-apply(rule,assumption)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
935  | 
apply(rule_tac x="a'" in bexI) using aa' `a'\<notin>s` by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
936  | 
    have "s \<in> ?A" using assms(1,2) by auto ultimately have  "?A \<supseteq> {s, insert a' (s - {a})}" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
937  | 
    moreover have "?A \<subseteq> {s, insert a' (s - {a})}" apply(rule) unfolding mem_Collect_eq proof(erule conjE)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
938  | 
      fix s' assume as:"ksimplex p n s'" and "\<exists>b\<in>s'. s' - {b} = s - {a}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
939  | 
from this(2) guess a'' .. note a''=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
940  | 
have "u\<noteq>v" unfolding expand_fun_eq unfolding l(2) k(2) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
941  | 
hence uv':"\<not> kle n v u" using uv using kle_antisym by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
942  | 
have "u\<noteq>a" "v\<noteq>a" unfolding expand_fun_eq k(2) l(2) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
943  | 
hence uvs':"u\<in>s'" "v\<in>s'" using `u\<in>s` `v\<in>s` using a'' by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
944  | 
have lem6:"a \<in> s' \<or> a' \<in> s'" proof(cases "\<forall>x\<in>s'. kle n x u \<or> kle n v x")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
945  | 
case False then guess w unfolding ball_simps .. note w=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
946  | 
hence "kle n u w" "kle n w v" using ksimplexD(6)[OF as] uvs' by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
947  | 
hence "w = a' \<or> w = a" using uxv[of w] uvs' w by auto thus ?thesis using w by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
948  | 
case True have "\<not> (\<forall>x\<in>s'. kle n x u)" unfolding ball_simps apply(rule_tac x=v in bexI)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
949  | 
using uv `u\<noteq>v` unfolding kle_antisym[of n u v,THEN sym] using `v\<in>s'` by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
950  | 
	hence "\<exists>y\<in>s'. \<exists>k\<in>{1..n}. \<forall>j. y j = (if j = k then u j + 1 else u j)" using ksimplex_successor[OF as `u\<in>s'`] by blast
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
951  | 
then guess w .. note w=this from this(2) guess kk .. note kk=this[rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
952  | 
have "\<not> kle n w u" apply-apply(rule,drule kle_imp_pointwise)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
953  | 
apply(erule_tac x=kk in allE) unfolding kk by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
954  | 
hence *:"kle n v w" using True[rule_format,OF w(1)] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
955  | 
hence False proof(cases "kk\<noteq>l") case True thus False using *[THEN kle_imp_pointwise, unfolded l(2) kk k(2)]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
956  | 
apply(erule_tac x=l in allE) using `k\<noteq>l` by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
957  | 
case False hence "kk\<noteq>k" using `k\<noteq>l` by auto thus False using *[THEN kle_imp_pointwise, unfolded l(2) kk k(2)]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
958  | 
apply(erule_tac x=k in allE) using `k\<noteq>l` by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
959  | 
thus ?thesis by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
960  | 
      thus "s' \<in> {s, insert a' (s - {a})}" proof(cases "a\<in>s'")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
961  | 
case True hence "s' = s" apply-apply(rule lem1[OF a''(2)]) using a'' `a\<in>s` 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  | 
thus ?thesis by auto next case False hence "a'\<in>s'" using lem6 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
963  | 
	hence "s' = insert a' (s - {a})" apply-apply(rule lem1[of _ a'' _ a'])
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
964  | 
unfolding a''(2)[THEN sym] using a'' using `a'\<notin>s` by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
965  | 
thus ?thesis by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
966  | 
    ultimately have *:"?A = {s, insert a' (s - {a})}" by blast
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
967  | 
    have "s \<noteq> insert a' (s - {a})" using `a'\<notin>s` by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
968  | 
hence ?thesis unfolding * by auto }  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
969  | 
ultimately show ?thesis by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
970  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
971  | 
subsection {* Hence another step towards concreteness. *}
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
972  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
973  | 
lemma kuhn_simplex_lemma:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
974  | 
  assumes "\<forall>s. ksimplex p (n + 1) s \<longrightarrow> (rl ` s \<subseteq>{0..n+1})"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
975  | 
  "odd (card{f. \<exists>s a. ksimplex p (n + 1) s \<and> a \<in> s \<and> (f = s - {a}) \<and>
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
976  | 
  (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))})"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
977  | 
  shows "odd(card {s\<in>{s. ksimplex p (n + 1) s}. rl ` s = {0..n+1} })" proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
978  | 
have *:"\<And>x y. x = y \<Longrightarrow> odd (card x) \<Longrightarrow> odd (card y)" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
979  | 
  have *:"odd(card {f\<in>{f. \<exists>s\<in>{s. ksimplex p (n + 1) s}. (\<exists>a\<in>s. f = s - {a})}. 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
980  | 
                (rl ` f = {0..n}) \<and>
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
981  | 
               ((\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = 0) \<or>
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
982  | 
                (\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = p))})" apply(rule *[OF _ assms(2)]) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
983  | 
show ?thesis apply(rule kuhn_complete_lemma[OF finite_simplices]) prefer 6 apply(rule *) apply(rule,rule,rule)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
984  | 
apply(subst ksimplex_def) defer apply(rule,rule assms(1)[rule_format]) unfolding mem_Collect_eq apply assumption  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
985  | 
apply default+ unfolding mem_Collect_eq apply(erule disjE bexE)+ defer apply(erule disjE bexE)+ defer  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
986  | 
apply default+ unfolding mem_Collect_eq apply(erule disjE bexE)+ unfolding mem_Collect_eq proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
987  | 
    fix f s a assume as:"ksimplex p (n + 1) s" "a\<in>s" "f = s - {a}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
988  | 
    let ?S = "{s. ksimplex p (n + 1) s \<and> (\<exists>a\<in>s. f = s - {a})}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
989  | 
    have S:"?S = {s'. ksimplex p (n + 1) s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})}" unfolding as by blast
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
990  | 
    { fix j assume j:"j \<in> {1..n + 1}" "\<forall>x\<in>f. x j = 0" thus "card {s. ksimplex p (n + 1) s \<and> (\<exists>a\<in>s. f = s - {a})} = 1" unfolding S
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
991  | 
apply-apply(rule ksimplex_replace_0) apply(rule as)+ unfolding as by auto }  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
992  | 
    { fix j assume j:"j \<in> {1..n + 1}" "\<forall>x\<in>f. x j = p" thus "card {s. ksimplex p (n + 1) s \<and> (\<exists>a\<in>s. f = s - {a})} = 1" unfolding S
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
993  | 
apply-apply(rule ksimplex_replace_1) apply(rule as)+ unfolding as by auto }  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
994  | 
    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"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
995  | 
unfolding S apply(rule ksimplex_replace_2) apply(rule as)+ unfolding as by auto qed auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
996  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
997  | 
subsection {* Reduced labelling. *}
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
998  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
999  | 
definition "reduced label (n::nat) (x::nat\<Rightarrow>nat) =  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1000  | 
(SOME k. k \<le> n \<and> (\<forall>i. 1\<le>i \<and> i<k+1 \<longrightarrow> label x i = 0) \<and> (k = n \<or> label x (k + 1) \<noteq> (0::nat)))"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1001  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1002  | 
lemma reduced_labelling: shows "reduced label n x \<le> n" (is ?t1) and  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1003  | 
"\<forall>i. 1\<le>i \<and> i < reduced label n x + 1 \<longrightarrow> (label x i = 0)" (is ?t2)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1004  | 
"(reduced label n x = n) \<or> (label x (reduced label n x + 1) \<noteq> 0)" (is ?t3) proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1005  | 
have num_WOP:"\<And>P k. P (k::nat) \<Longrightarrow> \<exists>n. P n \<and> (\<forall>m<n. \<not> P m)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1006  | 
apply(drule ex_has_least_nat[where m="\<lambda>x. x"]) apply(erule exE,rule_tac x=x in exI) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1007  | 
have *:"n \<le> n \<and> (label x (n + 1) \<noteq> 0 \<or> n = n)" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1008  | 
then guess N apply(drule_tac num_WOP[of "\<lambda>j. j\<le>n \<and> (label x (j+1) \<noteq> 0 \<or> n = j)"]) apply(erule exE) . note N=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1009  | 
have N':"N \<le> n" "\<forall>i. 1 \<le> i \<and> i < N + 1 \<longrightarrow> label x i = 0" "N = n \<or> label x (N + 1) \<noteq> 0" defer proof(rule,rule)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1010  | 
fix i assume i:"1\<le>i \<and> i<N+1" thus "label x i = 0" using N[THEN conjunct2,THEN spec[where x="i - 1"]] using N by auto qed(insert N, auto)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1011  | 
show ?t1 ?t2 ?t3 unfolding reduced_def apply(rule_tac[!] someI2_ex) using N' by(auto intro!: exI[where x=N]) qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1012  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1013  | 
lemma reduced_labelling_unique: fixes x::"nat \<Rightarrow> nat"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1014  | 
assumes "r \<le> n" "\<forall>i. 1 \<le> i \<and> i < r + 1 \<longrightarrow> (label x i = 0)" "(r = n) \<or> (label x (r + 1) \<noteq> 0)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1015  | 
shows "reduced label n x = r" apply(rule le_antisym) apply(rule_tac[!] ccontr) unfolding not_le  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1016  | 
using reduced_labelling[of label n x] 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
 | 
1017  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1018  | 
lemma reduced_labelling_0: assumes "j\<in>{1..n}" "label x j = 0" shows "reduced label n x \<noteq> j - 1"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1019  | 
using reduced_labelling[of label n x] using assms by fastsimp  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1020  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1021  | 
lemma reduced_labelling_1: assumes "j\<in>{1..n}" "label x j \<noteq> 0" shows "reduced label n x < j"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1022  | 
using assms and reduced_labelling[of label n x] apply(erule_tac x=j in allE) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1023  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1024  | 
lemma reduced_labelling_Suc:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1025  | 
assumes "reduced lab (n + 1) x \<noteq> n + 1" shows "reduced lab (n + 1) x = reduced lab n x"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1026  | 
apply(subst eq_commute) apply(rule reduced_labelling_unique)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1027  | 
using reduced_labelling[of lab "n+1" x] and 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
 | 
1028  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1029  | 
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
 | 
1030  | 
  assumes "\<forall>x\<in>f. \<forall>j\<in>{1..n+1}. x j = 0 \<longrightarrow> lab x j = 0"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1031  | 
          "\<forall>x\<in>f. \<forall>j\<in>{1..n+1}. x j = p \<longrightarrow> lab x j = 1"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1032  | 
  shows "((reduced lab (n + 1)) ` 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)) \<longleftrightarrow>
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1033  | 
  ((reduced lab (n + 1)) ` f = {0..n}) \<and> (\<forall>x\<in>f. x (n + 1) = p)" (is "?l = ?r") proof
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1034  | 
assume ?l (is "?as \<and> (?a \<or> ?b)") thus ?r apply-apply(rule,erule conjE,assumption) proof(cases ?a)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1035  | 
    case True then guess j .. note j=this {fix x assume x:"x\<in>f"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1036  | 
have "reduced lab (n+1) x \<noteq> j - 1" using j apply-apply(rule reduced_labelling_0) defer apply(rule assms(1)[rule_format]) using x by auto }  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1037  | 
    moreover have "j - 1 \<in> {0..n}" using j by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1038  | 
then guess y unfolding `?l`[THEN conjunct1,THEN sym] and image_iff .. note y = this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1039  | 
ultimately have False by auto thus "\<forall>x\<in>f. x (n + 1) = p" by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1040  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1041  | 
    case False hence ?b using `?l` by blast then guess j .. note j=this {fix x assume x:"x\<in>f"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1042  | 
have "reduced lab (n+1) x < j" using j apply-apply(rule reduced_labelling_1) using assms(2)[rule_format,of x j] and x by auto } note * = this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1043  | 
have "j = n + 1" proof(rule ccontr) case goal1 hence "j < n + 1" using j by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1044  | 
      have "n \<in> {0..n}" by auto then guess y unfolding `?l`[THEN conjunct1,THEN sym] image_iff ..
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1045  | 
ultimately show False using *[of y] by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1046  | 
thus "\<forall>x\<in>f. x (n + 1) = p" using j by auto qed qed(auto)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1047  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1048  | 
subsection {* Hence we get just about the nice induction. *}
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1049  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1050  | 
lemma kuhn_induction:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1051  | 
  assumes "0 < p" "\<forall>x. \<forall>j\<in>{1..n+1}. (\<forall>j. x j \<le> p) \<and> (x j = 0) \<longrightarrow> (lab x j = 0)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1052  | 
                  "\<forall>x. \<forall>j\<in>{1..n+1}. (\<forall>j. x j \<le> p) \<and> (x j = p) \<longrightarrow> (lab x j = 1)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1053  | 
        "odd (card {f. ksimplex p n f \<and> ((reduced lab n) ` f = {0..n})})"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1054  | 
  shows "odd (card {s. ksimplex p (n+1) s \<and>((reduced lab (n+1)) `  s = {0..n+1})})" proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1055  | 
  have *:"\<And>s t. odd (card s) \<Longrightarrow> s = t \<Longrightarrow> odd (card t)" "\<And>s f. (\<And>x. f x \<le> n +1 ) \<Longrightarrow> f ` s \<subseteq> {0..n+1}" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1056  | 
show ?thesis apply(rule kuhn_simplex_lemma[unfolded mem_Collect_eq]) apply(rule,rule,rule *,rule reduced_labelling)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1057  | 
apply(rule *(1)[OF assms(4)]) apply(rule set_ext) unfolding mem_Collect_eq apply(rule,erule conjE) defer apply(rule) proof-(*(rule,rule)*)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1058  | 
    fix f assume as:"ksimplex p n f" "reduced lab n ` f = {0..n}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1059  | 
    have *:"\<forall>x\<in>f. \<forall>j\<in>{1..n + 1}. x j = 0 \<longrightarrow> lab x j = 0" "\<forall>x\<in>f. \<forall>j\<in>{1..n + 1}. x j = p \<longrightarrow> lab x j = 1"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1060  | 
using assms(2-3) using as(1)[unfolded 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
 | 
1061  | 
have allp:"\<forall>x\<in>f. x (n + 1) = p" using assms(2) using as(1)[unfolded 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
 | 
1062  | 
    { fix x assume "x\<in>f" hence "reduced lab (n + 1) x < n + 1" apply-apply(rule reduced_labelling_1)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1063  | 
defer using assms(3) using as(1)[unfolded 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
 | 
1064  | 
hence "reduced lab (n + 1) x = reduced lab n x" apply-apply(rule reduced_labelling_Suc) using reduced_labelling(1) by auto }  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1065  | 
    hence "reduced lab (n + 1) ` f = {0..n}" unfolding as(2)[THEN sym] apply- apply(rule set_ext) unfolding image_iff by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1066  | 
moreover guess s using as(1)[unfolded simplex_top_face[OF assms(1) allp,THEN sym]] .. then guess a ..  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1067  | 
ultimately show "\<exists>s a. ksimplex p (n + 1) s \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1068  | 
      a \<in> s \<and> f = s - {a} \<and> reduced lab (n + 1) ` 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))" (is ?ex)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1069  | 
apply(rule_tac x=s in exI,rule_tac x=a in exI) unfolding complete_face_top[OF *] using allp as(1) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1070  | 
next fix f assume as:"\<exists>s a. ksimplex p (n + 1) s \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1071  | 
      a \<in> s \<and> f = s - {a} \<and> reduced lab (n + 1) ` 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))" (is ?ex)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1072  | 
then guess s .. then guess a apply-apply(erule exE,(erule conjE)+) . note sa=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1073  | 
    { fix x assume "x\<in>f" hence "reduced lab (n + 1) x \<in> reduced lab (n + 1) ` f" by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1074  | 
hence "reduced lab (n + 1) x < n + 1" using sa(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
 | 
1075  | 
hence "reduced lab (n + 1) x = reduced lab n x" apply-apply(rule reduced_labelling_Suc)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1076  | 
using reduced_labelling(1) by auto }  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1077  | 
    thus part1:"reduced lab n ` f = {0..n}" unfolding sa(4)[THEN sym] apply-apply(rule set_ext) unfolding image_iff by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1078  | 
    have *:"\<forall>x\<in>f. x (n + 1) = p" proof(cases "\<exists>j\<in>{1..n + 1}. \<forall>x\<in>f. x j = 0")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1079  | 
case True then guess j .. hence "\<And>x. x\<in>f \<Longrightarrow> reduced lab (n + 1) x \<noteq> j - 1" apply-apply(rule reduced_labelling_0) apply assumption  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1080  | 
apply(rule assms(2)[rule_format]) using sa(1)[unfolded ksimplex_def] unfolding sa by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1081  | 
      have "j - 1 \<in> {0..n}" using `j\<in>{1..n+1}` by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1082  | 
ultimately have False unfolding sa(4)[THEN sym] unfolding image_iff by fastsimp thus ?thesis by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1083  | 
      case False hence "\<exists>j\<in>{1..n + 1}. \<forall>x\<in>f. x j = p" using sa(5) by fastsimp then guess j .. note j=this
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1084  | 
thus ?thesis proof(cases "j = n+1")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1085  | 
	case False hence *:"j\<in>{1..n}" using j by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1086  | 
hence "\<And>x. x\<in>f \<Longrightarrow> reduced lab n x < j" apply(rule reduced_labelling_1) proof- fix x assume "x\<in>f"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1087  | 
hence "lab x j = 1" apply-apply(rule assms(3)[rule_format,OF j(1)])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1088  | 
using sa(1)[unfolded ksimplex_def] using j unfolding sa by auto thus "lab x j \<noteq> 0" by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1089  | 
	moreover have "j\<in>{0..n}" using * by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1090  | 
ultimately have False unfolding part1[THEN sym] using * unfolding image_iff by auto thus ?thesis by auto qed auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1091  | 
thus "ksimplex p n f" using as unfolding simplex_top_face[OF assms(1) *,THEN sym] by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1092  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1093  | 
lemma kuhn_induction_Suc:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1094  | 
  assumes "0 < p" "\<forall>x. \<forall>j\<in>{1..Suc n}. (\<forall>j. x j \<le> p) \<and> (x j = 0) \<longrightarrow> (lab x j = 0)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1095  | 
                  "\<forall>x. \<forall>j\<in>{1..Suc n}. (\<forall>j. x j \<le> p) \<and> (x j = p) \<longrightarrow> (lab x j = 1)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1096  | 
        "odd (card {f. ksimplex p n f \<and> ((reduced lab n) ` f = {0..n})})"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1097  | 
  shows "odd (card {s. ksimplex p (Suc n) s \<and>((reduced lab (Suc n)) `  s = {0..Suc n})})"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1098  | 
using assms unfolding Suc_eq_plus1 by(rule kuhn_induction)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1099  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1100  | 
subsection {* And so we get the final combinatorial result. *}
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1101  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1102  | 
lemma ksimplex_0: "ksimplex p 0 s \<longleftrightarrow> s = {(\<lambda>x. p)}" (is "?l = ?r") proof
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1103  | 
assume l:?l guess a using ksimplexD(3)[OF l, unfolded add_0] unfolding card_1_exists .. note a=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1104  | 
have "a = (\<lambda>x. p)" using ksimplexD(5)[OF l, rule_format, OF a(1)] by(rule,auto) thus ?r using a by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1105  | 
assume r:?r show ?l unfolding r ksimplex_eq by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1106  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1107  | 
lemma reduce_labelling_0[simp]: "reduced lab 0 x = 0" apply(rule reduced_labelling_unique) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1108  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1109  | 
lemma kuhn_combinatorial:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1110  | 
assumes "0 < p" "\<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)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1111  | 
"\<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)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1112  | 
  shows " odd (card {s. ksimplex p n s \<and> ((reduced lab n) ` s = {0..n})})" using assms proof(induct n)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1113  | 
  let ?M = "\<lambda>n. {s. ksimplex p n s \<and> ((reduced lab n) ` s = {0..n})}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1114  | 
  { case 0 have *:"?M 0 = {{(\<lambda>x. p)}}" unfolding ksimplex_0 by auto show ?case unfolding * by auto }
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1115  | 
case (Suc n) have "odd (card (?M n))" apply(rule Suc(1)[OF Suc(2)]) using Suc(3-) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1116  | 
thus ?case apply-apply(rule kuhn_induction_Suc) using Suc(2-) by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1117  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1118  | 
lemma kuhn_lemma: assumes "0 < (p::nat)" "0 < (n::nat)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1119  | 
  "\<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))"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1120  | 
  "\<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))"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1121  | 
  "\<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))"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1122  | 
  obtains q where "\<forall>i\<in>{1..n}. q i < p"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1123  | 
  "\<forall>i\<in>{1..n}. \<exists>r s. (\<forall>j\<in>{1..n}. q(j) \<le> r(j) \<and> r(j) \<le> q(j) + 1) \<and>
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1124  | 
                               (\<forall>j\<in>{1..n}. q(j) \<le> s(j) \<and> s(j) \<le> q(j) + 1) \<and>
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1125  | 
~(label r i = label s i)" proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1126  | 
  let ?A = "{s. ksimplex p n s \<and> reduced label n ` s = {0..n}}" have "n\<noteq>0" 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
 | 
1127  | 
have conjD:"\<And>P Q. P \<and> Q \<Longrightarrow> P" "\<And>P Q. P \<and> Q \<Longrightarrow> Q" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1128  | 
have "odd (card ?A)" apply(rule kuhn_combinatorial[of p n label]) 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
 | 
1129  | 
  hence "card ?A \<noteq> 0" apply-apply(rule ccontr) by auto hence "?A \<noteq> {}" unfolding card_eq_0_iff by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1130  | 
then obtain s where "s\<in>?A" by auto note s=conjD[OF this[unfolded mem_Collect_eq]]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1131  | 
guess a b apply(rule ksimplex_extrema_strong[OF s(1) `n\<noteq>0`]) . note ab=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1132  | 
  show ?thesis apply(rule that[of a]) proof(rule_tac[!] ballI) fix i assume "i\<in>{1..n}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1133  | 
hence "a i + 1 \<le> p" apply-apply(rule order_trans[of _ "b i"]) apply(subst ab(5)[THEN spec[where x=i]])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1134  | 
using s(1)[unfolded ksimplex_def] defer apply- apply(erule conjE)+ apply(drule_tac bspec[OF _ ab(2)])+ by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1135  | 
thus "a i < p" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1136  | 
case goal2 hence "i \<in> reduced label n ` s" using s by auto then guess u unfolding image_iff .. note u=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1137  | 
from goal2 have "i - 1 \<in> reduced label n ` s" using s by auto then guess v unfolding image_iff .. note v=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1138  | 
show ?case apply(rule_tac x=u in exI, rule_tac x=v in exI) apply(rule conjI) defer apply(rule conjI) defer 2 proof(rule_tac[1-2] ballI)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1139  | 
show "label u i \<noteq> label v i" using reduced_labelling[of label n u] reduced_labelling[of label n v]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1140  | 
unfolding u(2)[THEN sym] v(2)[THEN sym] using goal2 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1141  | 
      fix j assume j:"j\<in>{1..n}" show "a j \<le> u j \<and> u j \<le> a j + 1" "a j \<le> v j \<and> v j \<le> a j + 1"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1142  | 
using conjD[OF ab(4)[rule_format, OF u(1)]] and conjD[OF ab(4)[rule_format, OF v(1)]] apply-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1143  | 
apply(drule_tac[!] kle_imp_pointwise)+ apply(erule_tac[!] x=j in allE)+ unfolding ab(5)[rule_format] using j  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1144  | 
by auto qed qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1145  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1146  | 
subsection {* The main result for 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
 | 
1147  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1148  | 
lemma kuhn_labelling_lemma':  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1149  | 
assumes "(\<forall>x::nat\<Rightarrow>real. P x \<longrightarrow> P (f x))" "\<forall>x. P x \<longrightarrow> (\<forall>i::nat. Q i \<longrightarrow> 0 \<le> x i \<and> x i \<le> 1)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1150  | 
shows "\<exists>l. (\<forall>x i. l x i \<le> (1::nat)) \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1151  | 
(\<forall>x i. P x \<and> Q i \<and> (x i = 0) \<longrightarrow> (l x i = 0)) \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1152  | 
(\<forall>x i. P x \<and> Q i \<and> (x i = 1) \<longrightarrow> (l x i = 1)) \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1153  | 
(\<forall>x i. P x \<and> Q i \<and> (l x i = 0) \<longrightarrow> x i \<le> f(x) i) \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1154  | 
(\<forall>x i. P x \<and> Q i \<and> (l x i = 1) \<longrightarrow> f(x) i \<le> x i)" proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1155  | 
have and_forall_thm:"\<And>P Q. (\<forall>x. P x) \<and> (\<forall>x. Q x) \<longleftrightarrow> (\<forall>x. P x \<and> Q x)" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1156  | 
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)" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1157  | 
show ?thesis unfolding and_forall_thm apply(subst choice_iff[THEN sym])+ proof(rule,rule) case goal1  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1158  | 
let ?R = "\<lambda>y. (P x \<and> Q xa \<and> x xa = 0 \<longrightarrow> y = (0::nat)) \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1159  | 
(P x \<and> Q xa \<and> x xa = 1 \<longrightarrow> y = 1) \<and> (P x \<and> Q xa \<and> y = 0 \<longrightarrow> x xa \<le> (f x) xa) \<and> (P x \<and> Q xa \<and> y = 1 \<longrightarrow> (f x) xa \<le> x xa)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1160  | 
    { assume "P x" "Q xa" hence "0 \<le> (f x) xa \<and> (f x) xa \<le> 1" using assms(2)[rule_format,of "f x" xa]
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1161  | 
apply(drule_tac assms(1)[rule_format]) by auto }  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1162  | 
hence "?R 0 \<or> ?R 1" by auto thus ?case by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1163  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1164  | 
lemma brouwer_cube: fixes f::"real^'n \<Rightarrow> real^'n"  | 
| 
33741
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1165  | 
  assumes "continuous_on {0..1} f" "f ` {0..1} \<subseteq> {0..1}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1166  | 
  shows "\<exists>x\<in>{0..1}. f x = x" apply(rule ccontr) proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1167  | 
  def n \<equiv> "CARD('n)" have n:"1 \<le> n" "0 < n" "n \<noteq> 0" unfolding n_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
 | 
1168  | 
  assume "\<not> (\<exists>x\<in>{0..1}. f x = x)" hence *:"\<not> (\<exists>x\<in>{0..1}. f x - x = 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
 | 
1169  | 
guess d apply(rule brouwer_compactness_lemma[OF compact_interval _ *])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1170  | 
apply(rule continuous_on_intros assms)+ . note d=this[rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1171  | 
  have *:"\<forall>x. x \<in> {0..1} \<longrightarrow> f x \<in> {0..1}"  "\<forall>x. x \<in> {0..1::real^'n} \<longrightarrow> (\<forall>i. True \<longrightarrow> 0 \<le> x $ i \<and> x $ i \<le> 1)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1172  | 
using assms(2)[unfolded image_subset_iff Ball_def] unfolding mem_interval by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1173  | 
guess label using kuhn_labelling_lemma[OF *] apply-apply(erule exE,(erule conjE)+) . note label = this[rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1174  | 
  have lem1:"\<forall>x\<in>{0..1}.\<forall>y\<in>{0..1}.\<forall>i. label x i \<noteq> label y i
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1175  | 
\<longrightarrow> abs(f x $ i - x $ i) \<le> norm(f y - f x) + norm(y - x)" proof(rule,rule,rule,rule)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1176  | 
    fix x y assume xy:"x\<in>{0..1::real^'n}" "y\<in>{0..1::real^'n}" fix i::'n assume i:"label x i \<noteq> label y i"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1177  | 
have *:"\<And>x y fx fy::real. (x \<le> fx \<and> fy \<le> y \<or> fx \<le> x \<and> y \<le> fy)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1178  | 
\<Longrightarrow> abs(fx - x) \<le> abs(fy - fx) + abs(y - x)" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1179  | 
have "\<bar>(f x - x) $ i\<bar> \<le> abs((f y - f x)$i) + abs((y - x)$i)" unfolding vector_minus_component  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1180  | 
apply(rule *) apply(cases "label x i = 0") apply(rule disjI1,rule) prefer 3 proof(rule disjI2,rule)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1181  | 
assume lx:"label x i = 0" hence ly:"label y i = 1" using i label(1)[of y i] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1182  | 
show "x $ i \<le> f x $ i" apply(rule label(4)[rule_format]) using xy lx by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1183  | 
show "f y $ i \<le> y $ i" apply(rule label(5)[rule_format]) using xy ly by auto next  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1184  | 
assume "label x i \<noteq> 0" hence l:"label x i = 1" "label y i = 0"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1185  | 
using i label(1)[of x i] label(1)[of y i] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1186  | 
show "f x $ i \<le> x $ i" apply(rule label(5)[rule_format]) using xy l by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1187  | 
show "y $ i \<le> f y $ i" apply(rule label(4)[rule_format]) using xy l by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1188  | 
also have "\<dots> \<le> norm (f y - f x) + norm (y - x)" apply(rule add_mono) by(rule component_le_norm)+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1189  | 
finally show "\<bar>f x $ i - x $ i\<bar> \<le> norm (f y - f x) + norm (y - x)" unfolding vector_minus_component . qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1190  | 
  have "\<exists>e>0. \<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. \<forall>z\<in>{0..1}. \<forall>i.
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1191  | 
norm(x - z) < e \<and> norm(y - z) < e \<and> label x i \<noteq> label y i \<longrightarrow> abs((f(z) - z)$i) < d / (real n)" proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1192  | 
have d':"d / real n / 8 > 0" apply(rule divide_pos_pos)+ using d(1) unfolding n_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
 | 
1193  | 
    have *:"uniformly_continuous_on {0..1} f" by(rule compact_uniformly_continuous[OF assms(1) compact_interval])
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1194  | 
guess e using *[unfolded uniformly_continuous_on_def,rule_format,OF d'] apply-apply(erule exE,(erule conjE)+) .  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1195  | 
note e=this[rule_format,unfolded vector_dist_norm]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1196  | 
show ?thesis apply(rule_tac x="min (e/2) (d/real n/8)" in exI) apply(rule) defer  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1197  | 
apply(rule,rule,rule,rule,rule) apply(erule conjE)+ proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1198  | 
show "0 < min (e / 2) (d / real n / 8)" using d' e by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1199  | 
      fix x y z i assume as:"x \<in> {0..1}" "y \<in> {0..1}" "z \<in> {0..1}" "norm (x - z) < min (e / 2) (d / real n / 8)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1200  | 
"norm (y - z) < min (e / 2) (d / real n / 8)" "label x i \<noteq> label y i"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1201  | 
have *:"\<And>z fz x fx n1 n2 n3 n4 d4 d::real. abs(fx - x) \<le> n1 + n2 \<Longrightarrow> abs(fx - fz) \<le> n3 \<Longrightarrow> abs(x - z) \<le> n4 \<Longrightarrow>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1202  | 
n1 < d4 \<Longrightarrow> n2 < 2 * d4 \<Longrightarrow> n3 < d4 \<Longrightarrow> n4 < d4 \<Longrightarrow> (8 * d4 = d) \<Longrightarrow> abs(fz - z) < d" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1203  | 
show "\<bar>(f z - z) $ i\<bar> < d / real n" unfolding vector_minus_component proof(rule *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1204  | 
show "\<bar>f x $ i - x $ i\<bar> \<le> norm (f y -f x) + norm (y - x)" apply(rule lem1[rule_format]) using as by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1205  | 
show "\<bar>f x $ i - f z $ i\<bar> \<le> norm (f x - f z)" "\<bar>x $ i - z $ i\<bar> \<le> norm (x - z)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1206  | 
unfolding vector_minus_component[THEN sym] by(rule component_le_norm)+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1207  | 
have tria:"norm (y - x) \<le> norm (y - z) + norm (x - z)" using dist_triangle[of y x z,unfolded vector_dist_norm]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1208  | 
unfolding norm_minus_commute by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1209  | 
also have "\<dots> < e / 2 + e / 2" apply(rule add_strict_mono) using as(4,5) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1210  | 
finally show "norm (f y - f x) < d / real n / 8" apply- apply(rule e(2)) using as by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1211  | 
have "norm (y - z) + norm (x - z) < d / real n / 8 + d / real n / 8" apply(rule add_strict_mono) using as by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1212  | 
thus "norm (y - x) < 2 * (d / real n / 8)" using tria by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1213  | 
show "norm (f x - f z) < d / real n / 8" apply(rule e(2)) using as e(1) by auto qed(insert as, auto) qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1214  | 
then guess e apply-apply(erule exE,(erule conjE)+) . note e=this[rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1215  | 
guess p using real_arch_simple[of "1 + real n / e"] .. note p=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1216  | 
have "1 + real n / e > 0" apply(rule add_pos_pos) defer apply(rule divide_pos_pos) using e(1) n by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1217  | 
hence "p > 0" using p by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1218  | 
guess b using ex_bij_betw_nat_finite_1[OF finite_UNIV[where 'a='n]] .. note b=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1219  | 
  def b' \<equiv> "inv_into {1..n} b"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1220  | 
  have b':"bij_betw b' UNIV {1..n}" using bij_betw_inv_into[OF b] unfolding b'_def n_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
 | 
1221  | 
have bb'[simp]:"\<And>i. b (b' i) = i" unfolding b'_def apply(rule f_inv_into_f) unfolding n_def using b  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1222  | 
unfolding bij_betw_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
 | 
1223  | 
  have b'b[simp]:"\<And>i. i\<in>{1..n} \<Longrightarrow> b' (b i) = i" unfolding b'_def apply(rule inv_into_f_eq)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1224  | 
using b unfolding n_def bij_betw_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
 | 
1225  | 
have *:"\<And>x::nat. x=0 \<or> x=1 \<longleftrightarrow> x\<le>1" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1226  | 
  have q1:"0 < p" "0 < n"  "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow>
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1227  | 
    (\<forall>i\<in>{1..n}. (label (\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 0 \<or> (label (\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 1)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1228  | 
unfolding * using `p>0` `n>0` using label(1) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1229  | 
  have q2:"\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. x i = 0 \<longrightarrow> (label (\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 0)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1230  | 
    "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. x i = p \<longrightarrow> (label (\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 1)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1231  | 
apply(rule,rule,rule,rule) defer proof(rule,rule,rule,rule) fix x i  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1232  | 
    assume as:"\<forall>i\<in>{1..n}. x i \<le> p" "i \<in> {1..n}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1233  | 
    { assume "x i = p \<or> x i = 0" 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1234  | 
      have "(\<chi> i. real (x (b' i)) / real p) \<in> {0..1}" unfolding mem_interval Cart_lambda_beta proof(rule,rule)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1235  | 
        fix j::'n have j:"b' j \<in> {1..n}" using b' unfolding n_def bij_betw_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
 | 
1236  | 
show "0 $ j \<le> real (x (b' j)) / real p" unfolding zero_index  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1237  | 
apply(rule divide_nonneg_pos) using `p>0` using as(1)[rule_format,OF j] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1238  | 
show "real (x (b' j)) / real p \<le> 1 $ j" unfolding one_index divide_le_eq_1  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1239  | 
using as(1)[rule_format,OF j] by auto qed } note cube=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1240  | 
    { assume "x i = p" thus "(label (\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 1" unfolding o_def
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1241  | 
apply-apply(rule label(3)) using cube using as `p>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
 | 
1242  | 
    { assume "x i = 0" thus "(label (\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 0" unfolding o_def
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1243  | 
apply-apply(rule label(2)) using cube using as `p>0` by auto } qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1244  | 
guess q apply(rule kuhn_lemma[OF q1 q2]) . note q=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1245  | 
def z \<equiv> "\<chi> i. real (q (b' i)) / real p"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1246  | 
have "\<exists>i. d / real n \<le> abs((f z - z)$i)" proof(rule ccontr)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1247  | 
    have "\<forall>i. q (b' i) \<in> {0..<p}" using q(1) b'[unfolded bij_betw_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
 | 
1248  | 
    hence "\<forall>i. q (b' i) \<in> {0..p}" apply-apply(rule,erule_tac x=i in allE) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1249  | 
    hence "z\<in>{0..1}" unfolding z_def mem_interval unfolding one_index zero_index Cart_lambda_beta
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1250  | 
apply-apply(rule,rule) apply(rule divide_nonneg_pos) using `p>0` unfolding divide_le_eq_1 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1251  | 
hence d_fz_z:"d \<le> norm (f z - z)" apply(drule_tac d) .  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1252  | 
case goal1 hence as:"\<forall>i. \<bar>f z $ i - z $ i\<bar> < d / real n" using `n>0` by(auto simp add:not_le)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1253  | 
have "norm (f z - z) \<le> (\<Sum>i\<in>UNIV. \<bar>f z $ i - z $ i\<bar>)" unfolding vector_minus_component[THEN sym] by(rule norm_le_l1)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1254  | 
also have "\<dots> < (\<Sum>(i::'n)\<in>UNIV. d / real n)" apply(rule setsum_strict_mono) using as by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1255  | 
also have "\<dots> = d" unfolding real_eq_of_nat n_def using n by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1256  | 
finally show False using d_fz_z by auto qed then guess i .. note i=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1257  | 
  have *:"b' i \<in> {1..n}" using b'[unfolded bij_betw_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
 | 
1258  | 
guess r using q(2)[rule_format,OF *] .. then guess s apply-apply(erule exE,(erule conjE)+) . note rs=this[rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1259  | 
  have b'_im:"\<And>i. b' i \<in> {1..n}" using b' unfolding bij_betw_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
 | 
1260  | 
def r' \<equiv> "\<chi> i. real (r (b' i)) / real p"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1261  | 
have "\<And>i. r (b' i) \<le> p" apply(rule order_trans) apply(rule rs(1)[OF b'_im,THEN conjunct2])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1262  | 
using q(1)[rule_format,OF b'_im] by(auto simp add: Suc_le_eq)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1263  | 
  hence "r' \<in> {0..1::real^'n}" unfolding r'_def mem_interval Cart_lambda_beta one_index zero_index
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1264  | 
apply-apply(rule,rule,rule divide_nonneg_pos)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1265  | 
using rs(1)[OF b'_im] q(1)[rule_format,OF b'_im] `p>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
 | 
1266  | 
def s' \<equiv> "\<chi> i. real (s (b' i)) / real p"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1267  | 
have "\<And>i. s (b' i) \<le> p" apply(rule order_trans) apply(rule rs(2)[OF b'_im,THEN conjunct2])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1268  | 
using q(1)[rule_format,OF b'_im] by(auto simp add: Suc_le_eq)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1269  | 
  hence "s' \<in> {0..1::real^'n}" unfolding s'_def mem_interval Cart_lambda_beta one_index zero_index
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1270  | 
apply-apply(rule,rule,rule divide_nonneg_pos)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1271  | 
using rs(1)[OF b'_im] q(1)[rule_format,OF b'_im] `p>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
 | 
1272  | 
  have "z\<in>{0..1}" unfolding z_def mem_interval Cart_lambda_beta one_index zero_index 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1273  | 
apply(rule,rule,rule divide_nonneg_pos) using q(1)[rule_format,OF b'_im] `p>0` by(auto intro:less_imp_le)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1274  | 
have *:"\<And>x. 1 + real x = real (Suc x)" by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1275  | 
  { have "(\<Sum>i\<in>UNIV. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'n)\<in>UNIV. 1)" 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1276  | 
apply(rule setsum_mono) using rs(1)[OF b'_im] by(auto simp add:* field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1277  | 
also have "\<dots> < e * real p" using p `e>0` `p>0` unfolding n_def real_of_nat_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1278  | 
by(auto simp add:field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1279  | 
finally have "(\<Sum>i\<in>UNIV. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) < e * real p" . } moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1280  | 
  { have "(\<Sum>i\<in>UNIV. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'n)\<in>UNIV. 1)" 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1281  | 
apply(rule setsum_mono) using rs(2)[OF b'_im] by(auto simp add:* field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1282  | 
also have "\<dots> < e * real p" using p `e>0` `p>0` unfolding n_def real_of_nat_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1283  | 
by(auto simp add:field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1284  | 
finally have "(\<Sum>i\<in>UNIV. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) < e * real p" . } ultimately  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1285  | 
have "norm (r' - z) < e" "norm (s' - z) < e" unfolding r'_def s'_def z_def apply-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1286  | 
apply(rule_tac[!] le_less_trans[OF norm_le_l1]) using `p>0`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1287  | 
by(auto simp add:field_simps setsum_divide_distrib[THEN sym])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1288  | 
  hence "\<bar>(f z - z) $ i\<bar> < d / real n" apply-apply(rule e(2)[OF `r'\<in>{0..1}` `s'\<in>{0..1}` `z\<in>{0..1}`])
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1289  | 
using rs(3) unfolding r'_def[symmetric] s'_def[symmetric] o_def bb' by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1290  | 
thus False using i by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1291  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1292  | 
subsection {* Retractions. *}
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1293  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1294  | 
definition "retraction s t (r::real^'n\<Rightarrow>real^'n) \<longleftrightarrow>  | 
| 
33741
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1295  | 
t \<subseteq> s \<and> continuous_on s r \<and> (r ` s \<subseteq> t) \<and> (\<forall>x\<in>t. r x = x)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1296  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1297  | 
definition retract_of (infixl "retract'_of" 12) where  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1298  | 
"(t retract_of s) \<longleftrightarrow> (\<exists>r. retraction s t r)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1299  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1300  | 
lemma retraction_idempotent: "retraction s t r \<Longrightarrow> x \<in> s \<Longrightarrow> r(r x) = r x"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1301  | 
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
 | 
1302  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1303  | 
subsection {*preservation of fixpoints under (more general notion of) retraction. *}
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1304  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1305  | 
lemma invertible_fixpoint_property: fixes s::"(real^'n) set" and t::"(real^'m) set"  | 
| 
33741
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1306  | 
assumes "continuous_on t i" "i ` t \<subseteq> s" "continuous_on s r" "r ` s \<subseteq> t" "\<forall>y\<in>t. r (i y) = y"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1307  | 
"\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)" "continuous_on t g" "g ` t \<subseteq> t"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1308  | 
obtains y where "y\<in>t" "g y = y" proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1309  | 
have "\<exists>x\<in>s. (i \<circ> g \<circ> r) x = x" apply(rule assms(6)[rule_format],rule)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1310  | 
apply(rule continuous_on_compose assms)+ apply((rule continuous_on_subset)?,rule assms)+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1311  | 
using assms(2,4,8) unfolding image_compose by(auto,blast)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1312  | 
then guess x .. note x = this hence *:"g (r x) \<in> t" using assms(4,8) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1313  | 
have "r ((i \<circ> g \<circ> r) x) = r x" using x by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1314  | 
thus ?thesis apply(rule_tac that[of "r x"]) using x unfolding o_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1315  | 
unfolding assms(5)[rule_format,OF *] using assms(4) by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1316  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1317  | 
lemma homeomorphic_fixpoint_property:  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1318  | 
fixes s::"(real^'n) set" and t::"(real^'m) set" 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
 | 
1319  | 
shows "(\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)) \<longleftrightarrow>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1320  | 
(\<forall>g. continuous_on t g \<and> g ` t \<subseteq> t \<longrightarrow> (\<exists>y\<in>t. g y = y))" proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1321  | 
guess r using assms[unfolded homeomorphic_def homeomorphism_def] .. then guess i ..  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1322  | 
thus ?thesis apply- apply rule apply(rule_tac[!] allI impI)+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1323  | 
apply(rule_tac g=g in invertible_fixpoint_property[of t i s r]) prefer 10  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1324  | 
apply(rule_tac g=f in invertible_fixpoint_property[of s r t i]) by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1325  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1326  | 
lemma retract_fixpoint_property:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1327  | 
assumes "t retract_of s" "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)" "continuous_on t g" "g ` t \<subseteq> t"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1328  | 
obtains y where "y \<in> t" "g y = y" proof- guess h using assms(1) unfolding retract_of_def ..  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1329  | 
thus ?thesis unfolding retraction_def apply-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1330  | 
apply(rule invertible_fixpoint_property[OF continuous_on_id _ _ _ _ assms(2), of t h g]) prefer 7  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1331  | 
apply(rule_tac y=y in that) using assms by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1332  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1333  | 
subsection {*So the Brouwer theorem for any set with nonempty interior. *}
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1334  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1335  | 
lemma brouwer_weak: fixes f::"real^'n \<Rightarrow> real^'n"  | 
| 
33741
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1336  | 
  assumes "compact s" "convex s" "interior s \<noteq> {}" "continuous_on s f" "f ` s \<subseteq> s"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1337  | 
obtains x where "x \<in> s" "f x = x" proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1338  | 
  have *:"interior {0..1::real^'n} \<noteq> {}" unfolding interior_closed_interval interval_eq_empty by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1339  | 
  have *:"{0..1::real^'n} homeomorphic s" using homeomorphic_convex_compact[OF convex_interval(1) compact_interval * assms(2,1,3)] .
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1340  | 
  have "\<forall>f. continuous_on {0..1} f \<and> f ` {0..1} \<subseteq> {0..1} \<longrightarrow> (\<exists>x\<in>{0..1::real^'n}. f x = x)" using brouwer_cube by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1341  | 
thus ?thesis unfolding homeomorphic_fixpoint_property[OF *] apply(erule_tac x=f in allE)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1342  | 
apply(erule impE) defer apply(erule bexE) apply(rule_tac x=y in that) using assms by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1343  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1344  | 
subsection {* And in particular for a closed ball. *}
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1345  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1346  | 
lemma brouwer_ball: fixes f::"real^'n \<Rightarrow> real^'n"  | 
| 
33741
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1347  | 
assumes "0 < e" "continuous_on (cball a e) f" "f ` (cball a e) \<subseteq> (cball a e)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1348  | 
obtains x where "x \<in> cball a e" "f x = x"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1349  | 
using brouwer_weak[OF compact_cball convex_cball,of a e f] unfolding interior_cball ball_eq_empty  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1350  | 
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
 | 
1351  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1352  | 
text {*Still more general form; could derive this directly without using the 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1353  | 
rather involved HOMEOMORPHIC_CONVEX_COMPACT theorem, just using  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1354  | 
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
 | 
1355  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1356  | 
lemma brouwer: fixes f::"real^'n \<Rightarrow> real^'n"  | 
| 
33741
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1357  | 
  assumes "compact s" "convex s" "s \<noteq> {}" "continuous_on s f" "f ` s \<subseteq> s"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1358  | 
obtains x where "x \<in> s" "f x = x" proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1359  | 
have "\<exists>e>0. s \<subseteq> cball 0 e" using compact_imp_bounded[OF assms(1)] unfolding bounded_pos  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1360  | 
apply(erule_tac exE,rule_tac x=b in exI) by(auto simp add: vector_dist_norm)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1361  | 
then guess e apply-apply(erule exE,(erule conjE)+) . note e=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1362  | 
have "\<exists>x\<in> cball 0 e. (f \<circ> closest_point s) x = x"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1363  | 
apply(rule_tac brouwer_ball[OF e(1), of 0 "f \<circ> closest_point s"]) apply(rule continuous_on_compose )  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1364  | 
apply(rule continuous_on_closest_point[OF assms(2) compact_imp_closed[OF assms(1)] assms(3)])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1365  | 
apply(rule continuous_on_subset[OF assms(4)])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1366  | 
using closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)] apply - defer  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1367  | 
using assms(5)[unfolded subset_eq] using e(2)[unfolded subset_eq mem_cball] by(auto simp add:vector_dist_norm)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1368  | 
then guess x .. note x=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1369  | 
have *:"closest_point s x = x" apply(rule closest_point_self)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1370  | 
apply(rule assms(5)[unfolded subset_eq,THEN bspec[where x="x"],unfolded image_iff])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1371  | 
apply(rule_tac x="closest_point s x" in bexI) using x unfolding o_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1372  | 
using closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3), of x] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1373  | 
show thesis apply(rule_tac x="closest_point s x" in that) unfolding x(2)[unfolded o_def]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1374  | 
apply(rule closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)]) using * by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1375  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1376  | 
text {*So we get the no-retraction theorem. *}                                      
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1377  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1378  | 
lemma no_retraction_cball: assumes "0 < e"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1379  | 
shows "\<not> (frontier(cball a e) retract_of (cball a e))" proof case goal1  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1380  | 
have *:"\<And>xa. a - (2 *\<^sub>R a - xa) = -(a - xa)" using scaleR_left_distrib[of 1 1 a] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1381  | 
guess x apply(rule retract_fixpoint_property[OF goal1, of "\<lambda>x. scaleR 2 a - x"])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1382  | 
apply(rule,rule,erule conjE) apply(rule brouwer_ball[OF assms]) apply assumption+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1383  | 
apply(rule_tac x=x in bexI) apply assumption+ apply(rule continuous_on_intros)+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1384  | 
unfolding frontier_cball subset_eq Ball_def image_iff apply(rule,rule,erule bexE)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1385  | 
unfolding vector_dist_norm apply(simp add: * norm_minus_commute) . note x = this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1386  | 
hence "scaleR 2 a = scaleR 1 x + scaleR 1 x" by(auto simp add:group_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1387  | 
hence "a = x" unfolding scaleR_left_distrib[THEN sym] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1388  | 
thus False using x using assms by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1389  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1390  | 
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
 | 
1391  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1392  | 
definition "interval_bij = (\<lambda> (a,b) (u,v) (x::real^'n).  | 
| 
33741
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1393  | 
(\<chi> i. u$i + (x$i - a$i) / (b$i - a$i) * (v$i - u$i))::real^'n)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1394  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1395  | 
lemma interval_bij_affine:  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1396  | 
"interval_bij (a,b) (u,v) = (\<lambda>x. (\<chi> i. (v$i - u$i) / (b$i - a$i) * x$i) +  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1397  | 
(\<chi> i. u$i - (v$i - u$i) / (b$i - a$i) * a$i))"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1398  | 
apply rule unfolding Cart_eq interval_bij_def vector_component_simps  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1399  | 
by(auto simp add:group_simps field_simps add_divide_distrib[THEN sym])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1400  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1401  | 
lemma continuous_interval_bij:  | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1402  | 
"continuous (at x) (interval_bij (a,b::real^'n) (u,v))"  | 
| 
33741
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1403  | 
unfolding interval_bij_affine apply(rule continuous_intros)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1404  | 
apply(rule linear_continuous_at) unfolding linear_conv_bounded_linear[THEN sym]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1405  | 
unfolding linear_def unfolding Cart_eq unfolding Cart_lambda_beta defer  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1406  | 
apply(rule continuous_intros) by(auto simp add:field_simps add_divide_distrib[THEN sym])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1407  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1408  | 
lemma continuous_on_interval_bij: "continuous_on s (interval_bij (a,b) (u,v))"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1409  | 
apply(rule continuous_at_imp_continuous_on) by(rule, rule continuous_interval_bij)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1410  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1411  | 
(** move this **)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1412  | 
lemma divide_nonneg_nonneg:assumes "a \<ge> 0" "b \<ge> 0" shows "0 \<le> a / (b::real)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1413  | 
apply(cases "b=0") defer apply(rule divide_nonneg_pos) 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
 | 
1414  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1415  | 
lemma in_interval_interval_bij: assumes "x \<in> {a..b}" "{u..v} \<noteq> {}"
 | 
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1416  | 
  shows "interval_bij (a,b) (u,v) x \<in> {u..v::real^'n}" 
 | 
| 
33741
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1417  | 
unfolding interval_bij_def split_conv mem_interval Cart_lambda_beta proof(rule,rule)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1418  | 
  fix i::'n have "{a..b} \<noteq> {}" 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
 | 
1419  | 
hence *:"a$i \<le> b$i" "u$i \<le> v$i" using assms(2) unfolding interval_eq_empty not_ex apply-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1420  | 
apply(erule_tac[!] x=i in allE)+ by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1421  | 
have x:"a$i\<le>x$i" "x$i\<le>b$i" using assms(1)[unfolded mem_interval] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1422  | 
have "0 \<le> (x $ i - a $ i) / (b $ i - a $ i) * (v $ i - u $ i)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1423  | 
apply(rule mult_nonneg_nonneg) apply(rule divide_nonneg_nonneg)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1424  | 
using * x by(auto simp add:field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1425  | 
thus "u $ i \<le> u $ i + (x $ i - a $ i) / (b $ i - a $ i) * (v $ i - u $ i)" using * by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1426  | 
have "((x $ i - a $ i) / (b $ i - a $ i)) * (v $ i - u $ i) \<le> 1 * (v $ i - u $ i)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1427  | 
apply(rule mult_right_mono) unfolding divide_le_eq_1 using * x by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1428  | 
thus "u $ i + (x $ i - a $ i) / (b $ i - a $ i) * (v $ i - u $ i) \<le> v $ i" using * by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1429  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1430  | 
lemma interval_bij_bij: assumes "\<forall>i. a$i < b$i \<and> u$i < v$i"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1431  | 
shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1432  | 
unfolding interval_bij_def split_conv Cart_eq Cart_lambda_beta  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1433  | 
apply(rule,insert assms,erule_tac x=i in allE) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1434  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1435  | 
subsection {*Fashoda meet theorem. *}
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1436  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1437  | 
lemma infnorm_2: "infnorm (x::real^2) = max (abs(x$1)) (abs(x$2))"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1438  | 
unfolding infnorm_def UNIV_2 apply(rule Sup_eq) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1439  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1440  | 
lemma infnorm_eq_1_2: "infnorm (x::real^2) = 1 \<longleftrightarrow>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1441  | 
(abs(x$1) \<le> 1 \<and> abs(x$2) \<le> 1 \<and> (x$1 = -1 \<or> x$1 = 1 \<or> x$2 = -1 \<or> x$2 = 1))"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1442  | 
unfolding infnorm_2 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1443  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1444  | 
lemma infnorm_eq_1_imp: assumes "infnorm (x::real^2) = 1" shows "abs(x$1) \<le> 1" "abs(x$2) \<le> 1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1445  | 
using assms unfolding infnorm_eq_1_2 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1446  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1447  | 
lemma fashoda_unit: fixes f g::"real^1 \<Rightarrow> real^2"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1448  | 
  assumes "f ` {- 1..1} \<subseteq> {- 1..1}" "g ` {- 1..1} \<subseteq> {- 1..1}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1449  | 
  "continuous_on {- 1..1} f"  "continuous_on {- 1..1} g"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1450  | 
"f (- 1)$1 = - 1" "f 1$1 = 1" "g (- 1) $2 = -1" "g 1 $2 = 1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1451  | 
  shows "\<exists>s\<in>{- 1..1}. \<exists>t\<in>{- 1..1}. f s = g t" proof(rule ccontr)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1452  | 
case goal1 note as = this[unfolded bex_simps,rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1453  | 
def sqprojection \<equiv> "\<lambda>z::real^2. (inverse (infnorm z)) *\<^sub>R z"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1454  | 
def negatex \<equiv> "\<lambda>x::real^2. (vector [-(x$1), x$2])::real^2"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1455  | 
have lem1:"\<forall>z::real^2. infnorm(negatex z) = infnorm z"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1456  | 
unfolding negatex_def infnorm_2 vector_2 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1457  | 
have lem2:"\<forall>z. z\<noteq>0 \<longrightarrow> infnorm(sqprojection z) = 1" unfolding sqprojection_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1458  | 
unfolding infnorm_mul[unfolded smult_conv_scaleR] unfolding abs_inverse real_abs_infnorm  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1459  | 
unfolding infnorm_eq_0[THEN sym] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1460  | 
let ?F = "(\<lambda>w::real^2. (f \<circ> vec1 \<circ> (\<lambda>x. x$1)) w - (g \<circ> vec1 \<circ> (\<lambda>x. x$2)) w)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1461  | 
  have *:"\<And>i. vec1 ` (\<lambda>x::real^2. x $ i) ` {- 1..1} = {- 1..1::real^1}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1462  | 
apply(rule set_ext) unfolding image_iff Bex_def mem_interval apply rule defer  | 
| 34964 | 1463  | 
apply(rule_tac x="dest_vec1 x" in exI) apply rule apply(rule_tac x="vec (dest_vec1 x)" in exI) 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
 | 
1464  | 
  { fix x assume "x \<in> (\<lambda>w. (f \<circ> vec1 \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> vec1 \<circ> (\<lambda>x. x $ 2)) w) ` {- 1..1::real^2}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1465  | 
then guess w unfolding image_iff .. note w = this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1466  | 
hence "x \<noteq> 0" using as[of "vec1 (w$1)" "vec1 (w$2)"] unfolding mem_interval by auto} note x0=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1467  | 
have 21:"\<And>i::2. i\<noteq>1 \<Longrightarrow> i=2" using UNIV_2 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1468  | 
  have 1:"{- 1<..<1::real^2} \<noteq> {}" unfolding interval_eq_empty by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1469  | 
  have 2:"continuous_on {- 1..1} (negatex \<circ> sqprojection \<circ> ?F)" apply(rule continuous_on_intros continuous_on_component continuous_on_vec1)+
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1470  | 
prefer 2 apply(rule continuous_on_intros continuous_on_component continuous_on_vec1)+ unfolding *  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1471  | 
apply(rule assms)+ apply(rule continuous_on_compose,subst sqprojection_def)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1472  | 
apply(rule continuous_on_mul ) apply(rule continuous_at_imp_continuous_on,rule) apply(rule continuous_at_inv[unfolded o_def])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1473  | 
apply(rule continuous_at_infnorm) unfolding infnorm_eq_0 defer apply(rule continuous_on_id) apply(rule linear_continuous_on) proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1474  | 
show "bounded_linear negatex" apply(rule bounded_linearI') unfolding Cart_eq proof(rule_tac[!] allI) fix i::2 and x y::"real^2" and c::real  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1475  | 
show "negatex (x + y) $ i = (negatex x + negatex y) $ i" "negatex (c *s x) $ i = (c *s negatex x) $ i"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1476  | 
apply-apply(case_tac[!] "i\<noteq>1") prefer 3 apply(drule_tac[1-2] 21)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1477  | 
unfolding negatex_def by(auto simp add:vector_2 ) qed qed(insert x0, auto)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1478  | 
  have 3:"(negatex \<circ> sqprojection \<circ> ?F) ` {- 1..1} \<subseteq> {- 1..1}" unfolding subset_eq apply rule proof-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1479  | 
case goal1 then guess y unfolding image_iff .. note y=this have "?F y \<noteq> 0" apply(rule x0) using y(1) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1480  | 
hence *:"infnorm (sqprojection (?F y)) = 1" unfolding y o_def apply- by(rule lem2[rule_format])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1481  | 
have "infnorm x = 1" unfolding *[THEN sym] y o_def by(rule lem1[rule_format])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1482  | 
    thus "x\<in>{- 1..1}" unfolding mem_interval infnorm_2 apply- apply rule
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1483  | 
proof-case goal1 thus ?case apply(cases "i=1") defer apply(drule 21) by auto qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1484  | 
  guess x apply(rule brouwer_weak[of "{- 1..1::real^2}" "negatex \<circ> sqprojection \<circ> ?F"])
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1485  | 
apply(rule compact_interval convex_interval)+ unfolding interior_closed_interval  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1486  | 
apply(rule 1 2 3)+ . note x=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1487  | 
have "?F x \<noteq> 0" apply(rule x0) using x(1) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1488  | 
hence *:"infnorm (sqprojection (?F x)) = 1" unfolding o_def by(rule lem2[rule_format])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1489  | 
have nx:"infnorm x = 1" apply(subst x(2)[THEN sym]) unfolding *[THEN sym] o_def by(rule lem1[rule_format])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1490  | 
have "\<forall>x i. x \<noteq> 0 \<longrightarrow> (0 < (sqprojection x)$i \<longleftrightarrow> 0 < x$i)" "\<forall>x i. x \<noteq> 0 \<longrightarrow> ((sqprojection x)$i < 0 \<longleftrightarrow> x$i < 0)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1491  | 
apply- apply(rule_tac[!] allI impI)+ proof- fix x::"real^2" and i::2 assume x:"x\<noteq>0"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1492  | 
have "inverse (infnorm x) > 0" using x[unfolded infnorm_pos_lt[THEN sym]] by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1493  | 
thus "(0 < sqprojection x $ i) = (0 < x $ i)" "(sqprojection x $ i < 0) = (x $ i < 0)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1494  | 
unfolding sqprojection_def vector_component_simps Cart_nth.scaleR real_scaleR_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1495  | 
unfolding zero_less_mult_iff mult_less_0_iff by(auto simp add:field_simps) qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1496  | 
note lem3 = this[rule_format]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1497  | 
  have x1:"vec1 (x $ 1) \<in> {- 1..1::real^1}" "vec1 (x $ 2) \<in> {- 1..1::real^1}" using x(1) unfolding mem_interval by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1498  | 
hence nz:"f (vec1 (x $ 1)) - g (vec1 (x $ 2)) \<noteq> 0" unfolding right_minus_eq apply-apply(rule as) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1499  | 
have "x $ 1 = -1 \<or> x $ 1 = 1 \<or> x $ 2 = -1 \<or> x $ 2 = 1" using nx unfolding infnorm_eq_1_2 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1500  | 
thus False proof- fix P Q R S  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1501  | 
presume "P \<or> Q \<or> R \<or> S" "P\<Longrightarrow>False" "Q\<Longrightarrow>False" "R\<Longrightarrow>False" "S\<Longrightarrow>False" thus False by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1502  | 
next assume as:"x$1 = 1" hence "vec1 (x$1) = 1" unfolding Cart_eq by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1503  | 
hence *:"f (vec1 (x $ 1)) $ 1 = 1" using assms(6) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1504  | 
have "sqprojection (f (vec1 (x$1)) - g (vec1 (x$2))) $ 1 < 0"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1505  | 
using x(2)[unfolded o_def Cart_eq,THEN spec[where x=1]]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1506  | 
unfolding as negatex_def vector_2 by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1507  | 
    from x1 have "g (vec1 (x $ 2)) \<in> {- 1..1}" apply-apply(rule assms(2)[unfolded subset_eq,rule_format]) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1508  | 
ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1509  | 
apply(erule_tac x=1 in allE) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1510  | 
next assume as:"x$1 = -1" hence "vec1 (x$1) = - 1" unfolding Cart_eq by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1511  | 
hence *:"f (vec1 (x $ 1)) $ 1 = - 1" using assms(5) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1512  | 
have "sqprojection (f (vec1 (x$1)) - g (vec1 (x$2))) $ 1 > 0"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1513  | 
using x(2)[unfolded o_def Cart_eq,THEN spec[where x=1]]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1514  | 
unfolding as negatex_def vector_2 by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1515  | 
    from x1 have "g (vec1 (x $ 2)) \<in> {- 1..1}" apply-apply(rule assms(2)[unfolded subset_eq,rule_format]) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1516  | 
ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1517  | 
apply(erule_tac x=1 in allE) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1518  | 
next assume as:"x$2 = 1" hence "vec1 (x$2) = 1" unfolding Cart_eq by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1519  | 
hence *:"g (vec1 (x $ 2)) $ 2 = 1" using assms(8) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1520  | 
have "sqprojection (f (vec1 (x$1)) - g (vec1 (x$2))) $ 2 > 0"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1521  | 
using x(2)[unfolded o_def Cart_eq,THEN spec[where x=2]]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1522  | 
unfolding as negatex_def vector_2 by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1523  | 
    from x1 have "f (vec1 (x $ 1)) \<in> {- 1..1}" apply-apply(rule assms(1)[unfolded subset_eq,rule_format]) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1524  | 
ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1525  | 
apply(erule_tac x=2 in allE) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1526  | 
next assume as:"x$2 = -1" hence "vec1 (x$2) = - 1" unfolding Cart_eq by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1527  | 
hence *:"g (vec1 (x $ 2)) $ 2 = - 1" using assms(7) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1528  | 
have "sqprojection (f (vec1 (x$1)) - g (vec1 (x$2))) $ 2 < 0"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1529  | 
using x(2)[unfolded o_def Cart_eq,THEN spec[where x=2]]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1530  | 
unfolding as negatex_def vector_2 by auto moreover  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1531  | 
    from x1 have "f (vec1 (x $ 1)) \<in> {- 1..1}" apply-apply(rule assms(1)[unfolded subset_eq,rule_format]) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1532  | 
ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1533  | 
apply(erule_tac x=2 in allE) by auto qed(auto) qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1534  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1535  | 
lemma fashoda_unit_path: fixes f ::"real^1 \<Rightarrow> real^2" and g ::"real^1 \<Rightarrow> real^2"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1536  | 
  assumes "path f" "path g" "path_image f \<subseteq> {- 1..1}" "path_image g \<subseteq> {- 1..1}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1537  | 
"(pathstart f)$1 = -1" "(pathfinish f)$1 = 1" "(pathstart g)$2 = -1" "(pathfinish g)$2 = 1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1538  | 
obtains z where "z \<in> path_image f" "z \<in> path_image g" proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1539  | 
note assms=assms[unfolded path_def pathstart_def pathfinish_def path_image_def]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1540  | 
def iscale \<equiv> "\<lambda>z::real^1. inverse 2 *\<^sub>R (z + 1)"  | 
| 34964 | 1541  | 
  have isc:"iscale ` {- 1..1} \<subseteq> {0..1}" unfolding iscale_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
 | 
1542  | 
  have "\<exists>s\<in>{- 1..1}. \<exists>t\<in>{- 1..1}. (f \<circ> iscale) s = (g \<circ> iscale) t" proof(rule fashoda_unit) 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1543  | 
    show "(f \<circ> iscale) ` {- 1..1} \<subseteq> {- 1..1}" "(g \<circ> iscale) ` {- 1..1} \<subseteq> {- 1..1}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1544  | 
using isc and assms(3-4) unfolding image_compose by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1545  | 
    have *:"continuous_on {- 1..1} iscale" unfolding iscale_def by(rule continuous_on_intros)+
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1546  | 
    show "continuous_on {- 1..1} (f \<circ> iscale)" "continuous_on {- 1..1} (g \<circ> iscale)"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1547  | 
apply-apply(rule_tac[!] continuous_on_compose[OF *]) apply(rule_tac[!] continuous_on_subset[OF _ isc])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1548  | 
by(rule assms)+ have *:"(1 / 2) *\<^sub>R (1 + (1::real^1)) = 1" unfolding Cart_eq by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1549  | 
show "(f \<circ> iscale) (- 1) $ 1 = - 1" "(f \<circ> iscale) 1 $ 1 = 1" "(g \<circ> iscale) (- 1) $ 2 = -1" "(g \<circ> iscale) 1 $ 2 = 1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1550  | 
unfolding o_def iscale_def using assms by(auto simp add:*) qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1551  | 
then guess s .. from this(2) guess t .. note st=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1552  | 
show thesis apply(rule_tac z="f (iscale s)" in that)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1553  | 
    using st `s\<in>{- 1..1}` unfolding o_def path_image_def image_iff apply-
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1554  | 
apply(rule_tac x="iscale s" in bexI) prefer 3 apply(rule_tac x="iscale t" in bexI)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1555  | 
using isc[unfolded subset_eq, rule_format] by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1556  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1557  | 
lemma fashoda: fixes b::"real^2"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1558  | 
  assumes "path f" "path g" "path_image f \<subseteq> {a..b}" "path_image g \<subseteq> {a..b}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1559  | 
"(pathstart f)$1 = a$1" "(pathfinish f)$1 = b$1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1560  | 
"(pathstart g)$2 = a$2" "(pathfinish g)$2 = b$2"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1561  | 
obtains z where "z \<in> path_image f" "z \<in> path_image g" proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1562  | 
fix P Q S presume "P \<or> Q \<or> S" "P \<Longrightarrow> thesis" "Q \<Longrightarrow> thesis" "S \<Longrightarrow> thesis" thus thesis by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1563  | 
next have "{a..b} \<noteq> {}" using assms(3) using path_image_nonempty by auto
 | 
| 
33758
 
53078b0d21f5
Renamed vector_less_eq_def to the more usual name vector_le_def.
 
hoelzl 
parents: 
33741 
diff
changeset
 | 
1564  | 
hence "a \<le> b" unfolding interval_eq_empty vector_le_def by(auto simp add: not_less)  | 
| 
 
53078b0d21f5
Renamed vector_less_eq_def to the more usual name vector_le_def.
 
hoelzl 
parents: 
33741 
diff
changeset
 | 
1565  | 
thus "a$1 = b$1 \<or> a$2 = b$2 \<or> (a$1 < b$1 \<and> a$2 < b$2)" unfolding vector_le_def forall_2 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
 | 
1566  | 
next assume as:"a$1 = b$1" have "\<exists>z\<in>path_image g. z$2 = (pathstart f)$2" apply(rule connected_ivt_component)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1567  | 
apply(rule connected_path_image assms)+apply(rule pathstart_in_path_image,rule pathfinish_in_path_image)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1568  | 
unfolding assms using assms(3)[unfolded path_image_def subset_eq,rule_format,of "f 0"]  | 
| 
33758
 
53078b0d21f5
Renamed vector_less_eq_def to the more usual name vector_le_def.
 
hoelzl 
parents: 
33741 
diff
changeset
 | 
1569  | 
unfolding pathstart_def by(auto simp add: vector_le_def) then guess z .. note z=this  | 
| 
33741
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1570  | 
  have "z \<in> {a..b}" using z(1) assms(4) unfolding path_image_def by blast 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1571  | 
hence "z = f 0" unfolding Cart_eq forall_2 unfolding z(2) pathstart_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1572  | 
using assms(3)[unfolded path_image_def subset_eq mem_interval,rule_format,of "f 0" 1]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1573  | 
unfolding mem_interval apply(erule_tac x=1 in allE) using as by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1574  | 
thus thesis apply-apply(rule that[OF _ z(1)]) unfolding path_image_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
 | 
1575  | 
next assume as:"a$2 = b$2" have "\<exists>z\<in>path_image f. z$1 = (pathstart g)$1" apply(rule connected_ivt_component)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1576  | 
apply(rule connected_path_image assms)+apply(rule pathstart_in_path_image,rule pathfinish_in_path_image)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1577  | 
unfolding assms using assms(4)[unfolded path_image_def subset_eq,rule_format,of "g 0"]  | 
| 
33758
 
53078b0d21f5
Renamed vector_less_eq_def to the more usual name vector_le_def.
 
hoelzl 
parents: 
33741 
diff
changeset
 | 
1578  | 
unfolding pathstart_def by(auto simp add: vector_le_def) then guess z .. note z=this  | 
| 
33741
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1579  | 
  have "z \<in> {a..b}" using z(1) assms(3) unfolding path_image_def by blast 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1580  | 
hence "z = g 0" unfolding Cart_eq forall_2 unfolding z(2) pathstart_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1581  | 
using assms(4)[unfolded path_image_def subset_eq mem_interval,rule_format,of "g 0" 2]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1582  | 
unfolding mem_interval apply(erule_tac x=2 in allE) using as by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1583  | 
thus thesis apply-apply(rule that[OF z(1)]) unfolding path_image_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
 | 
1584  | 
next assume as:"a $ 1 < b $ 1 \<and> a $ 2 < b $ 2"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1585  | 
  have int_nem:"{- 1..1::real^2} \<noteq> {}" unfolding interval_eq_empty by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1586  | 
guess z apply(rule fashoda_unit_path[of "interval_bij (a,b) (- 1,1) \<circ> f" "interval_bij (a,b) (- 1,1) \<circ> g"])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1587  | 
unfolding path_def path_image_def pathstart_def pathfinish_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1588  | 
apply(rule_tac[1-2] continuous_on_compose) apply(rule assms[unfolded path_def] continuous_on_interval_bij)+  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1589  | 
unfolding subset_eq apply(rule_tac[1-2] ballI)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1590  | 
  proof- fix x assume "x \<in> (interval_bij (a, b) (- 1, 1) \<circ> f) ` {0..1}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1591  | 
then guess y unfolding image_iff .. note y=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1592  | 
    show "x \<in> {- 1..1}" unfolding y o_def apply(rule in_interval_interval_bij)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1593  | 
using y(1) using assms(3)[unfolded path_image_def subset_eq] int_nem by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1594  | 
  next fix x assume "x \<in> (interval_bij (a, b) (- 1, 1) \<circ> g) ` {0..1}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1595  | 
then guess y unfolding image_iff .. note y=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1596  | 
    show "x \<in> {- 1..1}" unfolding y o_def apply(rule in_interval_interval_bij)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1597  | 
using y(1) using assms(4)[unfolded path_image_def subset_eq] int_nem by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1598  | 
next show "(interval_bij (a, b) (- 1, 1) \<circ> f) 0 $ 1 = -1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1599  | 
"(interval_bij (a, b) (- 1, 1) \<circ> f) 1 $ 1 = 1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1600  | 
"(interval_bij (a, b) (- 1, 1) \<circ> g) 0 $ 2 = -1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1601  | 
"(interval_bij (a, b) (- 1, 1) \<circ> g) 1 $ 2 = 1" unfolding interval_bij_def Cart_lambda_beta vector_component_simps o_def split_conv  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1602  | 
unfolding assms[unfolded pathstart_def pathfinish_def] using as by auto qed note z=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1603  | 
from z(1) guess zf unfolding image_iff .. note zf=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1604  | 
from z(2) guess zg unfolding image_iff .. note zg=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1605  | 
have *:"\<forall>i. (- 1) $ i < (1::real^2) $ i \<and> a $ i < b $ i" unfolding forall_2 using as by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1606  | 
show thesis apply(rule_tac z="interval_bij (- 1,1) (a,b) z" in that)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1607  | 
apply(subst zf) defer apply(subst zg) unfolding o_def interval_bij_bij[OF *] path_image_def  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1608  | 
using zf(1) zg(1) by auto qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1609  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1610  | 
subsection {*Some slightly ad hoc lemmas I use below*}
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1611  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1612  | 
lemma segment_vertical: fixes a::"real^2" assumes "a$1 = b$1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1613  | 
shows "x \<in> closed_segment a b \<longleftrightarrow> (x$1 = a$1 \<and> x$1 = b$1 \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1614  | 
(a$2 \<le> x$2 \<and> x$2 \<le> b$2 \<or> b$2 \<le> x$2 \<and> x$2 \<le> a$2))" (is "_ = ?R")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1615  | 
proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1616  | 
let ?L = "\<exists>u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \<and> x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \<and> 0 \<le> u \<and> u \<le> 1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1617  | 
  { presume "?L \<Longrightarrow> ?R" "?R \<Longrightarrow> ?L" thus ?thesis unfolding closed_segment_def mem_Collect_eq
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1618  | 
unfolding Cart_eq forall_2 smult_conv_scaleR[THEN sym] vector_component_simps by blast }  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1619  | 
  { assume ?L then guess u apply-apply(erule exE)apply(erule conjE)+ . note u=this
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1620  | 
    { fix b a assume "b + u * a > a + u * b"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1621  | 
hence "(1 - u) * b > (1 - u) * a" by(auto simp add:field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1622  | 
hence "b \<ge> a" apply(drule_tac mult_less_imp_less_left) using u by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1623  | 
hence "u * a \<le> u * b" apply-apply(rule mult_left_mono[OF _ u(3)])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1624  | 
using u(3-4) by(auto simp add:field_simps) } note * = this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1625  | 
    { fix a b assume "u * b > u * a" hence "(1 - u) * a \<le> (1 - u) * b" apply-apply(rule mult_left_mono)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1626  | 
apply(drule mult_less_imp_less_left) using u by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1627  | 
hence "a + u * b \<le> b + u * a" by(auto simp add:field_simps) } note ** = this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1628  | 
thus ?R unfolding u assms using u by(auto simp add:field_simps not_le intro:* **) }  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1629  | 
  { assume ?R thus ?L proof(cases "x$2 = b$2")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1630  | 
case True thus ?L apply(rule_tac x="(x$2 - a$2) / (b$2 - a$2)" in exI) unfolding assms True  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1631  | 
using `?R` by(auto simp add:field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1632  | 
next case False thus ?L apply(rule_tac x="1 - (x$2 - b$2) / (a$2 - b$2)" in exI) unfolding assms using `?R`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1633  | 
by(auto simp add:field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1634  | 
qed } qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1635  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1636  | 
lemma segment_horizontal: fixes a::"real^2" assumes "a$2 = b$2"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1637  | 
shows "x \<in> closed_segment a b \<longleftrightarrow> (x$2 = a$2 \<and> x$2 = b$2 \<and>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1638  | 
(a$1 \<le> x$1 \<and> x$1 \<le> b$1 \<or> b$1 \<le> x$1 \<and> x$1 \<le> a$1))" (is "_ = ?R")  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1639  | 
proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1640  | 
let ?L = "\<exists>u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \<and> x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \<and> 0 \<le> u \<and> u \<le> 1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1641  | 
  { presume "?L \<Longrightarrow> ?R" "?R \<Longrightarrow> ?L" thus ?thesis unfolding closed_segment_def mem_Collect_eq
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1642  | 
unfolding Cart_eq forall_2 smult_conv_scaleR[THEN sym] vector_component_simps by blast }  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1643  | 
  { assume ?L then guess u apply-apply(erule exE)apply(erule conjE)+ . note u=this
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1644  | 
    { fix b a assume "b + u * a > a + u * b"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1645  | 
hence "(1 - u) * b > (1 - u) * a" by(auto simp add:field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1646  | 
hence "b \<ge> a" apply(drule_tac mult_less_imp_less_left) using u by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1647  | 
hence "u * a \<le> u * b" apply-apply(rule mult_left_mono[OF _ u(3)])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1648  | 
using u(3-4) by(auto simp add:field_simps) } note * = this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1649  | 
    { fix a b assume "u * b > u * a" hence "(1 - u) * a \<le> (1 - u) * b" apply-apply(rule mult_left_mono)
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1650  | 
apply(drule mult_less_imp_less_left) using u by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1651  | 
hence "a + u * b \<le> b + u * a" by(auto simp add:field_simps) } note ** = this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1652  | 
thus ?R unfolding u assms using u by(auto simp add:field_simps not_le intro:* **) }  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1653  | 
  { assume ?R thus ?L proof(cases "x$1 = b$1")
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1654  | 
case True thus ?L apply(rule_tac x="(x$1 - a$1) / (b$1 - a$1)" in exI) unfolding assms True  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1655  | 
using `?R` by(auto simp add:field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1656  | 
next case False thus ?L apply(rule_tac x="1 - (x$1 - b$1) / (a$1 - b$1)" in exI) unfolding assms using `?R`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1657  | 
by(auto simp add:field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1658  | 
qed } qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1659  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1660  | 
subsection {*useful Fashoda corollary pointed out to me by Tom Hales. *}
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1661  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1662  | 
lemma fashoda_interlace: fixes a::"real^2"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1663  | 
assumes "path f" "path g"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1664  | 
  "path_image f \<subseteq> {a..b}" "path_image g \<subseteq> {a..b}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1665  | 
"(pathstart f)$2 = a$2" "(pathfinish f)$2 = a$2"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1666  | 
"(pathstart g)$2 = a$2" "(pathfinish g)$2 = a$2"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1667  | 
"(pathstart f)$1 < (pathstart g)$1" "(pathstart g)$1 < (pathfinish f)$1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1668  | 
"(pathfinish f)$1 < (pathfinish g)$1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1669  | 
obtains z where "z \<in> path_image f" "z \<in> path_image g"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1670  | 
proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1671  | 
  have "{a..b} \<noteq> {}" using path_image_nonempty using assms(3) by auto
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1672  | 
note ab=this[unfolded interval_eq_empty not_ex forall_2 not_less]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1673  | 
  have "pathstart f \<in> {a..b}" "pathfinish f \<in> {a..b}" "pathstart g \<in> {a..b}" "pathfinish g \<in> {a..b}"
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1674  | 
using pathstart_in_path_image pathfinish_in_path_image using assms(3-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
 | 
1675  | 
note startfin = this[unfolded mem_interval forall_2]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1676  | 
let ?P1 = "linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2]) +++  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1677  | 
linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f) +++ f +++  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1678  | 
linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2]) +++  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1679  | 
linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2])"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1680  | 
let ?P2 = "linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g) +++ g +++  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1681  | 
linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1]) +++  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1682  | 
linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1]) +++  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1683  | 
linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3])"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1684  | 
let ?a = "vector[a$1 - 2, a$2 - 3]"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1685  | 
let ?b = "vector[b$1 + 2, b$2 + 3]"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1686  | 
have P1P2:"path_image ?P1 = path_image (linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2])) \<union>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1687  | 
path_image (linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f)) \<union> path_image f \<union>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1688  | 
path_image (linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2])) \<union>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1689  | 
path_image (linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2]))"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1690  | 
"path_image ?P2 = path_image(linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g)) \<union> path_image g \<union>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1691  | 
path_image(linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1])) \<union>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1692  | 
path_image(linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1])) \<union>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1693  | 
path_image(linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3]))" using assms(1-2)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1694  | 
by(auto simp add: pathstart_join pathfinish_join path_image_join path_image_linepath path_join path_linepath)  | 
| 
33758
 
53078b0d21f5
Renamed vector_less_eq_def to the more usual name vector_le_def.
 
hoelzl 
parents: 
33741 
diff
changeset
 | 
1695  | 
  have abab: "{a..b} \<subseteq> {?a..?b}" by(auto simp add:vector_le_def forall_2 vector_2)
 | 
| 
33741
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1696  | 
guess z apply(rule fashoda[of ?P1 ?P2 ?a ?b])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1697  | 
unfolding pathstart_join pathfinish_join pathstart_linepath pathfinish_linepath vector_2 proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1698  | 
show "path ?P1" "path ?P2" using assms by(auto simp add: pathstart_join pathfinish_join path_join)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1699  | 
    have "path_image ?P1 \<subseteq> {?a .. ?b}" unfolding P1P2 path_image_linepath apply(rule Un_least)+ defer 3
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1700  | 
apply(rule_tac[1-4] convex_interval(1)[unfolded convex_contains_segment,rule_format])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1701  | 
unfolding mem_interval forall_2 vector_2 using ab startfin abab assms(3)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1702  | 
using assms(9-) unfolding assms by(auto simp add:field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1703  | 
    thus "path_image ?P1  \<subseteq> {?a .. ?b}" .
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1704  | 
    have "path_image ?P2 \<subseteq> {?a .. ?b}" unfolding P1P2 path_image_linepath apply(rule Un_least)+ defer 2
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1705  | 
apply(rule_tac[1-4] convex_interval(1)[unfolded convex_contains_segment,rule_format])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1706  | 
unfolding mem_interval forall_2 vector_2 using ab startfin abab assms(4)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1707  | 
using assms(9-) unfolding assms by(auto simp add:field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1708  | 
    thus "path_image ?P2  \<subseteq> {?a .. ?b}" . 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1709  | 
show "a $ 1 - 2 = a $ 1 - 2" "b $ 1 + 2 = b $ 1 + 2" "pathstart g $ 2 - 3 = a $ 2 - 3" "b $ 2 + 3 = b $ 2 + 3"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1710  | 
by(auto simp add: assms)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1711  | 
qed note z=this[unfolded P1P2 path_image_linepath]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1712  | 
show thesis apply(rule that[of z]) proof-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1713  | 
have "(z \<in> closed_segment (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) \<or>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1714  | 
z \<in> closed_segment (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f)) \<or>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1715  | 
z \<in> closed_segment (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) \<or>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1716  | 
z \<in> closed_segment (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]) \<Longrightarrow>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1717  | 
(((z \<in> closed_segment (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g)) \<or>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1718  | 
z \<in> closed_segment (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1])) \<or>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1719  | 
z \<in> closed_segment (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1])) \<or>  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1720  | 
z \<in> closed_segment (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]) \<Longrightarrow> False"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1721  | 
apply(simp only: segment_vertical segment_horizontal vector_2) proof- case goal1 note as=this  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1722  | 
      have "pathfinish f \<in> {a..b}" using assms(3) pathfinish_in_path_image[of f] by auto 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1723  | 
hence "1 + b $ 1 \<le> pathfinish f $ 1 \<Longrightarrow> False" unfolding mem_interval forall_2 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1724  | 
hence "z$1 \<noteq> pathfinish f$1" using as(2) using assms ab by(auto simp add:field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1725  | 
      moreover have "pathstart f \<in> {a..b}" using assms(3) pathstart_in_path_image[of f] by auto 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1726  | 
hence "1 + b $ 1 \<le> pathstart f $ 1 \<Longrightarrow> False" unfolding mem_interval forall_2 by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1727  | 
hence "z$1 \<noteq> pathstart f$1" using as(2) using assms ab by(auto simp add:field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1728  | 
ultimately have *:"z$2 = a$2 - 2" using goal1(1) by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1729  | 
have "z$1 \<noteq> pathfinish g$1" using as(2) using assms ab by(auto simp add:field_simps *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1730  | 
      moreover have "pathstart g \<in> {a..b}" using assms(4) pathstart_in_path_image[of g] by auto 
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1731  | 
note this[unfolded mem_interval forall_2]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1732  | 
hence "z$1 \<noteq> pathstart g$1" using as(1) using assms ab by(auto simp add:field_simps *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1733  | 
ultimately have "a $ 2 - 1 \<le> z $ 2 \<and> z $ 2 \<le> b $ 2 + 3 \<or> b $ 2 + 3 \<le> z $ 2 \<and> z $ 2 \<le> a $ 2 - 1"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1734  | 
using as(2) unfolding * assms by(auto simp add:field_simps)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1735  | 
thus False unfolding * using ab by auto  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1736  | 
qed hence "z \<in> path_image f \<or> z \<in> path_image g" using z unfolding Un_iff by blast  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1737  | 
    hence z':"z\<in>{a..b}" using assms(3-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
 | 
1738  | 
have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart f $ 1 \<or> z $ 1 = pathfinish f $ 1) \<Longrightarrow> (z = pathstart f \<or> z = pathfinish f)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1739  | 
unfolding Cart_eq forall_2 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
 | 
1740  | 
with z' show "z\<in>path_image f" using z(1) unfolding Un_iff mem_interval forall_2 apply-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1741  | 
apply(simp only: segment_vertical segment_horizontal vector_2) unfolding 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
 | 
1742  | 
have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart g $ 1 \<or> z $ 1 = pathfinish g $ 1) \<Longrightarrow> (z = pathstart g \<or> z = pathfinish g)"  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1743  | 
unfolding Cart_eq forall_2 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
 | 
1744  | 
with z' show "z\<in>path_image g" using z(2) unfolding Un_iff mem_interval forall_2 apply-  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1745  | 
apply(simp only: segment_vertical segment_horizontal vector_2) unfolding 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
 | 
1746  | 
qed qed  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1747  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1748  | 
(** The Following still needs to be translated. Maybe I will do that later.  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1749  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1750  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1751  | 
(* Complement in dimension N >= 2 of set homeomorphic to any interval in *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1752  | 
(* any dimension is (path-)connected. This naively generalizes the argument *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1753  | 
(* in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1754  | 
(* fixed point theorem", American Mathematical Monthly 1984. *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1755  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1756  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1757  | 
let RETRACTION_INJECTIVE_IMAGE_INTERVAL = prove  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1758  | 
(`!p:real^M->real^N a b.  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1759  | 
        ~(interval[a,b] = {}) /\
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1760  | 
p continuous_on interval[a,b] /\  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1761  | 
(!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ p x = p y ==> x = y)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1762  | 
==> ?f. f continuous_on (:real^N) /\  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1763  | 
IMAGE f (:real^N) SUBSET (IMAGE p (interval[a,b])) /\  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1764  | 
(!x. x IN (IMAGE p (interval[a,b])) ==> f x = x)`,  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1765  | 
REPEAT STRIP_TAC THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1766  | 
FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [INJECTIVE_ON_LEFT_INVERSE]) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1767  | 
DISCH_THEN(X_CHOOSE_TAC `q:real^N->real^M`) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1768  | 
SUBGOAL_THEN `(q:real^N->real^M) continuous_on  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1769  | 
(IMAGE p (interval[a:real^M,b]))`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1770  | 
ASSUME_TAC THENL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1771  | 
[MATCH_MP_TAC CONTINUOUS_ON_INVERSE THEN ASM_REWRITE_TAC[COMPACT_INTERVAL];  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1772  | 
ALL_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1773  | 
MP_TAC(ISPECL [`q:real^N->real^M`;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1774  | 
`IMAGE (p:real^M->real^N)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1775  | 
(interval[a,b])`;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1776  | 
`a:real^M`; `b:real^M`]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1777  | 
TIETZE_CLOSED_INTERVAL) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1778  | 
ASM_SIMP_TAC[COMPACT_INTERVAL; COMPACT_CONTINUOUS_IMAGE;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1779  | 
COMPACT_IMP_CLOSED] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1780  | 
ANTS_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1781  | 
DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^M` STRIP_ASSUME_TAC) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1782  | 
EXISTS_TAC `(p:real^M->real^N) o (r:real^N->real^M)` THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1783  | 
REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; o_THM; IN_UNIV] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1784  | 
CONJ_TAC THENL [ALL_TAC; ASM SET_TAC[]] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1785  | 
MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN ASM_REWRITE_TAC[] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1786  | 
FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP(REWRITE_RULE[IMP_CONJ]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1787  | 
CONTINUOUS_ON_SUBSET)) THEN ASM SET_TAC[]);;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1788  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1789  | 
let UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1790  | 
(`!s:real^N->bool a b:real^M.  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1791  | 
s homeomorphic (interval[a,b])  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1792  | 
==> !x. ~(x IN s) ==> ~bounded(path_component((:real^N) DIFF s) x)`,  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1793  | 
REPEAT GEN_TAC THEN REWRITE_TAC[homeomorphic; homeomorphism] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1794  | 
REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1795  | 
MAP_EVERY X_GEN_TAC [`p':real^N->real^M`; `p:real^M->real^N`] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1796  | 
DISCH_TAC THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1797  | 
SUBGOAL_THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1798  | 
`!x y. x IN interval[a,b] /\ y IN interval[a,b] /\  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1799  | 
(p:real^M->real^N) x = p y ==> x = y`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1800  | 
ASSUME_TAC THENL [ASM_MESON_TAC[]; ALL_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1801  | 
FIRST_X_ASSUM(MP_TAC o funpow 4 CONJUNCT2) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1802  | 
DISCH_THEN(CONJUNCTS_THEN2 (SUBST1_TAC o SYM) ASSUME_TAC) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1803  | 
  ASM_CASES_TAC `interval[a:real^M,b] = {}` THEN
 | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1804  | 
ASM_REWRITE_TAC[IMAGE_CLAUSES; DIFF_EMPTY; PATH_COMPONENT_UNIV;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1805  | 
NOT_BOUNDED_UNIV] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1806  | 
ABBREV_TAC `s = (:real^N) DIFF (IMAGE p (interval[a:real^M,b]))` THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1807  | 
X_GEN_TAC `c:real^N` THEN REPEAT STRIP_TAC THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1808  | 
SUBGOAL_THEN `(c:real^N) IN s` ASSUME_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1809  | 
SUBGOAL_THEN `bounded((path_component s c) UNION  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1810  | 
(IMAGE (p:real^M->real^N) (interval[a,b])))`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1811  | 
MP_TAC THENL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1812  | 
[ASM_SIMP_TAC[BOUNDED_UNION; COMPACT_IMP_BOUNDED; COMPACT_IMP_BOUNDED;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1813  | 
COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1814  | 
ALL_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1815  | 
DISCH_THEN(MP_TAC o SPEC `c:real^N` o MATCH_MP BOUNDED_SUBSET_BALL) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1816  | 
REWRITE_TAC[UNION_SUBSET] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1817  | 
DISCH_THEN(X_CHOOSE_THEN `B:real` STRIP_ASSUME_TAC) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1818  | 
MP_TAC(ISPECL [`p:real^M->real^N`; `a:real^M`; `b:real^M`]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1819  | 
RETRACTION_INJECTIVE_IMAGE_INTERVAL) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1820  | 
ASM_REWRITE_TAC[SUBSET; IN_UNIV] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1821  | 
DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^N` MP_TAC) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1822  | 
DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1823  | 
(CONJUNCTS_THEN2 MP_TAC ASSUME_TAC)) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1824  | 
REWRITE_TAC[FORALL_IN_IMAGE; IN_UNIV] THEN DISCH_TAC THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1825  | 
ABBREV_TAC `q = \z:real^N. if z IN path_component s c then r(z) else z` THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1826  | 
SUBGOAL_THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1827  | 
`(q:real^N->real^N) continuous_on  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1828  | 
(closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)))`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1829  | 
MP_TAC THENL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1830  | 
[EXPAND_TAC "q" THEN MATCH_MP_TAC CONTINUOUS_ON_CASES THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1831  | 
REWRITE_TAC[CLOSED_CLOSURE; CONTINUOUS_ON_ID; GSYM OPEN_CLOSED] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1832  | 
REPEAT CONJ_TAC THENL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1833  | 
[MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1834  | 
ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1835  | 
COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1836  | 
ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV];  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1837  | 
ALL_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1838  | 
X_GEN_TAC `z:real^N` THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1839  | 
REWRITE_TAC[SET_RULE `~(z IN (s DIFF t) /\ z IN t)`] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1840  | 
STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1841  | 
MP_TAC(ISPECL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1842  | 
[`path_component s (z:real^N)`; `path_component s (c:real^N)`]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1843  | 
OPEN_INTER_CLOSURE_EQ_EMPTY) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1844  | 
ASM_REWRITE_TAC[GSYM DISJOINT; PATH_COMPONENT_DISJOINT] THEN ANTS_TAC THENL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1845  | 
[MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1846  | 
ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1847  | 
COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1848  | 
REWRITE_TAC[DISJOINT; EXTENSION; IN_INTER; NOT_IN_EMPTY] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1849  | 
DISCH_THEN(MP_TAC o SPEC `z:real^N`) THEN ASM_REWRITE_TAC[] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1850  | 
GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [IN] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1851  | 
REWRITE_TAC[PATH_COMPONENT_REFL_EQ] THEN ASM SET_TAC[]];  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1852  | 
ALL_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1853  | 
SUBGOAL_THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1854  | 
`closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)) =  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1855  | 
(:real^N)`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1856  | 
SUBST1_TAC THENL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1857  | 
[MATCH_MP_TAC(SET_RULE `s SUBSET t ==> t UNION (UNIV DIFF s) = UNIV`) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1858  | 
REWRITE_TAC[CLOSURE_SUBSET];  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1859  | 
DISCH_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1860  | 
MP_TAC(ISPECL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1861  | 
[`(\x. &2 % c - x) o  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1862  | 
(\x. c + B / norm(x - c) % (x - c)) o (q:real^N->real^N)`;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1863  | 
`cball(c:real^N,B)`]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1864  | 
BROUWER) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1865  | 
REWRITE_TAC[NOT_IMP; GSYM CONJ_ASSOC; COMPACT_CBALL; CONVEX_CBALL] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1866  | 
ASM_SIMP_TAC[CBALL_EQ_EMPTY; REAL_LT_IMP_LE; REAL_NOT_LT] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1867  | 
SUBGOAL_THEN `!x. ~((q:real^N->real^N) x = c)` ASSUME_TAC THENL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1868  | 
[X_GEN_TAC `x:real^N` THEN EXPAND_TAC "q" THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1869  | 
REWRITE_TAC[NORM_EQ_0; VECTOR_SUB_EQ] THEN COND_CASES_TAC THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1870  | 
ASM SET_TAC[PATH_COMPONENT_REFL_EQ];  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1871  | 
ALL_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1872  | 
REPEAT CONJ_TAC THENL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1873  | 
[MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1874  | 
SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1875  | 
MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN CONJ_TAC THENL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1876  | 
[ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; ALL_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1877  | 
MATCH_MP_TAC CONTINUOUS_ON_ADD THEN REWRITE_TAC[CONTINUOUS_ON_CONST] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1878  | 
MATCH_MP_TAC CONTINUOUS_ON_MUL THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1879  | 
SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1880  | 
REWRITE_TAC[o_DEF; real_div; LIFT_CMUL] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1881  | 
MATCH_MP_TAC CONTINUOUS_ON_CMUL THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1882  | 
MATCH_MP_TAC(REWRITE_RULE[o_DEF] CONTINUOUS_ON_INV) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1883  | 
ASM_REWRITE_TAC[FORALL_IN_IMAGE; NORM_EQ_0; VECTOR_SUB_EQ] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1884  | 
SUBGOAL_THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1885  | 
`(\x:real^N. lift(norm(x - c))) = (lift o norm) o (\x. x - c)`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1886  | 
SUBST1_TAC THENL [REWRITE_TAC[o_DEF]; ALL_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1887  | 
MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1888  | 
ASM_SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1889  | 
CONTINUOUS_ON_LIFT_NORM];  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1890  | 
REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; IN_CBALL; o_THM; dist] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1891  | 
X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1892  | 
REWRITE_TAC[VECTOR_ARITH `c - (&2 % c - (c + x)) = x`] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1893  | 
REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1894  | 
ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1895  | 
ASM_REAL_ARITH_TAC;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1896  | 
REWRITE_TAC[NOT_EXISTS_THM; TAUT `~(c /\ b) <=> c ==> ~b`] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1897  | 
REWRITE_TAC[IN_CBALL; o_THM; dist] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1898  | 
X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1899  | 
REWRITE_TAC[VECTOR_ARITH `&2 % c - (c + x') = x <=> --x' = x - c`] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1900  | 
ASM_CASES_TAC `(x:real^N) IN path_component s c` THENL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1901  | 
[MATCH_MP_TAC(NORM_ARITH `norm(y) < B /\ norm(x) = B ==> ~(--x = y)`) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1902  | 
REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1903  | 
ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1904  | 
ASM_SIMP_TAC[REAL_ARITH `&0 < B ==> abs B = B`] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1905  | 
UNDISCH_TAC `path_component s c SUBSET ball(c:real^N,B)` THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1906  | 
REWRITE_TAC[SUBSET; IN_BALL] THEN ASM_MESON_TAC[dist; NORM_SUB];  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1907  | 
EXPAND_TAC "q" THEN REWRITE_TAC[] THEN ASM_REWRITE_TAC[] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1908  | 
REWRITE_TAC[VECTOR_ARITH `--(c % x) = x <=> (&1 + c) % x = vec 0`] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1909  | 
ASM_REWRITE_TAC[DE_MORGAN_THM; VECTOR_SUB_EQ; VECTOR_MUL_EQ_0] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1910  | 
SUBGOAL_THEN `~(x:real^N = c)` ASSUME_TAC THENL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1911  | 
[ASM_MESON_TAC[PATH_COMPONENT_REFL; IN]; ALL_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1912  | 
ASM_REWRITE_TAC[] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1913  | 
MATCH_MP_TAC(REAL_ARITH `&0 < x ==> ~(&1 + x = &0)`) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1914  | 
ASM_SIMP_TAC[REAL_LT_DIV; NORM_POS_LT; VECTOR_SUB_EQ]]]);;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1915  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1916  | 
let PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1917  | 
(`!s:real^N->bool a b:real^M.  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1918  | 
2 <= dimindex(:N) /\ s homeomorphic interval[a,b]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1919  | 
==> path_connected((:real^N) DIFF s)`,  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1920  | 
REPEAT STRIP_TAC THEN REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1921  | 
FIRST_ASSUM(MP_TAC o MATCH_MP  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1922  | 
UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1923  | 
ASM_REWRITE_TAC[SET_RULE `~(x IN s) <=> x IN (UNIV DIFF s)`] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1924  | 
ABBREV_TAC `t = (:real^N) DIFF s` THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1925  | 
DISCH_TAC THEN MAP_EVERY X_GEN_TAC [`x:real^N`; `y:real^N`] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1926  | 
STRIP_TAC THEN FIRST_ASSUM(MP_TAC o MATCH_MP HOMEOMORPHIC_COMPACTNESS) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1927  | 
REWRITE_TAC[COMPACT_INTERVAL] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1928  | 
DISCH_THEN(MP_TAC o MATCH_MP COMPACT_IMP_BOUNDED) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1929  | 
REWRITE_TAC[BOUNDED_POS; LEFT_IMP_EXISTS_THM] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1930  | 
X_GEN_TAC `B:real` THEN STRIP_TAC THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1931  | 
SUBGOAL_THEN `(?u:real^N. u IN path_component t x /\ B < norm(u)) /\  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1932  | 
(?v:real^N. v IN path_component t y /\ B < norm(v))`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1933  | 
STRIP_ASSUME_TAC THENL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1934  | 
[ASM_MESON_TAC[BOUNDED_POS; REAL_NOT_LE]; ALL_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1935  | 
MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `u:real^N` THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1936  | 
CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1937  | 
MATCH_MP_TAC PATH_COMPONENT_SYM THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1938  | 
MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `v:real^N` THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1939  | 
CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1940  | 
MATCH_MP_TAC PATH_COMPONENT_OF_SUBSET THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1941  | 
EXISTS_TAC `(:real^N) DIFF cball(vec 0,B)` THEN CONJ_TAC THENL  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1942  | 
[EXPAND_TAC "t" THEN MATCH_MP_TAC(SET_RULE  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1943  | 
`s SUBSET t ==> (u DIFF t) SUBSET (u DIFF s)`) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1944  | 
ASM_REWRITE_TAC[SUBSET; IN_CBALL_0];  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1945  | 
MP_TAC(ISPEC `cball(vec 0:real^N,B)`  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1946  | 
PATH_CONNECTED_COMPLEMENT_BOUNDED_CONVEX) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1947  | 
ASM_REWRITE_TAC[BOUNDED_CBALL; CONVEX_CBALL] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1948  | 
REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1949  | 
DISCH_THEN MATCH_MP_TAC THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1950  | 
ASM_REWRITE_TAC[IN_DIFF; IN_UNIV; IN_CBALL_0; REAL_NOT_LE]]);;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1951  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1952  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1953  | 
(* In particular, apply all these to the special case of an arc. *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1954  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1955  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1956  | 
let RETRACTION_ARC = prove  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1957  | 
(`!p. arc p  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1958  | 
==> ?f. f continuous_on (:real^N) /\  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1959  | 
IMAGE f (:real^N) SUBSET path_image p /\  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1960  | 
(!x. x IN path_image p ==> f x = x)`,  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1961  | 
REWRITE_TAC[arc; path; path_image] THEN REPEAT STRIP_TAC THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1962  | 
MATCH_MP_TAC RETRACTION_INJECTIVE_IMAGE_INTERVAL THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1963  | 
ASM_REWRITE_TAC[INTERVAL_EQ_EMPTY_1; DROP_VEC; REAL_NOT_LT; REAL_POS]);;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1964  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1965  | 
let PATH_CONNECTED_ARC_COMPLEMENT = prove  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1966  | 
(`!p. 2 <= dimindex(:N) /\ arc p  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1967  | 
==> path_connected((:real^N) DIFF path_image p)`,  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1968  | 
REWRITE_TAC[arc; path] THEN REPEAT STRIP_TAC THEN SIMP_TAC[path_image] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1969  | 
MP_TAC(ISPECL [`path_image p:real^N->bool`; `vec 0:real^1`; `vec 1:real^1`]  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1970  | 
PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1971  | 
ASM_REWRITE_TAC[path_image] THEN DISCH_THEN MATCH_MP_TAC THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1972  | 
ONCE_REWRITE_TAC[HOMEOMORPHIC_SYM] THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1973  | 
MATCH_MP_TAC HOMEOMORPHIC_COMPACT THEN  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1974  | 
EXISTS_TAC `p:real^1->real^N` THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]);;  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1975  | 
|
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1976  | 
let CONNECTED_ARC_COMPLEMENT = prove  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1977  | 
(`!p. 2 <= dimindex(:N) /\ arc p  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1978  | 
==> connected((:real^N) DIFF path_image p)`,  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1979  | 
SIMP_TAC[PATH_CONNECTED_ARC_COMPLEMENT; PATH_CONNECTED_IMP_CONNECTED]);; *)  | 
| 
 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 
hoelzl 
parents:  
diff
changeset
 | 
1980  | 
|
| 
34291
 
4e896680897e
finite annotation on cartesian product is now implicit.
 
hoelzl 
parents: 
34289 
diff
changeset
 | 
1981  | 
end  |