src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy
author wenzelm
Mon, 07 Dec 2015 20:19:59 +0100
changeset 61808 fc1556774cfe
parent 61609 77b453bd616f
child 61945 1135b8de26c3
permissions -rw-r--r--
isabelle update_cartouches -c -t;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
     1
(*  Author:     John Harrison
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
     2
    Author:     Robert Himmelmann, TU Muenchen (Translation from HOL light)
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
     3
*)
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
     4
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
     5
(* ========================================================================= *)
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
     6
(* Results connected with topological dimension.                             *)
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
     7
(*                                                                           *)
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
     8
(* At the moment this is just Brouwer's fixpoint theorem. The proof is from  *)
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
     9
(* Kuhn: "some combinatorial lemmas in topology", IBM J. v4. (1960) p. 518   *)
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
    10
(* See "http://www.research.ibm.com/journal/rd/045/ibmrd0405K.pdf".          *)
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
    11
(*                                                                           *)
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
    12
(* The script below is quite messy, but at least we avoid formalizing any    *)
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
    13
(* topological machinery; we don't even use barycentric subdivision; this is *)
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
    14
(* the big advantage of Kuhn's proof over the usual Sperner's lemma one.     *)
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
    15
(*                                                                           *)
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
    16
(*              (c) Copyright, John Harrison 1998-2008                       *)
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
    17
(* ========================================================================= *)
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
    18
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
    19
section \<open>Results connected with topological dimension.\<close>
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
    20
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
    21
theory Brouwer_Fixpoint
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
    22
imports Convex_Euclidean_Space
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
    23
begin
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
    24
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    25
lemma bij_betw_singleton_eq:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    26
  assumes f: "bij_betw f A B" and g: "bij_betw g A B" and a: "a \<in> A"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    27
  assumes eq: "(\<And>x. x \<in> A \<Longrightarrow> x \<noteq> a \<Longrightarrow> f x = g x)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    28
  shows "f a = g a"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    29
proof -
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    30
  have "f ` (A - {a}) = g ` (A - {a})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    31
    by (intro image_cong) (simp_all add: eq)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    32
  then have "B - {f a} = B - {g a}"
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 59765
diff changeset
    33
    using f g a  by (auto simp: bij_betw_def inj_on_image_set_diff set_eq_iff Diff_subset)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    34
  moreover have "f a \<in> B" "g a \<in> B"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    35
    using f g a by (auto simp: bij_betw_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    36
  ultimately show ?thesis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    37
    by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    38
qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    39
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    40
lemma swap_image:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    41
  "Fun.swap i j f ` A = (if i \<in> A then (if j \<in> A then f ` A else f ` ((A - {i}) \<union> {j}))
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    42
                                  else (if j \<in> A then f ` ((A - {j}) \<union> {i}) else f ` A))"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    43
  apply (auto simp: Fun.swap_def image_iff)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    44
  apply metis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    45
  apply (metis member_remove remove_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    46
  apply (metis member_remove remove_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    47
  done
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    48
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    49
lemma swap_apply1: "Fun.swap x y f x = f y"
56545
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56541
diff changeset
    50
  by (simp add: Fun.swap_def)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
    51
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    52
lemma swap_apply2: "Fun.swap x y f y = f x"
56545
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56541
diff changeset
    53
  by (simp add: Fun.swap_def)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    54
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    55
lemma (in -) lessThan_empty_iff: "{..< n::nat} = {} \<longleftrightarrow> n = 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    56
  by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    57
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    58
lemma Zero_notin_Suc: "0 \<notin> Suc ` A"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    59
  by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    60
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    61
lemma atMost_Suc_eq_insert_0: "{.. Suc n} = insert 0 (Suc ` {.. n})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    62
  apply auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    63
  apply (case_tac x)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    64
  apply auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    65
  done
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    66
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56571
diff changeset
    67
lemma setsum_union_disjoint':
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    68
  assumes "finite A"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    69
    and "finite B"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    70
    and "A \<inter> B = {}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    71
    and "A \<union> B = C"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    72
  shows "setsum g C = setsum g A + setsum g B"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56571
diff changeset
    73
  using setsum.union_disjoint[OF assms(1-3)] and assms(4) by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    74
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    75
lemma pointwise_minimal_pointwise_maximal:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    76
  fixes s :: "(nat \<Rightarrow> nat) set"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    77
  assumes "finite s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    78
    and "s \<noteq> {}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    79
    and "\<forall>x\<in>s. \<forall>y\<in>s. x \<le> y \<or> y \<le> x"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    80
  shows "\<exists>a\<in>s. \<forall>x\<in>s. a \<le> x"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    81
    and "\<exists>a\<in>s. \<forall>x\<in>s. x \<le> a"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    82
  using assms
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    83
proof (induct s rule: finite_ne_induct)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    84
  case (insert b s)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    85
  assume *: "\<forall>x\<in>insert b s. \<forall>y\<in>insert b s. x \<le> y \<or> y \<le> x"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    86
  moreover then obtain u l where "l \<in> s" "\<forall>b\<in>s. l \<le> b" "u \<in> s" "\<forall>b\<in>s. b \<le> u"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    87
    using insert by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    88
  ultimately show "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. a \<le> x" "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. x \<le> a"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    89
    using *[rule_format, of b u] *[rule_format, of b l] by (metis insert_iff order.trans)+
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
    90
qed auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
    91
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
    92
lemma brouwer_compactness_lemma:
56226
29fd6bd9228e generalize some theorems
huffman
parents: 56188
diff changeset
    93
  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
    94
  assumes "compact s"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
    95
    and "continuous_on s f"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
    96
    and "\<not> (\<exists>x\<in>s. f x = 0)"
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
    97
  obtains d where "0 < d" and "\<forall>x\<in>s. d \<le> norm (f x)"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
    98
proof (cases "s = {}")
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
    99
  case True
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
   100
  show thesis
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
   101
    by (rule that [of 1]) (auto simp: True)
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
   102
next
49374
b08c6312782b tuned proofs;
wenzelm
parents: 44890
diff changeset
   103
  case False
b08c6312782b tuned proofs;
wenzelm
parents: 44890
diff changeset
   104
  have "continuous_on s (norm \<circ> f)"
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56273
diff changeset
   105
    by (rule continuous_intros continuous_on_norm assms(2))+
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
   106
  with False obtain x where x: "x \<in> s" "\<forall>y\<in>s. (norm \<circ> f) x \<le> (norm \<circ> f) y"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
   107
    using continuous_attains_inf[OF assms(1), of "norm \<circ> f"]
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
   108
    unfolding o_def
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
   109
    by auto
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
   110
  have "(norm \<circ> f) x > 0"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
   111
    using assms(3) and x(1)
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
   112
    by auto
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
   113
  then show ?thesis
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
   114
    by (rule that) (insert x(2), auto simp: o_def)
49555
fb2128470345 tuned proofs;
wenzelm
parents: 49374
diff changeset
   115
qed
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
   116
49555
fb2128470345 tuned proofs;
wenzelm
parents: 49374
diff changeset
   117
lemma kuhn_labelling_lemma:
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
   118
  fixes P Q :: "'a::euclidean_space \<Rightarrow> bool"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   119
  assumes "\<forall>x. P x \<longrightarrow> P (f x)"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
   120
    and "\<forall>x. P x \<longrightarrow> (\<forall>i\<in>Basis. Q i \<longrightarrow> 0 \<le> x\<bullet>i \<and> x\<bullet>i \<le> 1)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
   121
  shows "\<exists>l. (\<forall>x.\<forall>i\<in>Basis. l x i \<le> (1::nat)) \<and>
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
   122
             (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (x\<bullet>i = 0) \<longrightarrow> (l x i = 0)) \<and>
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
   123
             (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (x\<bullet>i = 1) \<longrightarrow> (l x i = 1)) \<and>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   124
             (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (l x i = 0) \<longrightarrow> x\<bullet>i \<le> f x\<bullet>i) \<and>
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   125
             (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (l x i = 1) \<longrightarrow> f x\<bullet>i \<le> x\<bullet>i)"
49374
b08c6312782b tuned proofs;
wenzelm
parents: 44890
diff changeset
   126
proof -
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   127
  { fix x i
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   128
    let ?R = "\<lambda>y. (P x \<and> Q i \<and> x \<bullet> i = 0 \<longrightarrow> y = (0::nat)) \<and>
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   129
        (P x \<and> Q i \<and> x \<bullet> i = 1 \<longrightarrow> y = 1) \<and>
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   130
        (P x \<and> Q i \<and> y = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i) \<and>
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   131
        (P x \<and> Q i \<and> y = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   132
    { assume "P x" "Q i" "i \<in> Basis" with assms have "0 \<le> f x \<bullet> i \<and> f x \<bullet> i \<le> 1" by auto }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   133
    then have "i \<in> Basis \<Longrightarrow> ?R 0 \<or> ?R 1" by auto }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   134
  then show ?thesis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   135
    unfolding all_conj_distrib[symmetric] Ball_def (* FIXME: shouldn't this work by metis? *)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   136
    by (subst choice_iff[symmetric])+ blast
49374
b08c6312782b tuned proofs;
wenzelm
parents: 44890
diff changeset
   137
qed
b08c6312782b tuned proofs;
wenzelm
parents: 44890
diff changeset
   138
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
   139
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   140
subsection \<open>The key "counting" observation, somewhat abstracted.\<close>
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
   141
53252
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
   142
lemma kuhn_counting_lemma:
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   143
  fixes bnd compo compo' face S F
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   144
  defines "nF s == card {f\<in>F. face f s \<and> compo' f}"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   145
  assumes [simp, intro]: "finite F" \<comment> "faces" and [simp, intro]: "finite S" \<comment> "simplices"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   146
    and "\<And>f. f \<in> F \<Longrightarrow> bnd f \<Longrightarrow> card {s\<in>S. face f s} = 1"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   147
    and "\<And>f. f \<in> F \<Longrightarrow> \<not> bnd f \<Longrightarrow> card {s\<in>S. face f s} = 2"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   148
    and "\<And>s. s \<in> S \<Longrightarrow> compo s \<Longrightarrow> nF s = 1"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   149
    and "\<And>s. s \<in> S \<Longrightarrow> \<not> compo s \<Longrightarrow> nF s = 0 \<or> nF s = 2"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   150
    and "odd (card {f\<in>F. compo' f \<and> bnd f})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   151
  shows "odd (card {s\<in>S. compo s})"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
   152
proof -
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   153
  have "(\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) + (\<Sum>s | s \<in> S \<and> compo s. nF s) = (\<Sum>s\<in>S. nF s)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56571
diff changeset
   154
    by (subst setsum.union_disjoint[symmetric]) (auto intro!: setsum.cong)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   155
  also have "\<dots> = (\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> bnd f}. face f s}) +
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   156
                  (\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> \<not> bnd f}. face f s})"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56571
diff changeset
   157
    unfolding setsum.distrib[symmetric]
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   158
    by (subst card_Un_disjoint[symmetric])
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56571
diff changeset
   159
       (auto simp: nF_def intro!: setsum.cong arg_cong[where f=card])
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   160
  also have "\<dots> = 1 * card {f\<in>F. compo' f \<and> bnd f} + 2 * card {f\<in>F. compo' f \<and> \<not> bnd f}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   161
    using assms(4,5) by (fastforce intro!: arg_cong2[where f="op +"] setsum_multicount)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   162
  finally have "odd ((\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) + card {s\<in>S. compo s})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   163
    using assms(6,8) by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   164
  moreover have "(\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) =
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   165
    (\<Sum>s | s \<in> S \<and> \<not> compo s \<and> nF s = 0. nF s) + (\<Sum>s | s \<in> S \<and> \<not> compo s \<and> nF s = 2. nF s)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56571
diff changeset
   166
    using assms(7) by (subst setsum.union_disjoint[symmetric]) (fastforce intro!: setsum.cong)+
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
   167
  ultimately show ?thesis
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
   168
    by auto
53186
0f4d9df1eaec tuned proofs;
wenzelm
parents: 53185
diff changeset
   169
qed
0f4d9df1eaec tuned proofs;
wenzelm
parents: 53185
diff changeset
   170
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   171
subsection \<open>The odd/even result for faces of complete vertices, generalized.\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   172
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   173
lemma kuhn_complete_lemma:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   174
  assumes [simp]: "finite simplices"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   175
    and face: "\<And>f s. face f s \<longleftrightarrow> (\<exists>a\<in>s. f = s - {a})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   176
    and card_s[simp]:  "\<And>s. s \<in> simplices \<Longrightarrow> card s = n + 2"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   177
    and rl_bd: "\<And>s. s \<in> simplices \<Longrightarrow> rl ` s \<subseteq> {..Suc n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   178
    and bnd: "\<And>f s. s \<in> simplices \<Longrightarrow> face f s \<Longrightarrow> bnd f \<Longrightarrow> card {s\<in>simplices. face f s} = 1"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   179
    and nbnd: "\<And>f s. s \<in> simplices \<Longrightarrow> face f s \<Longrightarrow> \<not> bnd f \<Longrightarrow> card {s\<in>simplices. face f s} = 2"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   180
    and odd_card: "odd (card {f. (\<exists>s\<in>simplices. face f s) \<and> rl ` f = {..n} \<and> bnd f})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   181
  shows "odd (card {s\<in>simplices. (rl ` s = {..Suc n})})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   182
proof (rule kuhn_counting_lemma)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   183
  have finite_s[simp]: "\<And>s. s \<in> simplices \<Longrightarrow> finite s"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
   184
    by (metis add_is_0 zero_neq_numeral card_infinite assms(3))
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   185
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   186
  let ?F = "{f. \<exists>s\<in>simplices. face f s}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   187
  have F_eq: "?F = (\<Union>s\<in>simplices. \<Union>a\<in>s. {s - {a}})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   188
    by (auto simp: face)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   189
  show "finite ?F"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   190
    using \<open>finite simplices\<close> unfolding F_eq by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   191
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   192
  show "card {s \<in> simplices. face f s} = 1" if "f \<in> ?F" "bnd f" for f
60449
229bad93377e renamed "prems" to "that";
wenzelm
parents: 60421
diff changeset
   193
    using bnd that by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   194
60421
92d9557fb78c misc tuning;
wenzelm
parents: 60420
diff changeset
   195
  show "card {s \<in> simplices. face f s} = 2" if "f \<in> ?F" "\<not> bnd f" for f
60449
229bad93377e renamed "prems" to "that";
wenzelm
parents: 60421
diff changeset
   196
    using nbnd that by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   197
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   198
  show "odd (card {f \<in> {f. \<exists>s\<in>simplices. face f s}. rl ` f = {..n} \<and> bnd f})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   199
    using odd_card by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   200
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   201
  fix s assume s[simp]: "s \<in> simplices"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   202
  let ?S = "{f \<in> {f. \<exists>s\<in>simplices. face f s}. face f s \<and> rl ` f = {..n}}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   203
  have "?S = (\<lambda>a. s - {a}) ` {a\<in>s. rl ` (s - {a}) = {..n}}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   204
    using s by (fastforce simp: face)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   205
  then have card_S: "card ?S = card {a\<in>s. rl ` (s - {a}) = {..n}}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   206
    by (auto intro!: card_image inj_onI)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   207
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   208
  { assume rl: "rl ` s = {..Suc n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   209
    then have inj_rl: "inj_on rl s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   210
      by (intro eq_card_imp_inj_on) auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   211
    moreover obtain a where "rl a = Suc n" "a \<in> s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   212
      by (metis atMost_iff image_iff le_Suc_eq rl)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   213
    ultimately have n: "{..n} = rl ` (s - {a})"
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 59765
diff changeset
   214
      by (auto simp add: inj_on_image_set_diff Diff_subset rl)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   215
    have "{a\<in>s. rl ` (s - {a}) = {..n}} = {a}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   216
      using inj_rl \<open>a \<in> s\<close> by (auto simp add: n inj_on_image_eq_iff[OF inj_rl] Diff_subset)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   217
    then show "card ?S = 1"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   218
      unfolding card_S by simp }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   219
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   220
  { assume rl: "rl ` s \<noteq> {..Suc n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   221
    show "card ?S = 0 \<or> card ?S = 2"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   222
    proof cases
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   223
      assume *: "{..n} \<subseteq> rl ` s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   224
      with rl rl_bd[OF s] have rl_s: "rl ` s = {..n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   225
        by (auto simp add: atMost_Suc subset_insert_iff split: split_if_asm)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   226
      then have "\<not> inj_on rl s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   227
        by (intro pigeonhole) simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   228
      then obtain a b where ab: "a \<in> s" "b \<in> s" "rl a = rl b" "a \<noteq> b"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   229
        by (auto simp: inj_on_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   230
      then have eq: "rl ` (s - {a}) = rl ` s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   231
        by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   232
      with ab have inj: "inj_on rl (s - {a})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   233
        by (intro eq_card_imp_inj_on) (auto simp add: rl_s card_Diff_singleton_if)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   234
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   235
      { fix x assume "x \<in> s" "x \<notin> {a, b}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   236
        then have "rl ` s - {rl x} = rl ` ((s - {a}) - {x})"
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 59765
diff changeset
   237
          by (auto simp: eq Diff_subset inj_on_image_set_diff[OF inj])
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   238
        also have "\<dots> = rl ` (s - {x})"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   239
          using ab \<open>x \<notin> {a, b}\<close> by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   240
        also assume "\<dots> = rl ` s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   241
        finally have False
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   242
          using \<open>x\<in>s\<close> by auto }
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   243
      moreover
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   244
      { fix x assume "x \<in> {a, b}" with ab have "x \<in> s \<and> rl ` (s - {x}) = rl ` s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   245
          by (simp add: set_eq_iff image_iff Bex_def) metis }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   246
      ultimately have "{a\<in>s. rl ` (s - {a}) = {..n}} = {a, b}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   247
        unfolding rl_s[symmetric] by fastforce
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   248
      with \<open>a \<noteq> b\<close> show "card ?S = 0 \<or> card ?S = 2"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   249
        unfolding card_S by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   250
    next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   251
      assume "\<not> {..n} \<subseteq> rl ` s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   252
      then have "\<And>x. rl ` (s - {x}) \<noteq> {..n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   253
        by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   254
      then show "card ?S = 0 \<or> card ?S = 2"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   255
        unfolding card_S by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   256
    qed }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   257
qed fact
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   258
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   259
locale kuhn_simplex =
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   260
  fixes p n and base upd and s :: "(nat \<Rightarrow> nat) set"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   261
  assumes base: "base \<in> {..< n} \<rightarrow> {..< p}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   262
  assumes base_out: "\<And>i. n \<le> i \<Longrightarrow> base i = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   263
  assumes upd: "bij_betw upd {..< n} {..< n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   264
  assumes s_pre: "s = (\<lambda>i j. if j \<in> upd`{..< i} then Suc (base j) else base j) ` {.. n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   265
begin
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   266
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   267
definition "enum i j = (if j \<in> upd`{..< i} then Suc (base j) else base j)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   268
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   269
lemma s_eq: "s = enum ` {.. n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   270
  unfolding s_pre enum_def[abs_def] ..
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   271
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   272
lemma upd_space: "i < n \<Longrightarrow> upd i < n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   273
  using upd by (auto dest!: bij_betwE)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   274
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   275
lemma s_space: "s \<subseteq> {..< n} \<rightarrow> {.. p}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   276
proof -
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   277
  { fix i assume "i \<le> n" then have "enum i \<in> {..< n} \<rightarrow> {.. p}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   278
    proof (induct i)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   279
      case 0 then show ?case
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   280
        using base by (auto simp: Pi_iff less_imp_le enum_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   281
    next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   282
      case (Suc i) with base show ?case
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   283
        by (auto simp: Pi_iff Suc_le_eq less_imp_le enum_def intro: upd_space)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   284
    qed }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   285
  then show ?thesis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   286
    by (auto simp: s_eq)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   287
qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   288
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   289
lemma inj_upd: "inj_on upd {..< n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   290
  using upd by (simp add: bij_betw_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   291
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   292
lemma inj_enum: "inj_on enum {.. n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   293
proof -
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   294
  { fix x y :: nat assume "x \<noteq> y" "x \<le> n" "y \<le> n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   295
    with upd have "upd ` {..< x} \<noteq> upd ` {..< y}"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
   296
      by (subst inj_on_image_eq_iff[where C="{..< n}"]) (auto simp: bij_betw_def)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   297
    then have "enum x \<noteq> enum y"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   298
      by (auto simp add: enum_def fun_eq_iff) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   299
  then show ?thesis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   300
    by (auto simp: inj_on_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   301
qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   302
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   303
lemma enum_0: "enum 0 = base"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   304
  by (simp add: enum_def[abs_def])
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   305
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   306
lemma base_in_s: "base \<in> s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   307
  unfolding s_eq by (subst enum_0[symmetric]) auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   308
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   309
lemma enum_in: "i \<le> n \<Longrightarrow> enum i \<in> s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   310
  unfolding s_eq by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   311
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   312
lemma one_step:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   313
  assumes a: "a \<in> s" "j < n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   314
  assumes *: "\<And>a'. a' \<in> s \<Longrightarrow> a' \<noteq> a \<Longrightarrow> a' j = p'"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   315
  shows "a j \<noteq> p'"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   316
proof
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   317
  assume "a j = p'"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   318
  with * a have "\<And>a'. a' \<in> s \<Longrightarrow> a' j = p'"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   319
    by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   320
  then have "\<And>i. i \<le> n \<Longrightarrow> enum i j = p'"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   321
    unfolding s_eq by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   322
  from this[of 0] this[of n] have "j \<notin> upd ` {..< n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   323
    by (auto simp: enum_def fun_eq_iff split: split_if_asm)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   324
  with upd \<open>j < n\<close> show False
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   325
    by (auto simp: bij_betw_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   326
qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   327
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   328
lemma upd_inj: "i < n \<Longrightarrow> j < n \<Longrightarrow> upd i = upd j \<longleftrightarrow> i = j"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61284
diff changeset
   329
  using upd by (auto simp: bij_betw_def inj_on_eq_iff)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   330
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   331
lemma upd_surj: "upd ` {..< n} = {..< n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   332
  using upd by (auto simp: bij_betw_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   333
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   334
lemma in_upd_image: "A \<subseteq> {..< n} \<Longrightarrow> i < n \<Longrightarrow> upd i \<in> upd ` A \<longleftrightarrow> i \<in> A"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61284
diff changeset
   335
  using inj_on_image_mem_iff[of upd "{..< n}"] upd
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   336
  by (auto simp: bij_betw_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   337
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   338
lemma enum_inj: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i = enum j \<longleftrightarrow> i = j"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61284
diff changeset
   339
  using inj_enum by (auto simp: inj_on_eq_iff)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   340
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   341
lemma in_enum_image: "A \<subseteq> {.. n} \<Longrightarrow> i \<le> n \<Longrightarrow> enum i \<in> enum ` A \<longleftrightarrow> i \<in> A"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61284
diff changeset
   342
  using inj_on_image_mem_iff[OF inj_enum] by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   343
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   344
lemma enum_mono: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i \<le> enum j \<longleftrightarrow> i \<le> j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   345
  by (auto simp: enum_def le_fun_def in_upd_image Ball_def[symmetric])
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   346
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   347
lemma enum_strict_mono: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i < enum j \<longleftrightarrow> i < j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   348
  using enum_mono[of i j] enum_inj[of i j] by (auto simp add: le_less)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   349
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   350
lemma chain: "a \<in> s \<Longrightarrow> b \<in> s \<Longrightarrow> a \<le> b \<or> b \<le> a"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   351
  by (auto simp: s_eq enum_mono)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   352
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   353
lemma less: "a \<in> s \<Longrightarrow> b \<in> s \<Longrightarrow> a i < b i \<Longrightarrow> a < b"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   354
  using chain[of a b] by (auto simp: less_fun_def le_fun_def not_le[symmetric])
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   355
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   356
lemma enum_0_bot: "a \<in> s \<Longrightarrow> a = enum 0 \<longleftrightarrow> (\<forall>a'\<in>s. a \<le> a')"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   357
  unfolding s_eq by (auto simp: enum_mono Ball_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   358
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   359
lemma enum_n_top: "a \<in> s \<Longrightarrow> a = enum n \<longleftrightarrow> (\<forall>a'\<in>s. a' \<le> a)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   360
  unfolding s_eq by (auto simp: enum_mono Ball_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   361
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   362
lemma enum_Suc: "i < n \<Longrightarrow> enum (Suc i) = (enum i)(upd i := Suc (enum i (upd i)))"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   363
  by (auto simp: fun_eq_iff enum_def upd_inj)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   364
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   365
lemma enum_eq_p: "i \<le> n \<Longrightarrow> n \<le> j \<Longrightarrow> enum i j = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   366
  by (induct i) (auto simp: enum_Suc enum_0 base_out upd_space not_less[symmetric])
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   367
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   368
lemma out_eq_p: "a \<in> s \<Longrightarrow> n \<le> j \<Longrightarrow> a j = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   369
  unfolding s_eq by (auto simp add: enum_eq_p)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   370
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   371
lemma s_le_p: "a \<in> s \<Longrightarrow> a j \<le> p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   372
  using out_eq_p[of a j] s_space by (cases "j < n") auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   373
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   374
lemma le_Suc_base: "a \<in> s \<Longrightarrow> a j \<le> Suc (base j)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   375
  unfolding s_eq by (auto simp: enum_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   376
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   377
lemma base_le: "a \<in> s \<Longrightarrow> base j \<le> a j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   378
  unfolding s_eq by (auto simp: enum_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   379
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   380
lemma enum_le_p: "i \<le> n \<Longrightarrow> j < n \<Longrightarrow> enum i j \<le> p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   381
  using enum_in[of i] s_space by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   382
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   383
lemma enum_less: "a \<in> s \<Longrightarrow> i < n \<Longrightarrow> enum i < a \<longleftrightarrow> enum (Suc i) \<le> a"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   384
  unfolding s_eq by (auto simp: enum_strict_mono enum_mono)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   385
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   386
lemma ksimplex_0:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   387
  "n = 0 \<Longrightarrow> s = {(\<lambda>x. p)}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   388
  using s_eq enum_def base_out by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   389
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   390
lemma replace_0:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   391
  assumes "j < n" "a \<in> s" and p: "\<forall>x\<in>s - {a}. x j = 0" and "x \<in> s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   392
  shows "x \<le> a"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   393
proof cases
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   394
  assume "x \<noteq> a"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   395
  have "a j \<noteq> 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   396
    using assms by (intro one_step[where a=a]) auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   397
  with less[OF \<open>x\<in>s\<close> \<open>a\<in>s\<close>, of j] p[rule_format, of x] \<open>x \<in> s\<close> \<open>x \<noteq> a\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   398
  show ?thesis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   399
    by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   400
qed simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   401
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   402
lemma replace_1:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   403
  assumes "j < n" "a \<in> s" and p: "\<forall>x\<in>s - {a}. x j = p" and "x \<in> s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   404
  shows "a \<le> x"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   405
proof cases
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   406
  assume "x \<noteq> a"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   407
  have "a j \<noteq> p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   408
    using assms by (intro one_step[where a=a]) auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   409
  with enum_le_p[of _ j] \<open>j < n\<close> \<open>a\<in>s\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   410
  have "a j < p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   411
    by (auto simp: less_le s_eq)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   412
  with less[OF \<open>a\<in>s\<close> \<open>x\<in>s\<close>, of j] p[rule_format, of x] \<open>x \<in> s\<close> \<open>x \<noteq> a\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   413
  show ?thesis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   414
    by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   415
qed simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   416
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   417
end
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   418
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   419
locale kuhn_simplex_pair = s: kuhn_simplex p n b_s u_s s + t: kuhn_simplex p n b_t u_t t
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   420
  for p n b_s u_s s b_t u_t t
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   421
begin
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   422
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   423
lemma enum_eq:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   424
  assumes l: "i \<le> l" "l \<le> j" and "j + d \<le> n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   425
  assumes eq: "s.enum ` {i .. j} = t.enum ` {i + d .. j + d}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   426
  shows "s.enum l = t.enum (l + d)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   427
using l proof (induct l rule: dec_induct)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   428
  case base
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   429
  then have s: "s.enum i \<in> t.enum ` {i + d .. j + d}" and t: "t.enum (i + d) \<in> s.enum ` {i .. j}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   430
    using eq by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   431
  from t \<open>i \<le> j\<close> \<open>j + d \<le> n\<close> have "s.enum i \<le> t.enum (i + d)"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   432
    by (auto simp: s.enum_mono)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   433
  moreover from s \<open>i \<le> j\<close> \<open>j + d \<le> n\<close> have "t.enum (i + d) \<le> s.enum i"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   434
    by (auto simp: t.enum_mono)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   435
  ultimately show ?case
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   436
    by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   437
next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   438
  case (step l)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   439
  moreover from step.prems \<open>j + d \<le> n\<close> have
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   440
      "s.enum l < s.enum (Suc l)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   441
      "t.enum (l + d) < t.enum (Suc l + d)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   442
    by (simp_all add: s.enum_strict_mono t.enum_strict_mono)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   443
  moreover have
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   444
      "s.enum (Suc l) \<in> t.enum ` {i + d .. j + d}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   445
      "t.enum (Suc l + d) \<in> s.enum ` {i .. j}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   446
    using step \<open>j + d \<le> n\<close> eq by (auto simp: s.enum_inj t.enum_inj)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   447
  ultimately have "s.enum (Suc l) = t.enum (Suc (l + d))"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   448
    using \<open>j + d \<le> n\<close>
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
   449
    by (intro antisym s.enum_less[THEN iffD1] t.enum_less[THEN iffD1])
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   450
       (auto intro!: s.enum_in t.enum_in)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   451
  then show ?case by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   452
qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   453
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   454
lemma ksimplex_eq_bot:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   455
  assumes a: "a \<in> s" "\<And>a'. a' \<in> s \<Longrightarrow> a \<le> a'"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   456
  assumes b: "b \<in> t" "\<And>b'. b' \<in> t \<Longrightarrow> b \<le> b'"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   457
  assumes eq: "s - {a} = t - {b}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   458
  shows "s = t"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   459
proof cases
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   460
  assume "n = 0" with s.ksimplex_0 t.ksimplex_0 show ?thesis by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   461
next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   462
  assume "n \<noteq> 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   463
  have "s.enum 0 = (s.enum (Suc 0)) (u_s 0 := s.enum (Suc 0) (u_s 0) - 1)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   464
       "t.enum 0 = (t.enum (Suc 0)) (u_t 0 := t.enum (Suc 0) (u_t 0) - 1)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   465
    using \<open>n \<noteq> 0\<close> by (simp_all add: s.enum_Suc t.enum_Suc)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   466
  moreover have e0: "a = s.enum 0" "b = t.enum 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   467
    using a b by (simp_all add: s.enum_0_bot t.enum_0_bot)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   468
  moreover
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
   469
  { fix j assume "0 < j" "j \<le> n"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   470
    moreover have "s - {a} = s.enum ` {Suc 0 .. n}" "t - {b} = t.enum ` {Suc 0 .. n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   471
      unfolding s.s_eq t.s_eq e0 by (auto simp: s.enum_inj t.enum_inj)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   472
    ultimately have "s.enum j = t.enum j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   473
      using enum_eq[of "1" j n 0] eq by auto }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   474
  note enum_eq = this
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   475
  then have "s.enum (Suc 0) = t.enum (Suc 0)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   476
    using \<open>n \<noteq> 0\<close> by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   477
  moreover
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   478
  { fix j assume "Suc j < n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   479
    with enum_eq[of "Suc j"] enum_eq[of "Suc (Suc j)"]
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   480
    have "u_s (Suc j) = u_t (Suc j)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   481
      using s.enum_Suc[of "Suc j"] t.enum_Suc[of "Suc j"]
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   482
      by (auto simp: fun_eq_iff split: split_if_asm) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   483
  then have "\<And>j. 0 < j \<Longrightarrow> j < n \<Longrightarrow> u_s j = u_t j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   484
    by (auto simp: gr0_conv_Suc)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   485
  with \<open>n \<noteq> 0\<close> have "u_t 0 = u_s 0"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   486
    by (intro bij_betw_singleton_eq[OF t.upd s.upd, of 0]) auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   487
  ultimately have "a = b"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   488
    by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   489
  with assms show "s = t"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   490
    by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   491
qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   492
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   493
lemma ksimplex_eq_top:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   494
  assumes a: "a \<in> s" "\<And>a'. a' \<in> s \<Longrightarrow> a' \<le> a"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   495
  assumes b: "b \<in> t" "\<And>b'. b' \<in> t \<Longrightarrow> b' \<le> b"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   496
  assumes eq: "s - {a} = t - {b}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   497
  shows "s = t"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   498
proof (cases n)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   499
  assume "n = 0" with s.ksimplex_0 t.ksimplex_0 show ?thesis by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   500
next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   501
  case (Suc n')
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   502
  have "s.enum n = (s.enum n') (u_s n' := Suc (s.enum n' (u_s n')))"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   503
       "t.enum n = (t.enum n') (u_t n' := Suc (t.enum n' (u_t n')))"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   504
    using Suc by (simp_all add: s.enum_Suc t.enum_Suc)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   505
  moreover have en: "a = s.enum n" "b = t.enum n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   506
    using a b by (simp_all add: s.enum_n_top t.enum_n_top)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   507
  moreover
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
   508
  { fix j assume "j < n"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   509
    moreover have "s - {a} = s.enum ` {0 .. n'}" "t - {b} = t.enum ` {0 .. n'}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   510
      unfolding s.s_eq t.s_eq en by (auto simp: s.enum_inj t.enum_inj Suc)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   511
    ultimately have "s.enum j = t.enum j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   512
      using enum_eq[of "0" j n' 0] eq Suc by auto }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   513
  note enum_eq = this
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   514
  then have "s.enum n' = t.enum n'"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   515
    using Suc by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   516
  moreover
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   517
  { fix j assume "j < n'"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   518
    with enum_eq[of j] enum_eq[of "Suc j"]
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   519
    have "u_s j = u_t j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   520
      using s.enum_Suc[of j] t.enum_Suc[of j]
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   521
      by (auto simp: Suc fun_eq_iff split: split_if_asm) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   522
  then have "\<And>j. j < n' \<Longrightarrow> u_s j = u_t j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   523
    by (auto simp: gr0_conv_Suc)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   524
  then have "u_t n' = u_s n'"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   525
    by (intro bij_betw_singleton_eq[OF t.upd s.upd, of n']) (auto simp: Suc)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   526
  ultimately have "a = b"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   527
    by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   528
  with assms show "s = t"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   529
    by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   530
qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   531
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   532
end
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   533
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   534
inductive ksimplex for p n :: nat where
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   535
  ksimplex: "kuhn_simplex p n base upd s \<Longrightarrow> ksimplex p n s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   536
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   537
lemma finite_ksimplexes: "finite {s. ksimplex p n s}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   538
proof (rule finite_subset)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   539
  { fix a s assume "ksimplex p n s" "a \<in> s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   540
    then obtain b u where "kuhn_simplex p n b u s" by (auto elim: ksimplex.cases)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   541
    then interpret kuhn_simplex p n b u s .
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   542
    from s_space \<open>a \<in> s\<close> out_eq_p[OF \<open>a \<in> s\<close>]
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   543
    have "a \<in> (\<lambda>f x. if n \<le> x then p else f x) ` ({..< n} \<rightarrow>\<^sub>E {.. p})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   544
      by (auto simp: image_iff subset_eq Pi_iff split: split_if_asm
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   545
               intro!: bexI[of _ "restrict a {..< n}"]) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   546
  then show "{s. ksimplex p n s} \<subseteq> Pow ((\<lambda>f x. if n \<le> x then p else f x) ` ({..< n} \<rightarrow>\<^sub>E {.. p}))"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   547
    by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   548
qed (simp add: finite_PiE)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   549
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   550
lemma ksimplex_card:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   551
  assumes "ksimplex p n s" shows "card s = Suc n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   552
using assms proof cases
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   553
  case (ksimplex u b)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   554
  then interpret kuhn_simplex p n u b s .
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   555
  show ?thesis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   556
    by (simp add: card_image s_eq inj_enum)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   557
qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   558
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   559
lemma simplex_top_face:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   560
  assumes "0 < p" "\<forall>x\<in>s'. x n = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   561
  shows "ksimplex p n s' \<longleftrightarrow> (\<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> s' = s - {a})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   562
  using assms
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   563
proof safe
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   564
  fix s a assume "ksimplex p (Suc n) s" and a: "a \<in> s" and na: "\<forall>x\<in>s - {a}. x n = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   565
  then show "ksimplex p n (s - {a})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   566
  proof cases
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   567
    case (ksimplex base upd)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   568
    then interpret kuhn_simplex p "Suc n" base upd "s" .
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   569
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   570
    have "a n < p"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   571
      using one_step[of a n p] na \<open>a\<in>s\<close> s_space by (auto simp: less_le)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   572
    then have "a = enum 0"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   573
      using \<open>a \<in> s\<close> na by (subst enum_0_bot) (auto simp: le_less intro!: less[of a _ n])
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   574
    then have s_eq: "s - {a} = enum ` Suc ` {.. n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   575
      using s_eq by (simp add: atMost_Suc_eq_insert_0 insert_ident Zero_notin_Suc in_enum_image subset_eq)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   576
    then have "enum 1 \<in> s - {a}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   577
      by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   578
    then have "upd 0 = n"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   579
      using \<open>a n < p\<close> \<open>a = enum 0\<close> na[rule_format, of "enum 1"]
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   580
      by (auto simp: fun_eq_iff enum_Suc split: split_if_asm)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   581
    then have "bij_betw upd (Suc ` {..< n}) {..< n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   582
      using upd
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   583
      by (subst notIn_Un_bij_betw3[where b=0])
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   584
         (auto simp: lessThan_Suc[symmetric] lessThan_Suc_eq_insert_0)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   585
    then have "bij_betw (upd\<circ>Suc) {..<n} {..<n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   586
      by (rule bij_betw_trans[rotated]) (auto simp: bij_betw_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   587
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   588
    have "a n = p - 1"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   589
      using enum_Suc[of 0] na[rule_format, OF \<open>enum 1 \<in> s - {a}\<close>] \<open>a = enum 0\<close> by (auto simp: \<open>upd 0 = n\<close>)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   590
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   591
    show ?thesis
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60580
diff changeset
   592
    proof (rule ksimplex.intros, standard)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   593
      show "bij_betw (upd\<circ>Suc) {..< n} {..< n}" by fact
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   594
      show "base(n := p) \<in> {..<n} \<rightarrow> {..<p}" "\<And>i. n\<le>i \<Longrightarrow> (base(n := p)) i = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   595
        using base base_out by (auto simp: Pi_iff)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   596
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   597
      have "\<And>i. Suc ` {..< i} = {..< Suc i} - {0}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   598
        by (auto simp: image_iff Ball_def) arith
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   599
      then have upd_Suc: "\<And>i. i \<le> n \<Longrightarrow> (upd\<circ>Suc) ` {..< i} = upd ` {..< Suc i} - {n}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   600
        using \<open>upd 0 = n\<close> upd_inj
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 59765
diff changeset
   601
        by (auto simp add: image_comp[symmetric] inj_on_image_set_diff[OF inj_upd])
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   602
      have n_in_upd: "\<And>i. n \<in> upd ` {..< Suc i}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   603
        using \<open>upd 0 = n\<close> by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   604
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   605
      def f' \<equiv> "\<lambda>i j. if j \<in> (upd\<circ>Suc)`{..< i} then Suc ((base(n := p)) j) else (base(n := p)) j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   606
      { fix x i assume i[arith]: "i \<le> n" then have "enum (Suc i) x = f' i x"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   607
          unfolding f'_def enum_def using \<open>a n < p\<close> \<open>a = enum 0\<close> \<open>upd 0 = n\<close> \<open>a n = p - 1\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   608
          by (simp add: upd_Suc enum_0 n_in_upd) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   609
      then show "s - {a} = f' ` {.. n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   610
        unfolding s_eq image_comp by (intro image_cong) auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   611
    qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   612
  qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   613
next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   614
  assume "ksimplex p n s'" and *: "\<forall>x\<in>s'. x n = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   615
  then show "\<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> s' = s - {a}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   616
  proof cases
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   617
    case (ksimplex base upd)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   618
    then interpret kuhn_simplex p n base upd s' .
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   619
    def b \<equiv> "base (n := p - 1)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   620
    def u \<equiv> "\<lambda>i. case i of 0 \<Rightarrow> n | Suc i \<Rightarrow> upd i"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   621
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   622
    have "ksimplex p (Suc n) (s' \<union> {b})"
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60580
diff changeset
   623
    proof (rule ksimplex.intros, standard)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   624
      show "b \<in> {..<Suc n} \<rightarrow> {..<p}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   625
        using base \<open>0 < p\<close> unfolding lessThan_Suc b_def by (auto simp: PiE_iff)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   626
      show "\<And>i. Suc n \<le> i \<Longrightarrow> b i = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   627
        using base_out by (auto simp: b_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   628
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   629
      have "bij_betw u (Suc ` {..< n} \<union> {0}) ({..<n} \<union> {u 0})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   630
        using upd
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   631
        by (intro notIn_Un_bij_betw) (auto simp: u_def bij_betw_def image_comp comp_def inj_on_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   632
      then show "bij_betw u {..<Suc n} {..<Suc n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   633
        by (simp add: u_def lessThan_Suc[symmetric] lessThan_Suc_eq_insert_0)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   634
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   635
      def f' \<equiv> "\<lambda>i j. if j \<in> u`{..< i} then Suc (b j) else b j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   636
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   637
      have u_eq: "\<And>i. i \<le> n \<Longrightarrow> u ` {..< Suc i} = upd ` {..< i} \<union> { n }"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   638
        by (auto simp: u_def image_iff upd_inj Ball_def split: nat.split) arith
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   639
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   640
      { fix x have "x \<le> n \<Longrightarrow> n \<notin> upd ` {..<x}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   641
          using upd_space by (simp add: image_iff neq_iff) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   642
      note n_not_upd = this
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   643
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   644
      have *: "f' ` {.. Suc n} = f' ` (Suc ` {.. n} \<union> {0})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   645
        unfolding atMost_Suc_eq_insert_0 by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   646
      also have "\<dots> = (f' \<circ> Suc) ` {.. n} \<union> {b}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   647
        by (auto simp: f'_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   648
      also have "(f' \<circ> Suc) ` {.. n} = s'"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   649
        using \<open>0 < p\<close> base_out[of n]
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   650
        unfolding s_eq enum_def[abs_def] f'_def[abs_def] upd_space
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   651
        by (intro image_cong) (simp_all add: u_eq b_def fun_eq_iff n_not_upd)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   652
      finally show "s' \<union> {b} = f' ` {.. Suc n}" ..
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   653
    qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   654
    moreover have "b \<notin> s'"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   655
      using * \<open>0 < p\<close> by (auto simp: b_def)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   656
    ultimately show ?thesis by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   657
  qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   658
qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   659
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   660
lemma ksimplex_replace_0:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   661
  assumes s: "ksimplex p n s" and a: "a \<in> s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   662
  assumes j: "j < n" and p: "\<forall>x\<in>s - {a}. x j = 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   663
  shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 1"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   664
  using s
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   665
proof cases
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   666
  case (ksimplex b_s u_s)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   667
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
   668
  { fix t b assume "ksimplex p n t"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   669
    then obtain b_t u_t where "kuhn_simplex p n b_t u_t t"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   670
      by (auto elim: ksimplex.cases)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   671
    interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   672
      by intro_locales fact+
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   673
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   674
    assume b: "b \<in> t" "t - {b} = s - {a}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   675
    with a j p s.replace_0[of _ a] t.replace_0[of _ b] have "s = t"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   676
      by (intro ksimplex_eq_top[of a b]) auto }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   677
  then have "{s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = {s}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   678
    using s \<open>a \<in> s\<close> by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   679
  then show ?thesis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   680
    by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   681
qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   682
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   683
lemma ksimplex_replace_1:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   684
  assumes s: "ksimplex p n s" and a: "a \<in> s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   685
  assumes j: "j < n" and p: "\<forall>x\<in>s - {a}. x j = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   686
  shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 1"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   687
  using s
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   688
proof cases
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   689
  case (ksimplex b_s u_s)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   690
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
   691
  { fix t b assume "ksimplex p n t"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   692
    then obtain b_t u_t where "kuhn_simplex p n b_t u_t t"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   693
      by (auto elim: ksimplex.cases)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   694
    interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   695
      by intro_locales fact+
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   696
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   697
    assume b: "b \<in> t" "t - {b} = s - {a}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   698
    with a j p s.replace_1[of _ a] t.replace_1[of _ b] have "s = t"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   699
      by (intro ksimplex_eq_bot[of a b]) auto }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   700
  then have "{s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = {s}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   701
    using s \<open>a \<in> s\<close> by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   702
  then show ?thesis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   703
    by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   704
qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   705
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   706
lemma card_2_exists: "card s = 2 \<longleftrightarrow> (\<exists>x\<in>s. \<exists>y\<in>s. x \<noteq> y \<and> (\<forall>z\<in>s. z = x \<or> z = y))"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   707
  by (auto simp add: card_Suc_eq eval_nat_numeral)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   708
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   709
lemma ksimplex_replace_2:
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   710
  assumes s: "ksimplex p n s" and "a \<in> s" and "n \<noteq> 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   711
    and lb: "\<forall>j<n. \<exists>x\<in>s - {a}. x j \<noteq> 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   712
    and ub: "\<forall>j<n. \<exists>x\<in>s - {a}. x j \<noteq> p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   713
  shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 2"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   714
  using s
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   715
proof cases
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   716
  case (ksimplex base upd)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   717
  then interpret kuhn_simplex p n base upd s .
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   718
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   719
  from \<open>a \<in> s\<close> obtain i where "i \<le> n" "a = enum i"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   720
    unfolding s_eq by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   721
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   722
  from \<open>i \<le> n\<close> have "i = 0 \<or> i = n \<or> (0 < i \<and> i < n)"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   723
    by linarith
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   724
  then have "\<exists>!s'. s' \<noteq> s \<and> ksimplex p n s' \<and> (\<exists>b\<in>s'. s - {a} = s'- {b})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   725
  proof (elim disjE conjE)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   726
    assume "i = 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   727
    def rot \<equiv> "\<lambda>i. if i + 1 = n then 0 else i + 1"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   728
    let ?upd = "upd \<circ> rot"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   729
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   730
    have rot: "bij_betw rot {..< n} {..< n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   731
      by (auto simp: bij_betw_def inj_on_def image_iff Ball_def rot_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   732
         arith+
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   733
    from rot upd have "bij_betw ?upd {..<n} {..<n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   734
      by (rule bij_betw_trans)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   735
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   736
    def f' \<equiv> "\<lambda>i j. if j \<in> ?upd`{..< i} then Suc (enum (Suc 0) j) else enum (Suc 0) j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   737
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   738
    interpret b: kuhn_simplex p n "enum (Suc 0)" "upd \<circ> rot" "f' ` {.. n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   739
    proof
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   740
      from \<open>a = enum i\<close> ub \<open>n \<noteq> 0\<close> \<open>i = 0\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   741
      obtain i' where "i' \<le> n" "enum i' \<noteq> enum 0" "enum i' (upd 0) \<noteq> p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   742
        unfolding s_eq by (auto intro: upd_space simp: enum_inj)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   743
      then have "enum 1 \<le> enum i'" "enum i' (upd 0) < p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   744
        using enum_le_p[of i' "upd 0"] by (auto simp add: enum_inj enum_mono upd_space)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   745
      then have "enum 1 (upd 0) < p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   746
        by (auto simp add: le_fun_def intro: le_less_trans)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   747
      then show "enum (Suc 0) \<in> {..<n} \<rightarrow> {..<p}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   748
        using base \<open>n \<noteq> 0\<close> by (auto simp add: enum_0 enum_Suc PiE_iff extensional_def upd_space)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   749
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   750
      { fix i assume "n \<le> i" then show "enum (Suc 0) i = p"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   751
        using \<open>n \<noteq> 0\<close> by (auto simp: enum_eq_p) }
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   752
      show "bij_betw ?upd {..<n} {..<n}" by fact
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   753
    qed (simp add: f'_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   754
    have ks_f': "ksimplex p n (f' ` {.. n})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   755
      by rule unfold_locales
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   756
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   757
    have b_enum: "b.enum = f'" unfolding f'_def b.enum_def[abs_def] ..
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   758
    with b.inj_enum have inj_f': "inj_on f' {.. n}" by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   759
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   760
    have [simp]: "\<And>j. j < n \<Longrightarrow> rot ` {..< j} = {0 <..< Suc j}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   761
      by (auto simp: rot_def image_iff Ball_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   762
         arith
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   763
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   764
    { fix j assume j: "j < n"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   765
      from j \<open>n \<noteq> 0\<close> have "f' j = enum (Suc j)"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   766
        by (auto simp add: f'_def enum_def upd_inj in_upd_image image_comp[symmetric] fun_eq_iff) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   767
    note f'_eq_enum = this
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   768
    then have "enum ` Suc ` {..< n} = f' ` {..< n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   769
      by (force simp: enum_inj)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   770
    also have "Suc ` {..< n} = {.. n} - {0}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   771
      by (auto simp: image_iff Ball_def) arith
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   772
    also have "{..< n} = {.. n} - {n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   773
      by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   774
    finally have eq: "s - {a} = f' ` {.. n} - {f' n}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   775
      unfolding s_eq \<open>a = enum i\<close> \<open>i = 0\<close>
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 59765
diff changeset
   776
      by (simp add: Diff_subset inj_on_image_set_diff[OF inj_enum] inj_on_image_set_diff[OF inj_f'])
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   777
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   778
    have "enum 0 < f' 0"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   779
      using \<open>n \<noteq> 0\<close> by (simp add: enum_strict_mono f'_eq_enum)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   780
    also have "\<dots> < f' n"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   781
      using \<open>n \<noteq> 0\<close> b.enum_strict_mono[of 0 n] unfolding b_enum by simp
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   782
    finally have "a \<noteq> f' n"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   783
      using \<open>a = enum i\<close> \<open>i = 0\<close> by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   784
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   785
    { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   786
      obtain b u where "kuhn_simplex p n b u t"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   787
        using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   788
      then interpret t: kuhn_simplex p n b u t .
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   789
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   790
      { fix x assume "x \<in> s" "x \<noteq> a"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   791
         then have "x (upd 0) = enum (Suc 0) (upd 0)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   792
           by (auto simp: \<open>a = enum i\<close> \<open>i = 0\<close> s_eq enum_def enum_inj) }
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   793
      then have eq_upd0: "\<forall>x\<in>t-{c}. x (upd 0) = enum (Suc 0) (upd 0)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   794
        unfolding eq_sma[symmetric] by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   795
      then have "c (upd 0) \<noteq> enum (Suc 0) (upd 0)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   796
        using \<open>n \<noteq> 0\<close> by (intro t.one_step[OF \<open>c\<in>t\<close> ]) (auto simp: upd_space)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   797
      then have "c (upd 0) < enum (Suc 0) (upd 0) \<or> c (upd 0) > enum (Suc 0) (upd 0)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   798
        by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   799
      then have "t = s \<or> t = f' ` {..n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   800
      proof (elim disjE conjE)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   801
        assume *: "c (upd 0) < enum (Suc 0) (upd 0)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   802
        interpret st: kuhn_simplex_pair p n base upd s b u t ..
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   803
        { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "c \<le> x"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   804
            by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   805
        note top = this
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   806
        have "s = t"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   807
          using \<open>a = enum i\<close> \<open>i = 0\<close> \<open>c \<in> t\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   808
          by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq_sma])
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   809
             (auto simp: s_eq enum_mono t.s_eq t.enum_mono top)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   810
        then show ?thesis by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   811
      next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   812
        assume *: "c (upd 0) > enum (Suc 0) (upd 0)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   813
        interpret st: kuhn_simplex_pair p n "enum (Suc 0)" "upd \<circ> rot" "f' ` {.. n}" b u t ..
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   814
        have eq: "f' ` {..n} - {f' n} = t - {c}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   815
          using eq_sma eq by simp
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   816
        { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "x \<le> c"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   817
            by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   818
        note top = this
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   819
        have "f' ` {..n} = t"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   820
          using \<open>a = enum i\<close> \<open>i = 0\<close> \<open>c \<in> t\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   821
          by (intro st.ksimplex_eq_top[OF _ _ _ _ eq])
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   822
             (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono b_enum[symmetric] top)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   823
        then show ?thesis by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   824
      qed }
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   825
    with ks_f' eq \<open>a \<noteq> f' n\<close> \<open>n \<noteq> 0\<close> show ?thesis
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   826
      apply (intro ex1I[of _ "f' ` {.. n}"])
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   827
      apply auto []
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   828
      apply metis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   829
      done
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   830
  next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   831
    assume "i = n"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   832
    from \<open>n \<noteq> 0\<close> obtain n' where n': "n = Suc n'"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   833
      by (cases n) auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   834
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   835
    def rot \<equiv> "\<lambda>i. case i of 0 \<Rightarrow> n' | Suc i \<Rightarrow> i"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   836
    let ?upd = "upd \<circ> rot"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   837
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   838
    have rot: "bij_betw rot {..< n} {..< n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   839
      by (auto simp: bij_betw_def inj_on_def image_iff Bex_def rot_def n' split: nat.splits)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   840
         arith
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   841
    from rot upd have "bij_betw ?upd {..<n} {..<n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   842
      by (rule bij_betw_trans)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   843
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   844
    def b \<equiv> "base (upd n' := base (upd n') - 1)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   845
    def f' \<equiv> "\<lambda>i j. if j \<in> ?upd`{..< i} then Suc (b j) else b j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   846
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   847
    interpret b: kuhn_simplex p n b "upd \<circ> rot" "f' ` {.. n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   848
    proof
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   849
      { fix i assume "n \<le> i" then show "b i = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   850
          using base_out[of i] upd_space[of n'] by (auto simp: b_def n') }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   851
      show "b \<in> {..<n} \<rightarrow> {..<p}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   852
        using base \<open>n \<noteq> 0\<close> upd_space[of n']
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   853
        by (auto simp: b_def PiE_def Pi_iff Ball_def upd_space extensional_def n')
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   854
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   855
      show "bij_betw ?upd {..<n} {..<n}" by fact
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   856
    qed (simp add: f'_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   857
    have f': "b.enum = f'" unfolding f'_def b.enum_def[abs_def] ..
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   858
    have ks_f': "ksimplex p n (b.enum ` {.. n})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   859
      unfolding f' by rule unfold_locales
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   860
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
   861
    have "0 < n"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   862
      using \<open>n \<noteq> 0\<close> by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   863
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   864
    { from \<open>a = enum i\<close> \<open>n \<noteq> 0\<close> \<open>i = n\<close> lb upd_space[of n']
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   865
      obtain i' where "i' \<le> n" "enum i' \<noteq> enum n" "0 < enum i' (upd n')"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   866
        unfolding s_eq by (auto simp: enum_inj n')
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   867
      moreover have "enum i' (upd n') = base (upd n')"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   868
        unfolding enum_def using \<open>i' \<le> n\<close> \<open>enum i' \<noteq> enum n\<close> by (auto simp: n' upd_inj enum_inj)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   869
      ultimately have "0 < base (upd n')"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   870
        by auto }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   871
    then have benum1: "b.enum (Suc 0) = base"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   872
      unfolding b.enum_Suc[OF \<open>0<n\<close>] b.enum_0 by (auto simp: b_def rot_def)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   873
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   874
    have [simp]: "\<And>j. Suc j < n \<Longrightarrow> rot ` {..< Suc j} = {n'} \<union> {..< j}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   875
      by (auto simp: rot_def image_iff Ball_def split: nat.splits)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   876
    have rot_simps: "\<And>j. rot (Suc j) = j" "rot 0 = n'"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   877
      by (simp_all add: rot_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   878
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   879
    { fix j assume j: "Suc j \<le> n" then have "b.enum (Suc j) = enum j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   880
        by (induct j) (auto simp add: benum1 enum_0 b.enum_Suc enum_Suc rot_simps) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   881
    note b_enum_eq_enum = this
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   882
    then have "enum ` {..< n} = b.enum ` Suc ` {..< n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   883
      by (auto simp add: image_comp intro!: image_cong)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   884
    also have "Suc ` {..< n} = {.. n} - {0}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   885
      by (auto simp: image_iff Ball_def) arith
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   886
    also have "{..< n} = {.. n} - {n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   887
      by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   888
    finally have eq: "s - {a} = b.enum ` {.. n} - {b.enum 0}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   889
      unfolding s_eq \<open>a = enum i\<close> \<open>i = n\<close>
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 59765
diff changeset
   890
      using inj_on_image_set_diff[OF inj_enum Diff_subset, of "{n}"]
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 59765
diff changeset
   891
            inj_on_image_set_diff[OF b.inj_enum Diff_subset, of "{0}"]
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 59765
diff changeset
   892
      by (simp add: comp_def )
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   893
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   894
    have "b.enum 0 \<le> b.enum n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   895
      by (simp add: b.enum_mono)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   896
    also have "b.enum n < enum n"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   897
      using \<open>n \<noteq> 0\<close> by (simp add: enum_strict_mono b_enum_eq_enum n')
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   898
    finally have "a \<noteq> b.enum 0"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   899
      using \<open>a = enum i\<close> \<open>i = n\<close> by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   900
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   901
    { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   902
      obtain b' u where "kuhn_simplex p n b' u t"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   903
        using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   904
      then interpret t: kuhn_simplex p n b' u t .
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   905
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   906
      { fix x assume "x \<in> s" "x \<noteq> a"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   907
         then have "x (upd n') = enum n' (upd n')"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   908
           by (auto simp: \<open>a = enum i\<close> n' \<open>i = n\<close> s_eq enum_def enum_inj in_upd_image) }
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   909
      then have eq_upd0: "\<forall>x\<in>t-{c}. x (upd n') = enum n' (upd n')"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   910
        unfolding eq_sma[symmetric] by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   911
      then have "c (upd n') \<noteq> enum n' (upd n')"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   912
        using \<open>n \<noteq> 0\<close> by (intro t.one_step[OF \<open>c\<in>t\<close> ]) (auto simp: n' upd_space[unfolded n'])
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   913
      then have "c (upd n') < enum n' (upd n') \<or> c (upd n') > enum n' (upd n')"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   914
        by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   915
      then have "t = s \<or> t = b.enum ` {..n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   916
      proof (elim disjE conjE)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   917
        assume *: "c (upd n') > enum n' (upd n')"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   918
        interpret st: kuhn_simplex_pair p n base upd s b' u t ..
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   919
        { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "x \<le> c"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   920
            by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   921
        note top = this
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   922
        have "s = t"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   923
          using \<open>a = enum i\<close> \<open>i = n\<close> \<open>c \<in> t\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   924
          by (intro st.ksimplex_eq_top[OF _ _ _ _ eq_sma])
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   925
             (auto simp: s_eq enum_mono t.s_eq t.enum_mono top)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   926
        then show ?thesis by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   927
      next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   928
        assume *: "c (upd n') < enum n' (upd n')"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   929
        interpret st: kuhn_simplex_pair p n b "upd \<circ> rot" "f' ` {.. n}" b' u t ..
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   930
        have eq: "f' ` {..n} - {b.enum 0} = t - {c}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   931
          using eq_sma eq f' by simp
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   932
        { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "c \<le> x"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   933
            by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   934
        note bot = this
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   935
        have "f' ` {..n} = t"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   936
          using \<open>a = enum i\<close> \<open>i = n\<close> \<open>c \<in> t\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   937
          by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq])
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   938
             (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono bot)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   939
        with f' show ?thesis by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   940
      qed }
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   941
    with ks_f' eq \<open>a \<noteq> b.enum 0\<close> \<open>n \<noteq> 0\<close> show ?thesis
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   942
      apply (intro ex1I[of _ "b.enum ` {.. n}"])
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   943
      apply auto []
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   944
      apply metis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   945
      done
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   946
  next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   947
    assume i: "0 < i" "i < n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   948
    def i' \<equiv> "i - 1"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   949
    with i have "Suc i' < n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   950
      by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   951
    with i have Suc_i': "Suc i' = i"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   952
      by (simp add: i'_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   953
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   954
    let ?upd = "Fun.swap i' i upd"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   955
    from i upd have "bij_betw ?upd {..< n} {..< n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   956
      by (subst bij_betw_swap_iff) (auto simp: i'_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   957
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   958
    def f' \<equiv> "\<lambda>i j. if j \<in> ?upd`{..< i} then Suc (base j) else base j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   959
    interpret b: kuhn_simplex p n base ?upd "f' ` {.. n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   960
    proof
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   961
      show "base \<in> {..<n} \<rightarrow> {..<p}" by fact
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   962
      { fix i assume "n \<le> i" then show "base i = p" by fact }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   963
      show "bij_betw ?upd {..<n} {..<n}" by fact
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   964
    qed (simp add: f'_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   965
    have f': "b.enum = f'" unfolding f'_def b.enum_def[abs_def] ..
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   966
    have ks_f': "ksimplex p n (b.enum ` {.. n})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   967
      unfolding f' by rule unfold_locales
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   968
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   969
    have "{i} \<subseteq> {..n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   970
      using i by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   971
    { fix j assume "j \<le> n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   972
      moreover have "j < i \<or> i = j \<or> i < j" by arith
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   973
      moreover note i
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   974
      ultimately have "enum j = b.enum j \<longleftrightarrow> j \<noteq> i"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   975
        unfolding enum_def[abs_def] b.enum_def[abs_def]
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   976
        by (auto simp add: fun_eq_iff swap_image i'_def
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   977
                           in_upd_image inj_on_image_set_diff[OF inj_upd]) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   978
    note enum_eq_benum = this
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   979
    then have "enum ` ({.. n} - {i}) = b.enum ` ({.. n} - {i})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   980
      by (intro image_cong) auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   981
    then have eq: "s - {a} = b.enum ` {.. n} - {b.enum i}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   982
      unfolding s_eq \<open>a = enum i\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   983
      using inj_on_image_set_diff[OF inj_enum Diff_subset \<open>{i} \<subseteq> {..n}\<close>]
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   984
            inj_on_image_set_diff[OF b.inj_enum Diff_subset \<open>{i} \<subseteq> {..n}\<close>]
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   985
      by (simp add: comp_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   986
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   987
    have "a \<noteq> b.enum i"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   988
      using \<open>a = enum i\<close> enum_eq_benum i by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   989
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   990
    { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   991
      obtain b' u where "kuhn_simplex p n b' u t"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   992
        using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   993
      then interpret t: kuhn_simplex p n b' u t .
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   994
      have "enum i' \<in> s - {a}" "enum (i + 1) \<in> s - {a}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   995
        using \<open>a = enum i\<close> i enum_in by (auto simp: enum_inj i'_def)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   996
      then obtain l k where
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   997
        l: "t.enum l = enum i'" "l \<le> n" "t.enum l \<noteq> c" and
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   998
        k: "t.enum k = enum (i + 1)" "k \<le> n" "t.enum k \<noteq> c"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
   999
        unfolding eq_sma by (auto simp: t.s_eq)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1000
      with i have "t.enum l < t.enum k"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1001
        by (simp add: enum_strict_mono i'_def)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1002
      with \<open>l \<le> n\<close> \<open>k \<le> n\<close> have "l < k"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1003
        by (simp add: t.enum_strict_mono)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1004
      { assume "Suc l = k"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1005
        have "enum (Suc (Suc i')) = t.enum (Suc l)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1006
          using i by (simp add: k \<open>Suc l = k\<close> i'_def)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1007
        then have False
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1008
          using \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1009
          by (auto simp: t.enum_Suc enum_Suc l upd_inj fun_eq_iff split: split_if_asm)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1010
             (metis Suc_lessD n_not_Suc_n upd_inj) }
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1011
      with \<open>l < k\<close> have "Suc l < k"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1012
        by arith
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1013
      have c_eq: "c = t.enum (Suc l)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1014
      proof (rule ccontr)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1015
        assume "c \<noteq> t.enum (Suc l)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1016
        then have "t.enum (Suc l) \<in> s - {a}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1017
          using \<open>l < k\<close> \<open>k \<le> n\<close> by (simp add: t.s_eq eq_sma)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1018
        then obtain j where "t.enum (Suc l) = enum j" "j \<le> n" "enum j \<noteq> enum i"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1019
          unfolding s_eq \<open>a = enum i\<close> by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1020
        with i have "t.enum (Suc l) \<le> t.enum l \<or> t.enum k \<le> t.enum (Suc l)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1021
          by (auto simp add: i'_def enum_mono enum_inj l k)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1022
        with \<open>Suc l < k\<close> \<open>k \<le> n\<close> show False
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1023
          by (simp add: t.enum_mono)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1024
      qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1025
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1026
      { have "t.enum (Suc (Suc l)) \<in> s - {a}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1027
          unfolding eq_sma c_eq t.s_eq using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (auto simp: t.enum_inj)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1028
        then obtain j where eq: "t.enum (Suc (Suc l)) = enum j" and "j \<le> n" "j \<noteq> i"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1029
          by (auto simp: s_eq \<open>a = enum i\<close>)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1030
        moreover have "enum i' < t.enum (Suc (Suc l))"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1031
          unfolding l(1)[symmetric] using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (auto simp: t.enum_strict_mono)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1032
        ultimately have "i' < j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1033
          using i by (simp add: enum_strict_mono i'_def)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1034
        with \<open>j \<noteq> i\<close> \<open>j \<le> n\<close> have "t.enum k \<le> t.enum (Suc (Suc l))"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1035
          unfolding i'_def by (simp add: enum_mono k eq)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1036
        then have "k \<le> Suc (Suc l)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1037
          using \<open>k \<le> n\<close> \<open>Suc l < k\<close> by (simp add: t.enum_mono) }
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1038
      with \<open>Suc l < k\<close> have "Suc (Suc l) = k" by simp
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1039
      then have "enum (Suc (Suc i')) = t.enum (Suc (Suc l))"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1040
        using i by (simp add: k i'_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1041
      also have "\<dots> = (enum i') (u l := Suc (enum i' (u l)), u (Suc l) := Suc (enum i' (u (Suc l))))"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1042
        using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (simp add: t.enum_Suc l t.upd_inj)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
  1043
      finally have "(u l = upd i' \<and> u (Suc l) = upd (Suc i')) \<or>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1044
        (u l = upd (Suc i') \<and> u (Suc l) = upd i')"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1045
        using \<open>Suc i' < n\<close> by (auto simp: enum_Suc fun_eq_iff split: split_if_asm)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1046
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1047
      then have "t = s \<or> t = b.enum ` {..n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1048
      proof (elim disjE conjE)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1049
        assume u: "u l = upd i'"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1050
        have "c = t.enum (Suc l)" unfolding c_eq ..
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1051
        also have "t.enum (Suc l) = enum (Suc i')"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1052
          using u \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close> by (simp add: enum_Suc t.enum_Suc l)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1053
        also have "\<dots> = a"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1054
          using \<open>a = enum i\<close> i by (simp add: i'_def)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1055
        finally show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1056
          using eq_sma \<open>a \<in> s\<close> \<open>c \<in> t\<close> by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1057
      next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1058
        assume u: "u l = upd (Suc i')"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1059
        def B \<equiv> "b.enum ` {..n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1060
        have "b.enum i' = enum i'"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1061
          using enum_eq_benum[of i'] i by (auto simp add: i'_def gr0_conv_Suc)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1062
        have "c = t.enum (Suc l)" unfolding c_eq ..
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1063
        also have "t.enum (Suc l) = b.enum (Suc i')"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1064
          using u \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1065
          by (simp_all add: enum_Suc t.enum_Suc l b.enum_Suc \<open>b.enum i' = enum i'\<close> swap_apply1)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1066
             (simp add: Suc_i')
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1067
        also have "\<dots> = b.enum i"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1068
          using i by (simp add: i'_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1069
        finally have "c = b.enum i" .
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1070
        then have "t - {c} = B - {c}" "c \<in> B"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1071
          unfolding eq_sma[symmetric] eq B_def using i by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1072
        with \<open>c \<in> t\<close> have "t = B"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1073
          by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1074
        then show ?thesis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1075
          by (simp add: B_def)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1076
      qed }
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1077
    with ks_f' eq \<open>a \<noteq> b.enum i\<close> \<open>n \<noteq> 0\<close> \<open>i \<le> n\<close> show ?thesis
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1078
      apply (intro ex1I[of _ "b.enum ` {.. n}"])
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1079
      apply auto []
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1080
      apply metis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1081
      done
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1082
  qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1083
  then show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1084
    using s \<open>a \<in> s\<close> by (simp add: card_2_exists Ex1_def) metis
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1085
qed
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1086
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1087
text \<open>Hence another step towards concreteness.\<close>
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1088
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1089
lemma kuhn_simplex_lemma:
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1090
  assumes "\<forall>s. ksimplex p (Suc n) s \<longrightarrow> rl ` s \<subseteq> {.. Suc n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1091
    and "odd (card {f. \<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> (f = s - {a}) \<and>
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1092
      rl ` f = {..n} \<and> ((\<exists>j\<le>n. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>f. x j = p))})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1093
  shows "odd (card {s. ksimplex p (Suc n) s \<and> rl ` s = {..Suc n}})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1094
proof (rule kuhn_complete_lemma[OF finite_ksimplexes refl, unfolded mem_Collect_eq,
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1095
    where bnd="\<lambda>f. (\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = p)"],
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1096
    safe del: notI)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1097
53186
0f4d9df1eaec tuned proofs;
wenzelm
parents: 53185
diff changeset
  1098
  have *: "\<And>x y. x = y \<Longrightarrow> odd (card x) \<Longrightarrow> odd (card y)"
0f4d9df1eaec tuned proofs;
wenzelm
parents: 53185
diff changeset
  1099
    by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1100
  show "odd (card {f. (\<exists>s\<in>{s. ksimplex p (Suc n) s}. \<exists>a\<in>s. f = s - {a}) \<and>
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1101
    rl ` f = {..n} \<and> ((\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = p))})"
53186
0f4d9df1eaec tuned proofs;
wenzelm
parents: 53185
diff changeset
  1102
    apply (rule *[OF _ assms(2)])
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1103
    apply (auto simp: atLeast0AtMost)
53186
0f4d9df1eaec tuned proofs;
wenzelm
parents: 53185
diff changeset
  1104
    done
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1105
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1106
next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1107
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1108
  fix s assume s: "ksimplex p (Suc n) s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1109
  then show "card s = n + 2"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1110
    by (simp add: ksimplex_card)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1111
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1112
  fix a assume a: "a \<in> s" then show "rl a \<le> Suc n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1113
    using assms(1) s by (auto simp: subset_eq)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1114
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1115
  let ?S = "{t. ksimplex p (Suc n) t \<and> (\<exists>b\<in>t. s - {a} = t - {b})}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1116
  { fix j assume j: "j \<le> n" "\<forall>x\<in>s - {a}. x j = 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1117
    with s a show "card ?S = 1"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1118
      using ksimplex_replace_0[of p "n + 1" s a j]
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1119
      by (subst eq_commute) simp }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1120
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1121
  { fix j assume j: "j \<le> n" "\<forall>x\<in>s - {a}. x j = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1122
    with s a show "card ?S = 1"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1123
      using ksimplex_replace_1[of p "n + 1" s a j]
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1124
      by (subst eq_commute) simp }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1125
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1126
  { assume "card ?S \<noteq> 2" "\<not> (\<exists>j\<in>{..n}. \<forall>x\<in>s - {a}. x j = p)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1127
    with s a show "\<exists>j\<in>{..n}. \<forall>x\<in>s - {a}. x j = 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1128
      using ksimplex_replace_2[of p "n + 1" s a]
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1129
      by (subst (asm) eq_commute) auto }
53186
0f4d9df1eaec tuned proofs;
wenzelm
parents: 53185
diff changeset
  1130
qed
0f4d9df1eaec tuned proofs;
wenzelm
parents: 53185
diff changeset
  1131
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1132
subsection \<open>Reduced labelling\<close>
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1133
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1134
definition reduced :: "nat \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> nat" where "reduced n x = (LEAST k. k = n \<or> x k \<noteq> 0)"
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1135
53186
0f4d9df1eaec tuned proofs;
wenzelm
parents: 53185
diff changeset
  1136
lemma reduced_labelling:
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1137
  shows "reduced n x \<le> n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1138
    and "\<forall>i<reduced n x. x i = 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1139
    and "reduced n x = n \<or> x (reduced n x) \<noteq> 0"
53186
0f4d9df1eaec tuned proofs;
wenzelm
parents: 53185
diff changeset
  1140
proof -
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1141
  show "reduced n x \<le> n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1142
    unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1143
  show "\<forall>i<reduced n x. x i = 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1144
    unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) fastforce+
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1145
  show "reduced n x = n \<or> x (reduced n x) \<noteq> 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1146
    unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) fastforce+
53186
0f4d9df1eaec tuned proofs;
wenzelm
parents: 53185
diff changeset
  1147
qed
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1148
53186
0f4d9df1eaec tuned proofs;
wenzelm
parents: 53185
diff changeset
  1149
lemma reduced_labelling_unique:
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1150
  "r \<le> n \<Longrightarrow> \<forall>i<r. x i = 0 \<Longrightarrow> r = n \<or> x r \<noteq> 0 \<Longrightarrow> reduced n x = r"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1151
 unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) (metis le_less not_le)+
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1152
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1153
lemma reduced_labelling_zero: "j < n \<Longrightarrow> x j = 0 \<Longrightarrow> reduced n x \<noteq> j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1154
  using reduced_labelling[of n x] by auto
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1155
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1156
lemma reduce_labelling_zero[simp]: "reduced 0 x = 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1157
  by (rule reduced_labelling_unique) auto
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1158
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1159
lemma reduced_labelling_nonzero: "j < n \<Longrightarrow> x j \<noteq> 0 \<Longrightarrow> reduced n x \<le> j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1160
  using reduced_labelling[of n x] by (elim allE[where x=j]) auto
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1161
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1162
lemma reduced_labelling_Suc: "reduced (Suc n) x \<noteq> Suc n \<Longrightarrow> reduced (Suc n) x = reduced n x"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1163
  using reduced_labelling[of "Suc n" x]
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1164
  by (intro reduced_labelling_unique[symmetric]) auto
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1165
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1166
lemma complete_face_top:
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1167
  assumes "\<forall>x\<in>f. \<forall>j\<le>n. x j = 0 \<longrightarrow> lab x j = 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1168
    and "\<forall>x\<in>f. \<forall>j\<le>n. x j = p \<longrightarrow> lab x j = 1"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1169
    and eq: "(reduced (Suc n) \<circ> lab) ` f = {..n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1170
  shows "((\<exists>j\<le>n. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>f. x j = p)) \<longleftrightarrow> (\<forall>x\<in>f. x n = p)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1171
proof (safe del: disjCI)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1172
  fix x j assume j: "j \<le> n" "\<forall>x\<in>f. x j = 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1173
  { fix x assume "x \<in> f" with assms j have "reduced (Suc n) (lab x) \<noteq> j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1174
      by (intro reduced_labelling_zero) auto }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1175
  moreover have "j \<in> (reduced (Suc n) \<circ> lab) ` f"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1176
    using j eq by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1177
  ultimately show "x n = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1178
    by force
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1179
next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1180
  fix x j assume j: "j \<le> n" "\<forall>x\<in>f. x j = p" and x: "x \<in> f"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1181
  have "j = n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1182
  proof (rule ccontr)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1183
    assume "\<not> ?thesis"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1184
    { fix x assume "x \<in> f"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1185
      with assms j have "reduced (Suc n) (lab x) \<le> j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1186
        by (intro reduced_labelling_nonzero) auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1187
      then have "reduced (Suc n) (lab x) \<noteq> n"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1188
        using \<open>j \<noteq> n\<close> \<open>j \<le> n\<close> by simp }
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1189
    moreover
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
  1190
    have "n \<in> (reduced (Suc n) \<circ> lab) ` f"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1191
      using eq by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1192
    ultimately show False
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1193
      by force
53186
0f4d9df1eaec tuned proofs;
wenzelm
parents: 53185
diff changeset
  1194
  qed
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1195
  moreover have "j \<in> (reduced (Suc n) \<circ> lab) ` f"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1196
    using j eq by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1197
  ultimately show "x n = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1198
    using j x by auto
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1199
qed auto
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1200
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1201
text \<open>Hence we get just about the nice induction.\<close>
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1202
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1203
lemma kuhn_induction:
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1204
  assumes "0 < p"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1205
    and lab_0: "\<forall>x. \<forall>j\<le>n. (\<forall>j. x j \<le> p) \<and> x j = 0 \<longrightarrow> lab x j = 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1206
    and lab_1: "\<forall>x. \<forall>j\<le>n. (\<forall>j. x j \<le> p) \<and> x j = p \<longrightarrow> lab x j = 1"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1207
    and odd: "odd (card {s. ksimplex p n s \<and> (reduced n\<circ>lab) ` s = {..n}})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1208
  shows "odd (card {s. ksimplex p (Suc n) s \<and> (reduced (Suc n)\<circ>lab) ` s = {..Suc n}})"
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1209
proof -
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1210
  let ?rl = "reduced (Suc n) \<circ> lab" and ?ext = "\<lambda>f v. \<exists>j\<le>n. \<forall>x\<in>f. x j = v"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1211
  let ?ext = "\<lambda>s. (\<exists>j\<le>n. \<forall>x\<in>s. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>s. x j = p)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1212
  have "\<forall>s. ksimplex p (Suc n) s \<longrightarrow> ?rl ` s \<subseteq> {..Suc n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1213
    by (simp add: reduced_labelling subset_eq)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1214
  moreover
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1215
  have "{s. ksimplex p n s \<and> (reduced n \<circ> lab) ` s = {..n}} =
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1216
        {f. \<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> f = s - {a} \<and> ?rl ` f = {..n} \<and> ?ext f}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1217
  proof (intro set_eqI, safe del: disjCI equalityI disjE)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1218
    fix s assume s: "ksimplex p n s" and rl: "(reduced n \<circ> lab) ` s = {..n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1219
    from s obtain u b where "kuhn_simplex p n u b s" by (auto elim: ksimplex.cases)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1220
    then interpret kuhn_simplex p n u b s .
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1221
    have all_eq_p: "\<forall>x\<in>s. x n = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1222
      by (auto simp: out_eq_p)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1223
    moreover
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1224
    { fix x assume "x \<in> s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1225
      with lab_1[rule_format, of n x] all_eq_p s_le_p[of x]
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1226
      have "?rl x \<le> n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1227
        by (auto intro!: reduced_labelling_nonzero)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1228
      then have "?rl x = reduced n (lab x)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1229
        by (auto intro!: reduced_labelling_Suc) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1230
    then have "?rl ` s = {..n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1231
      using rl by (simp cong: image_cong)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1232
    moreover
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1233
    obtain t a where "ksimplex p (Suc n) t" "a \<in> t" "s = t - {a}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1234
      using s unfolding simplex_top_face[OF \<open>0 < p\<close> all_eq_p] by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1235
    ultimately
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1236
    show "\<exists>t a. ksimplex p (Suc n) t \<and> a \<in> t \<and> s = t - {a} \<and> ?rl ` s = {..n} \<and> ?ext s"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1237
      by auto
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1238
  next
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1239
    fix x s a assume s: "ksimplex p (Suc n) s" and rl: "?rl ` (s - {a}) = {.. n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1240
      and a: "a \<in> s" and "?ext (s - {a})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1241
    from s obtain u b where "kuhn_simplex p (Suc n) u b s" by (auto elim: ksimplex.cases)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1242
    then interpret kuhn_simplex p "Suc n" u b s .
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1243
    have all_eq_p: "\<forall>x\<in>s. x (Suc n) = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1244
      by (auto simp: out_eq_p)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1245
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1246
    { fix x assume "x \<in> s - {a}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1247
      then have "?rl x \<in> ?rl ` (s - {a})"
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1248
        by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1249
      then have "?rl x \<le> n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1250
        unfolding rl by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1251
      then have "?rl x = reduced n (lab x)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1252
        by (auto intro!: reduced_labelling_Suc) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1253
    then show rl': "(reduced n\<circ>lab) ` (s - {a}) = {..n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1254
      unfolding rl[symmetric] by (intro image_cong) auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1255
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1256
    from \<open>?ext (s - {a})\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1257
    have all_eq_p: "\<forall>x\<in>s - {a}. x n = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1258
    proof (elim disjE exE conjE)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1259
      fix j assume "j \<le> n" "\<forall>x\<in>s - {a}. x j = 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1260
      with lab_0[rule_format, of j] all_eq_p s_le_p
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1261
      have "\<And>x. x \<in> s - {a} \<Longrightarrow> reduced (Suc n) (lab x) \<noteq> j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1262
        by (intro reduced_labelling_zero) auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1263
      moreover have "j \<in> ?rl ` (s - {a})"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1264
        using \<open>j \<le> n\<close> unfolding rl by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1265
      ultimately show ?thesis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1266
        by force
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1267
    next
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1268
      fix j assume "j \<le> n" and eq_p: "\<forall>x\<in>s - {a}. x j = p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1269
      show ?thesis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1270
      proof cases
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1271
        assume "j = n" with eq_p show ?thesis by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1272
      next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1273
        assume "j \<noteq> n"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1274
        { fix x assume x: "x \<in> s - {a}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1275
          have "reduced n (lab x) \<le> j"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1276
          proof (rule reduced_labelling_nonzero)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1277
            show "lab x j \<noteq> 0"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1278
              using lab_1[rule_format, of j x] x s_le_p[of x] eq_p \<open>j \<le> n\<close> by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1279
            show "j < n"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1280
              using \<open>j \<le> n\<close> \<open>j \<noteq> n\<close> by simp
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1281
          qed
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1282
          then have "reduced n (lab x) \<noteq> n"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1283
            using \<open>j \<le> n\<close> \<open>j \<noteq> n\<close> by simp }
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1284
        moreover have "n \<in> (reduced n\<circ>lab) ` (s - {a})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1285
          unfolding rl' by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1286
        ultimately show ?thesis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1287
          by force
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1288
      qed
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1289
    qed
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1290
    show "ksimplex p n (s - {a})"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1291
      unfolding simplex_top_face[OF \<open>0 < p\<close> all_eq_p] using s a by auto
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1292
  qed
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1293
  ultimately show ?thesis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1294
    using assms by (intro kuhn_simplex_lemma) auto
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1295
qed
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1296
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1297
text \<open>And so we get the final combinatorial result.\<close>
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1298
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1299
lemma ksimplex_0: "ksimplex p 0 s \<longleftrightarrow> s = {(\<lambda>x. p)}"
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1300
proof
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1301
  assume "ksimplex p 0 s" then show "s = {(\<lambda>x. p)}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1302
    by (blast dest: kuhn_simplex.ksimplex_0 elim: ksimplex.cases)
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1303
next
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1304
  assume s: "s = {(\<lambda>x. p)}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1305
  show "ksimplex p 0 s"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1306
  proof (intro ksimplex, unfold_locales)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1307
    show "(\<lambda>_. p) \<in> {..<0::nat} \<rightarrow> {..<p}" by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1308
    show "bij_betw id {..<0} {..<0}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1309
      by simp
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1310
  qed (auto simp: s)
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1311
qed
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1312
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1313
lemma kuhn_combinatorial:
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1314
  assumes "0 < p"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1315
    and "\<forall>x j. (\<forall>j. x j \<le> p) \<and> j < n \<and> x j = 0 \<longrightarrow> lab x j = 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1316
    and "\<forall>x j. (\<forall>j. x j \<le> p) \<and> j < n  \<and> x j = p \<longrightarrow> lab x j = 1"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1317
  shows "odd (card {s. ksimplex p n s \<and> (reduced n\<circ>lab) ` s = {..n}})"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1318
    (is "odd (card (?M n))")
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1319
  using assms
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1320
proof (induct n)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1321
  case 0 then show ?case
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1322
    by (simp add: ksimplex_0 cong: conj_cong)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1323
next
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1324
  case (Suc n)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1325
  then have "odd (card (?M n))"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1326
    by force
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1327
  with Suc show ?case
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1328
    using kuhn_induction[of p n] by (auto simp: comp_def)
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1329
qed
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1330
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1331
lemma kuhn_lemma:
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1332
  fixes n p :: nat
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1333
  assumes "0 < p"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1334
    and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. label x i = (0::nat) \<or> label x i = 1)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1335
    and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = 0 \<longrightarrow> label x i = 0)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1336
    and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = p \<longrightarrow> label x i = 1)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1337
  obtains q where "\<forall>i<n. q i < p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1338
    and "\<forall>i<n. \<exists>r s. (\<forall>j<n. q j \<le> r j \<and> r j \<le> q j + 1) \<and> (\<forall>j<n. q j \<le> s j \<and> s j \<le> q j + 1) \<and> label r i \<noteq> label s i"
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1339
proof -
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1340
  let ?rl = "reduced n \<circ> label"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1341
  let ?A = "{s. ksimplex p n s \<and> ?rl ` s = {..n}}"
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1342
  have "odd (card ?A)"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1343
    using assms by (intro kuhn_combinatorial[of p n label]) auto
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1344
  then have "?A \<noteq> {}"
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1345
    by fastforce
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1346
  then obtain s b u where "kuhn_simplex p n b u s" and rl: "?rl ` s = {..n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1347
    by (auto elim: ksimplex.cases)
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1348
  interpret kuhn_simplex p n b u s by fact
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1349
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1350
  show ?thesis
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1351
  proof (intro that[of b] allI impI)
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1352
    fix i
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1353
    assume "i < n"
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1354
    then show "b i < p"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1355
      using base by auto
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1356
  next
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1357
    fix i
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1358
    assume "i < n"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1359
    then have "i \<in> {.. n}" "Suc i \<in> {.. n}"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1360
      by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1361
    then obtain u v where u: "u \<in> s" "Suc i = ?rl u" and v: "v \<in> s" "i = ?rl v"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1362
      unfolding rl[symmetric] by blast
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1363
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1364
    have "label u i \<noteq> label v i"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1365
      using reduced_labelling [of n "label u"] reduced_labelling [of n "label v"]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1366
        u(2)[symmetric] v(2)[symmetric] \<open>i < n\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1367
      by auto
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1368
    moreover
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1369
    have "b j \<le> u j" "u j \<le> b j + 1" "b j \<le> v j" "v j \<le> b j + 1" if "j < n" for j
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1370
      using that base_le[OF \<open>u\<in>s\<close>] le_Suc_base[OF \<open>u\<in>s\<close>] base_le[OF \<open>v\<in>s\<close>] le_Suc_base[OF \<open>v\<in>s\<close>]
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1371
      by auto
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1372
    ultimately show "\<exists>r s. (\<forall>j<n. b j \<le> r j \<and> r j \<le> b j + 1) \<and>
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1373
        (\<forall>j<n. b j \<le> s j \<and> s j \<le> b j + 1) \<and> label r i \<noteq> label s i"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1374
      by blast
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1375
  qed
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1376
qed
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1377
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1378
subsection \<open>The main result for the unit cube\<close>
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1379
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1380
lemma kuhn_labelling_lemma':
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1381
  assumes "(\<forall>x::nat\<Rightarrow>real. P x \<longrightarrow> P (f x))"
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1382
    and "\<forall>x. P x \<longrightarrow> (\<forall>i::nat. Q i \<longrightarrow> 0 \<le> x i \<and> x i \<le> 1)"
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1383
  shows "\<exists>l. (\<forall>x i. l x i \<le> (1::nat)) \<and>
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1384
             (\<forall>x i. P x \<and> Q i \<and> x i = 0 \<longrightarrow> l x i = 0) \<and>
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1385
             (\<forall>x i. P x \<and> Q i \<and> x i = 1 \<longrightarrow> l x i = 1) \<and>
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1386
             (\<forall>x i. P x \<and> Q i \<and> l x i = 0 \<longrightarrow> x i \<le> f x i) \<and>
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1387
             (\<forall>x i. P x \<and> Q i \<and> l x i = 1 \<longrightarrow> f x i \<le> x i)"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1388
proof -
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1389
  have and_forall_thm: "\<And>P Q. (\<forall>x. P x) \<and> (\<forall>x. Q x) \<longleftrightarrow> (\<forall>x. P x \<and> Q x)"
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1390
    by auto
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1391
  have *: "\<forall>x y::real. 0 \<le> x \<and> x \<le> 1 \<and> 0 \<le> y \<and> y \<le> 1 \<longrightarrow> x \<noteq> 1 \<and> x \<le> y \<or> x \<noteq> 0 \<and> y \<le> x"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1392
    by auto
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1393
  show ?thesis
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1394
    unfolding and_forall_thm
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1395
    apply (subst choice_iff[symmetric])+
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1396
    apply rule
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1397
    apply rule
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1398
  proof -
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1399
    fix x x'
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1400
    let ?R = "\<lambda>y::nat.
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1401
      (P x \<and> Q x' \<and> x x' = 0 \<longrightarrow> y = 0) \<and>
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1402
      (P x \<and> Q x' \<and> x x' = 1 \<longrightarrow> y = 1) \<and>
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1403
      (P x \<and> Q x' \<and> y = 0 \<longrightarrow> x x' \<le> (f x) x') \<and>
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1404
      (P x \<and> Q x' \<and> y = 1 \<longrightarrow> (f x) x' \<le> x x')"
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1405
    have "0 \<le> f x x' \<and> f x x' \<le> 1" if "P x" "Q x'"
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1406
      using assms(2)[rule_format,of "f x" x'] that
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1407
      apply (drule_tac assms(1)[rule_format])
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1408
      apply auto
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1409
      done
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1410
    then have "?R 0 \<or> ?R 1"
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1411
      by auto
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1412
    then show "\<exists>y\<le>1. ?R y"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1413
      by auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1414
  qed
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1415
qed
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1416
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1417
definition unit_cube :: "'a::euclidean_space set"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1418
  where "unit_cube = {x. \<forall>i\<in>Basis. 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1}"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1419
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1420
lemma mem_unit_cube: "x \<in> unit_cube \<longleftrightarrow> (\<forall>i\<in>Basis. 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1421
  unfolding unit_cube_def by simp
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1422
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1423
lemma bounded_unit_cube: "bounded unit_cube"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1424
  unfolding bounded_def
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1425
proof (intro exI ballI)
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1426
  fix y :: 'a assume y: "y \<in> unit_cube"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1427
  have "dist 0 y = norm y" by (rule dist_0_norm)
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1428
  also have "\<dots> = norm (\<Sum>i\<in>Basis. (y \<bullet> i) *\<^sub>R i)" unfolding euclidean_representation ..
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1429
  also have "\<dots> \<le> (\<Sum>i\<in>Basis. norm ((y \<bullet> i) *\<^sub>R i))" by (rule norm_setsum)
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1430
  also have "\<dots> \<le> (\<Sum>i::'a\<in>Basis. 1)"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1431
    by (rule setsum_mono, simp add: y [unfolded mem_unit_cube])
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1432
  finally show "dist 0 y \<le> (\<Sum>i::'a\<in>Basis. 1)" .
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1433
qed
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1434
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1435
lemma closed_unit_cube: "closed unit_cube"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1436
  unfolding unit_cube_def Collect_ball_eq Collect_conj_eq
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1437
  by (rule closed_INT, auto intro!: closed_Collect_le)
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1438
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1439
lemma compact_unit_cube: "compact unit_cube" (is "compact ?C")
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1440
  unfolding compact_eq_seq_compact_metric
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1441
  using bounded_unit_cube closed_unit_cube
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1442
  by (rule bounded_closed_imp_seq_compact)
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1443
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1444
lemma brouwer_cube:
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1445
  fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1446
  assumes "continuous_on unit_cube f"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1447
    and "f ` unit_cube \<subseteq> unit_cube"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1448
  shows "\<exists>x\<in>unit_cube. f x = x"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1449
proof (rule ccontr)
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1450
  def n \<equiv> "DIM('a)"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1451
  have n: "1 \<le> n" "0 < n" "n \<noteq> 0"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1452
    unfolding n_def by (auto simp add: Suc_le_eq DIM_positive)
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1453
  assume "\<not> ?thesis"
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1454
  then have *: "\<not> (\<exists>x\<in>unit_cube. f x - x = 0)"
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1455
    by auto
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1456
  obtain d where
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1457
      d: "d > 0" "\<And>x. x \<in> unit_cube \<Longrightarrow> d \<le> norm (f x - x)"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1458
    apply (rule brouwer_compactness_lemma[OF compact_unit_cube _ *])
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56273
diff changeset
  1459
    apply (rule continuous_intros assms)+
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1460
    apply blast
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1461
    done
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1462
  have *: "\<forall>x. x \<in> unit_cube \<longrightarrow> f x \<in> unit_cube"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1463
    "\<forall>x. x \<in> (unit_cube::'a set) \<longrightarrow> (\<forall>i\<in>Basis. True \<longrightarrow> 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1464
    using assms(2)[unfolded image_subset_iff Ball_def]
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1465
    unfolding mem_unit_cube
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1466
    by auto
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1467
  obtain label :: "'a \<Rightarrow> 'a \<Rightarrow> nat" where
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1468
    "\<forall>x. \<forall>i\<in>Basis. label x i \<le> 1"
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1469
    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> x \<bullet> i = 0 \<longrightarrow> label x i = 0"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1470
    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> x \<bullet> i = 1 \<longrightarrow> label x i = 1"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1471
    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> label x i = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1472
    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> label x i = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i"
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1473
    using kuhn_labelling_lemma[OF *] by blast
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1474
  note label = this [rule_format]
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1475
  have lem1: "\<forall>x\<in>unit_cube. \<forall>y\<in>unit_cube. \<forall>i\<in>Basis. label x i \<noteq> label y i \<longrightarrow>
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1476
    abs (f x \<bullet> i - x \<bullet> i) \<le> norm (f y - f x) + norm (y - x)"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1477
  proof safe
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1478
    fix x y :: 'a
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1479
    assume x: "x \<in> unit_cube"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1480
    assume y: "y \<in> unit_cube"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1481
    fix i
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1482
    assume i: "label x i \<noteq> label y i" "i \<in> Basis"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1483
    have *: "\<And>x y fx fy :: real. x \<le> fx \<and> fy \<le> y \<or> fx \<le> x \<and> y \<le> fy \<Longrightarrow>
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1484
      abs (fx - x) \<le> abs (fy - fx) + abs (y - x)" by auto
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1485
    have "\<bar>(f x - x) \<bullet> i\<bar> \<le> abs ((f y - f x)\<bullet>i) + abs ((y - x)\<bullet>i)"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1486
      unfolding inner_simps
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1487
      apply (rule *)
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1488
      apply (cases "label x i = 0")
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1489
      apply (rule disjI1)
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1490
      apply rule
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1491
      prefer 3
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1492
      apply (rule disjI2)
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1493
      apply rule
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1494
    proof -
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1495
      assume lx: "label x i = 0"
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1496
      then have ly: "label y i = 1"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1497
        using i label(1)[of i y]
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1498
        by auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1499
      show "x \<bullet> i \<le> f x \<bullet> i"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1500
        apply (rule label(4)[rule_format])
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1501
        using x y lx i(2)
53252
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1502
        apply auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1503
        done
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1504
      show "f y \<bullet> i \<le> y \<bullet> i"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1505
        apply (rule label(5)[rule_format])
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1506
        using x y ly i(2)
53252
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1507
        apply auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1508
        done
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1509
    next
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1510
      assume "label x i \<noteq> 0"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1511
      then have l: "label x i = 1" "label y i = 0"
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1512
        using i label(1)[of i x] label(1)[of i y]
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1513
        by auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1514
      show "f x \<bullet> i \<le> x \<bullet> i"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1515
        apply (rule label(5)[rule_format])
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1516
        using x y l i(2)
53252
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1517
        apply auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1518
        done
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1519
      show "y \<bullet> i \<le> f y \<bullet> i"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1520
        apply (rule label(4)[rule_format])
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1521
        using x y l i(2)
53252
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1522
        apply auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1523
        done
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1524
    qed
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1525
    also have "\<dots> \<le> norm (f y - f x) + norm (y - x)"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1526
      apply (rule add_mono)
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1527
      apply (rule Basis_le_norm[OF i(2)])+
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1528
      done
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1529
    finally show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1530
      unfolding inner_simps .
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1531
  qed
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1532
  have "\<exists>e>0. \<forall>x\<in>unit_cube. \<forall>y\<in>unit_cube. \<forall>z\<in>unit_cube. \<forall>i\<in>Basis.
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1533
    norm (x - z) < e \<and> norm (y - z) < e \<and> label x i \<noteq> label y i \<longrightarrow>
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1534
      abs ((f(z) - z)\<bullet>i) < d / (real n)"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1535
  proof -
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1536
    have d': "d / real n / 8 > 0"
56541
0e3abadbef39 made divide_pos_pos a simp rule
nipkow
parents: 56371
diff changeset
  1537
      using d(1) by (simp add: n_def DIM_positive)
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1538
    have *: "uniformly_continuous_on unit_cube f"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1539
      by (rule compact_uniformly_continuous[OF assms(1) compact_unit_cube])
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1540
    obtain e where e:
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1541
        "e > 0"
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1542
        "\<And>x x'. x \<in> unit_cube \<Longrightarrow>
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1543
          x' \<in> unit_cube \<Longrightarrow>
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1544
          norm (x' - x) < e \<Longrightarrow>
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1545
          norm (f x' - f x) < d / real n / 8"
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1546
      using *[unfolded uniformly_continuous_on_def,rule_format,OF d']
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1547
      unfolding dist_norm
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1548
      by blast
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1549
    show ?thesis
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1550
      apply (rule_tac x="min (e/2) (d/real n/8)" in exI)
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1551
      apply safe
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1552
    proof -
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1553
      show "0 < min (e / 2) (d / real n / 8)"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1554
        using d' e by auto
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1555
      fix x y z i
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1556
      assume as:
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1557
        "x \<in> unit_cube" "y \<in> unit_cube" "z \<in> unit_cube"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36587
diff changeset
  1558
        "norm (x - z) < min (e / 2) (d / real n / 8)"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1559
        "norm (y - z) < min (e / 2) (d / real n / 8)"
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1560
        "label x i \<noteq> label y i"
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1561
      assume i: "i \<in> Basis"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1562
      have *: "\<And>z fz x fx n1 n2 n3 n4 d4 d :: real. abs(fx - x) \<le> n1 + n2 \<Longrightarrow>
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1563
        abs (fx - fz) \<le> n3 \<Longrightarrow> abs (x - z) \<le> n4 \<Longrightarrow>
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1564
        n1 < d4 \<Longrightarrow> n2 < 2 * d4 \<Longrightarrow> n3 < d4 \<Longrightarrow> n4 < d4 \<Longrightarrow>
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1565
        (8 * d4 = d) \<Longrightarrow> abs(fz - z) < d"
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1566
        by auto
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1567
      show "\<bar>(f z - z) \<bullet> i\<bar> < d / real n"
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1568
        unfolding inner_simps
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1569
      proof (rule *)
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1570
        show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y -f x) + norm (y - x)"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1571
          apply (rule lem1[rule_format])
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1572
          using as i
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1573
          apply auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1574
          done
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1575
        show "\<bar>f x \<bullet> i - f z \<bullet> i\<bar> \<le> norm (f x - f z)" "\<bar>x \<bullet> i - z \<bullet> i\<bar> \<le> norm (x - z)"
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1576
          unfolding inner_diff_left[symmetric]
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1577
          by (rule Basis_le_norm[OF i])+
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1578
        have tria: "norm (y - x) \<le> norm (y - z) + norm (x - z)"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1579
          using dist_triangle[of y x z, unfolded dist_norm]
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1580
          unfolding norm_minus_commute
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1581
          by auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1582
        also have "\<dots> < e / 2 + e / 2"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1583
          apply (rule add_strict_mono)
53252
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1584
          using as(4,5)
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1585
          apply auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1586
          done
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1587
        finally show "norm (f y - f x) < d / real n / 8"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1588
          apply -
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1589
          apply (rule e(2))
53252
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1590
          using as
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1591
          apply auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1592
          done
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1593
        have "norm (y - z) + norm (x - z) < d / real n / 8 + d / real n / 8"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1594
          apply (rule add_strict_mono)
53252
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1595
          using as
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1596
          apply auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1597
          done
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1598
        then show "norm (y - x) < 2 * (d / real n / 8)"
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1599
          using tria
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1600
          by auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1601
        show "norm (f x - f z) < d / real n / 8"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1602
          apply (rule e(2))
53252
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1603
          using as e(1)
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1604
          apply auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1605
          done
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1606
      qed (insert as, auto)
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1607
    qed
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1608
  qed
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1609
  then
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1610
  obtain e where e:
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1611
    "e > 0"
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1612
    "\<And>x y z i. x \<in> unit_cube \<Longrightarrow>
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1613
      y \<in> unit_cube \<Longrightarrow>
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1614
      z \<in> unit_cube \<Longrightarrow>
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1615
      i \<in> Basis \<Longrightarrow>
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1616
      norm (x - z) < e \<and> norm (y - z) < e \<and> label x i \<noteq> label y i \<Longrightarrow>
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1617
      \<bar>(f z - z) \<bullet> i\<bar> < d / real n"
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1618
    by blast
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1619
  obtain p :: nat where p: "1 + real n / e \<le> real p"
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1620
    using real_arch_simple ..
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1621
  have "1 + real n / e > 0"
56541
0e3abadbef39 made divide_pos_pos a simp rule
nipkow
parents: 56371
diff changeset
  1622
    using e(1) n by (simp add: add_pos_pos)
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1623
  then have "p > 0"
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1624
    using p by auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1625
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1626
  obtain b :: "nat \<Rightarrow> 'a" where b: "bij_betw b {..< n} Basis"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1627
    by atomize_elim (auto simp: n_def intro!: finite_same_card_bij)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1628
  def b' \<equiv> "inv_into {..< n} b"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1629
  then have b': "bij_betw b' Basis {..< n}"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1630
    using bij_betw_inv_into[OF b] by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1631
  then have b'_Basis: "\<And>i. i \<in> Basis \<Longrightarrow> b' i \<in> {..< n}"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1632
    unfolding bij_betw_def by (auto simp: set_eq_iff)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1633
  have bb'[simp]:"\<And>i. i \<in> Basis \<Longrightarrow> b (b' i) = i"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1634
    unfolding b'_def
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1635
    using b
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1636
    by (auto simp: f_inv_into_f bij_betw_def)
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1637
  have b'b[simp]:"\<And>i. i < n \<Longrightarrow> b' (b i) = i"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1638
    unfolding b'_def
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1639
    using b
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1640
    by (auto simp: inv_into_f_eq bij_betw_def)
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1641
  have *: "\<And>x :: nat. x = 0 \<or> x = 1 \<longleftrightarrow> x \<le> 1"
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1642
    by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1643
  have b'': "\<And>j. j < n \<Longrightarrow> b j \<in> Basis"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1644
    using b unfolding bij_betw_def by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1645
  have q1: "0 < p" "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow>
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1646
    (\<forall>i<n. (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0 \<or>
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1647
           (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1648
    unfolding *
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1649
    using \<open>p > 0\<close> \<open>n > 0\<close>
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1650
    using label(1)[OF b'']
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1651
    by auto
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1652
  { fix x :: "nat \<Rightarrow> nat" and i assume "\<forall>i<n. x i \<le> p" "i < n" "x i = p \<or> x i = 0"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1653
    then have "(\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<in> (unit_cube::'a set)"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1654
      using b'_Basis
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1655
      by (auto simp add: mem_unit_cube inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) }
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1656
  note cube = this
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1657
  have q2: "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = 0 \<longrightarrow>
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1658
      (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1659
    unfolding o_def using cube \<open>p > 0\<close> by (intro allI impI label(2)) (auto simp add: b'')
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
  1660
  have q3: "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = p \<longrightarrow>
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1661
      (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1662
    using cube \<open>p > 0\<close> unfolding o_def by (intro allI impI label(3)) (auto simp add: b'')
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1663
  obtain q where q:
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1664
      "\<forall>i<n. q i < p"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1665
      "\<forall>i<n.
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1666
         \<exists>r s. (\<forall>j<n. q j \<le> r j \<and> r j \<le> q j + 1) \<and>
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1667
               (\<forall>j<n. q j \<le> s j \<and> s j \<le> q j + 1) \<and>
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1668
               (label (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i) \<circ> b) i \<noteq>
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1669
               (label (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i) \<circ> b) i"
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1670
    by (rule kuhn_lemma[OF q1 q2 q3])
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1671
  def z \<equiv> "(\<Sum>i\<in>Basis. (real (q (b' i)) / real p) *\<^sub>R i)::'a"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1672
  have "\<exists>i\<in>Basis. d / real n \<le> abs ((f z - z)\<bullet>i)"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1673
  proof (rule ccontr)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1674
    have "\<forall>i\<in>Basis. q (b' i) \<in> {0..p}"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1675
      using q(1) b'
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1676
      by (auto intro: less_imp_le simp: bij_betw_def)
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1677
    then have "z \<in> unit_cube"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1678
      unfolding z_def mem_unit_cube
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1679
      using b'_Basis
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1680
      by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1)
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1681
    then have d_fz_z: "d \<le> norm (f z - z)"
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1682
      by (rule d)
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1683
    assume "\<not> ?thesis"
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1684
    then have as: "\<forall>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar> < d / real n"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1685
      using \<open>n > 0\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1686
      by (auto simp add: not_le inner_diff)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1687
    have "norm (f z - z) \<le> (\<Sum>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar>)"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1688
      unfolding inner_diff_left[symmetric]
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1689
      by (rule norm_le_l1)
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1690
    also have "\<dots> < (\<Sum>(i::'a) \<in> Basis. d / real n)"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1691
      apply (rule setsum_strict_mono)
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1692
      using as
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1693
      apply auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1694
      done
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1695
    also have "\<dots> = d"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1696
      using DIM_positive[where 'a='a]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
  1697
      by (auto simp: n_def)
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1698
    finally show False
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1699
      using d_fz_z by auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1700
  qed
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1701
  then obtain i where i: "i \<in> Basis" "d / real n \<le> \<bar>(f z - z) \<bullet> i\<bar>" ..
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1702
  have *: "b' i < n"
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1703
    using i and b'[unfolded bij_betw_def]
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1704
    by auto
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1705
  obtain r s where rs:
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1706
    "\<And>j. j < n \<Longrightarrow> q j \<le> r j \<and> r j \<le> q j + 1"
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1707
    "\<And>j. j < n \<Longrightarrow> q j \<le> s j \<and> s j \<le> q j + 1"
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1708
    "(label (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i) \<circ> b) (b' i) \<noteq>
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1709
      (label (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i) \<circ> b) (b' i)"
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1710
    using q(2)[rule_format,OF *] by blast
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1711
  have b'_im: "\<And>i. i \<in> Basis \<Longrightarrow>  b' i < n"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1712
    using b' unfolding bij_betw_def by auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1713
  def r' \<equiv> "(\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i)::'a"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1714
  have "\<And>i. i \<in> Basis \<Longrightarrow> r (b' i) \<le> p"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1715
    apply (rule order_trans)
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1716
    apply (rule rs(1)[OF b'_im,THEN conjunct2])
53252
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1717
    using q(1)[rule_format,OF b'_im]
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1718
    apply (auto simp add: Suc_le_eq)
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1719
    done
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1720
  then have "r' \<in> unit_cube"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1721
    unfolding r'_def mem_unit_cube
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1722
    using b'_Basis
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1723
    by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1724
  def s' \<equiv> "(\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i)::'a"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1725
  have "\<And>i. i \<in> Basis \<Longrightarrow> s (b' i) \<le> p"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1726
    apply (rule order_trans)
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1727
    apply (rule rs(2)[OF b'_im, THEN conjunct2])
53252
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1728
    using q(1)[rule_format,OF b'_im]
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1729
    apply (auto simp add: Suc_le_eq)
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1730
    done
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1731
  then have "s' \<in> unit_cube"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1732
    unfolding s'_def mem_unit_cube
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1733
    using b'_Basis
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1734
    by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1)
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1735
  have "z \<in> unit_cube"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1736
    unfolding z_def mem_unit_cube
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1737
    using b'_Basis q(1)[rule_format,OF b'_im] \<open>p > 0\<close>
56273
def3bbe6f2a5 cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents: 56226
diff changeset
  1738
    by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le)
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1739
  have *: "\<And>x. 1 + real x = real (Suc x)"
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1740
    by auto
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1741
  {
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1742
    have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1743
      apply (rule setsum_mono)
53252
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1744
      using rs(1)[OF b'_im]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
  1745
      apply (auto simp add:* field_simps simp del: of_nat_Suc)
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1746
      done
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1747
    also have "\<dots> < e * real p"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1748
      using p \<open>e > 0\<close> \<open>p > 0\<close>
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
  1749
      by (auto simp add: field_simps n_def)
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1750
    finally have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) < e * real p" .
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1751
  }
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1752
  moreover
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1753
  {
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1754
    have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1755
      apply (rule setsum_mono)
53252
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1756
      using rs(2)[OF b'_im]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
  1757
      apply (auto simp add:* field_simps simp del: of_nat_Suc)
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1758
      done
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1759
    also have "\<dots> < e * real p"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1760
      using p \<open>e > 0\<close> \<open>p > 0\<close>
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
  1761
      by (auto simp add: field_simps n_def)
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1762
    finally have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) < e * real p" .
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1763
  }
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1764
  ultimately
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1765
  have "norm (r' - z) < e" and "norm (s' - z) < e"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1766
    unfolding r'_def s'_def z_def
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1767
    using \<open>p > 0\<close>
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1768
    apply (rule_tac[!] le_less_trans[OF norm_le_l1])
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1769
    apply (auto simp add: field_simps setsum_divide_distrib[symmetric] inner_diff_left)
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1770
    done
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1771
  then have "\<bar>(f z - z) \<bullet> i\<bar> < d / real n"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1772
    using rs(3) i
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1773
    unfolding r'_def[symmetric] s'_def[symmetric] o_def bb'
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1774
    by (intro e(2)[OF \<open>r'\<in>unit_cube\<close> \<open>s'\<in>unit_cube\<close> \<open>z\<in>unit_cube\<close>]) auto
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1775
  then show False
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1776
    using i by auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1777
qed
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1778
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1779
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1780
subsection \<open>Retractions\<close>
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1781
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1782
definition "retraction s t r \<longleftrightarrow> t \<subseteq> s \<and> continuous_on s r \<and> r ` s \<subseteq> t \<and> (\<forall>x\<in>t. r x = x)"
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1783
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1784
definition retract_of (infixl "retract'_of" 12)
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1785
  where "(t retract_of s) \<longleftrightarrow> (\<exists>r. retraction s t r)"
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1786
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1787
lemma retraction_idempotent: "retraction s t r \<Longrightarrow> x \<in> s \<Longrightarrow>  r (r x) = r x"
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1788
  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
  1789
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1790
subsection \<open>Preservation of fixpoints under (more general notion of) retraction\<close>
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1791
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1792
lemma invertible_fixpoint_property:
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1793
  fixes s :: "'a::euclidean_space set"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1794
    and t :: "'b::euclidean_space set"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1795
  assumes "continuous_on t i"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1796
    and "i ` t \<subseteq> s"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1797
    and "continuous_on s r"
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1798
    and "r ` s \<subseteq> t"
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1799
    and "\<forall>y\<in>t. r (i y) = y"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1800
    and "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1801
    and "continuous_on t g"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1802
    and "g ` t \<subseteq> t"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1803
  obtains y where "y \<in> t" and "g y = y"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1804
proof -
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1805
  have "\<exists>x\<in>s. (i \<circ> g \<circ> r) x = x"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1806
    apply (rule assms(6)[rule_format])
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1807
    apply rule
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1808
    apply (rule continuous_on_compose assms)+
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1809
    apply ((rule continuous_on_subset)?, rule assms)+
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1810
    using assms(2,4,8)
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1811
    apply auto
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1812
    apply blast
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1813
    done
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1814
  then obtain x where x: "x \<in> s" "(i \<circ> g \<circ> r) x = x" ..
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1815
  then have *: "g (r x) \<in> t"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1816
    using assms(4,8) by auto
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1817
  have "r ((i \<circ> g \<circ> r) x) = r x"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1818
    using x by auto
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1819
  then show ?thesis
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1820
    apply (rule_tac that[of "r x"])
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1821
    using x
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1822
    unfolding o_def
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1823
    unfolding assms(5)[rule_format,OF *]
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1824
    using assms(4)
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1825
    apply auto
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1826
    done
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1827
qed
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1828
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1829
lemma homeomorphic_fixpoint_property:
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1830
  fixes s :: "'a::euclidean_space set"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1831
    and t :: "'b::euclidean_space set"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1832
  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
  1833
  shows "(\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)) \<longleftrightarrow>
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1834
    (\<forall>g. continuous_on t g \<and> g ` t \<subseteq> t \<longrightarrow> (\<exists>y\<in>t. g y = y))"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1835
proof -
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1836
  obtain r i where
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1837
      "\<forall>x\<in>s. i (r x) = x"
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1838
      "r ` s = t"
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1839
      "continuous_on s r"
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1840
      "\<forall>y\<in>t. r (i y) = y"
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1841
      "i ` t = s"
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1842
      "continuous_on t i"
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1843
    using assms
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1844
    unfolding homeomorphic_def homeomorphism_def
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1845
    by blast
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1846
  then show ?thesis
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1847
    apply -
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1848
    apply rule
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1849
    apply (rule_tac[!] allI impI)+
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1850
    apply (rule_tac g=g in invertible_fixpoint_property[of t i s r])
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1851
    prefer 10
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1852
    apply (rule_tac g=f in invertible_fixpoint_property[of s r t i])
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1853
    apply auto
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1854
    done
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1855
qed
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1856
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1857
lemma retract_fixpoint_property:
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1858
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1859
    and s :: "'a set"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1860
  assumes "t retract_of s"
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1861
    and "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1862
    and "continuous_on t g"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1863
    and "g ` t \<subseteq> t"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1864
  obtains y where "y \<in> t" and "g y = y"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1865
proof -
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1866
  obtain h where "retraction s t h"
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1867
    using assms(1) unfolding retract_of_def ..
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1868
  then show ?thesis
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1869
    unfolding retraction_def
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1870
    apply -
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1871
    apply (rule invertible_fixpoint_property[OF continuous_on_id _ _ _ _ assms(2), of t h g])
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1872
    prefer 7
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1873
    apply (rule_tac y = y in that)
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1874
    using assms
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1875
    apply auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1876
    done
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1877
qed
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1878
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1879
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1880
subsection \<open>The Brouwer theorem for any set with nonempty interior\<close>
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1881
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1882
lemma convex_unit_cube: "convex unit_cube"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1883
  apply (rule is_interval_convex)
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1884
  apply (clarsimp simp add: is_interval_def mem_unit_cube)
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1885
  apply (drule (1) bspec)+
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1886
  apply auto
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1887
  done
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1888
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50514
diff changeset
  1889
lemma brouwer_weak:
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1890
  fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1891
  assumes "compact s"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1892
    and "convex s"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1893
    and "interior s \<noteq> {}"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1894
    and "continuous_on s f"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1895
    and "f ` s \<subseteq> s"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1896
  obtains x where "x \<in> s" and "f x = x"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1897
proof -
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1898
  let ?U = "unit_cube :: 'a set"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1899
  have "\<Sum>Basis /\<^sub>R 2 \<in> interior ?U"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1900
  proof (rule interiorI)
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1901
    let ?I = "(\<Inter>i\<in>Basis. {x::'a. 0 < x \<bullet> i} \<inter> {x. x \<bullet> i < 1})"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1902
    show "open ?I"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1903
      by (intro open_INT finite_Basis ballI open_Int, auto intro: open_Collect_less)
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1904
    show "\<Sum>Basis /\<^sub>R 2 \<in> ?I"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1905
      by simp
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1906
    show "?I \<subseteq> unit_cube"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1907
      unfolding unit_cube_def by force
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1908
  qed
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1909
  then have *: "interior ?U \<noteq> {}" by fast
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1910
  have *: "?U homeomorphic s"
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1911
    using homeomorphic_convex_compact[OF convex_unit_cube compact_unit_cube * assms(2,1,3)] .
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1912
  have "\<forall>f. continuous_on ?U f \<and> f ` ?U \<subseteq> ?U \<longrightarrow>
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1913
    (\<exists>x\<in>?U. f x = x)"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36587
diff changeset
  1914
    using brouwer_cube by auto
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1915
  then show ?thesis
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1916
    unfolding homeomorphic_fixpoint_property[OF *]
53252
4766fbe322b5 tuned proofs;
wenzelm
parents: 53248
diff changeset
  1917
    using assms
59765
26d1c71784f1 tweaked a few slow or very ugly proofs
paulson <lp15@cam.ac.uk>
parents: 58877
diff changeset
  1918
    by (auto simp: intro: that)
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1919
qed
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1920
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1921
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1922
text \<open>And in particular for a closed ball.\<close>
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1923
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1924
lemma brouwer_ball:
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1925
  fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1926
  assumes "e > 0"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1927
    and "continuous_on (cball a e) f"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1928
    and "f ` cball a e \<subseteq> cball a e"
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1929
  obtains x where "x \<in> cball a e" and "f x = x"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1930
  using brouwer_weak[OF compact_cball convex_cball, of a e f]
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1931
  unfolding interior_cball ball_eq_empty
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1932
  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
  1933
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1934
text \<open>Still more general form; could derive this directly without using the
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
  1935
  rather involved \<open>HOMEOMORPHIC_CONVEX_COMPACT\<close> theorem, just using
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1936
  a scaling and translation to put the set inside the unit cube.\<close>
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1937
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  1938
lemma brouwer:
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1939
  fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1940
  assumes "compact s"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1941
    and "convex s"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1942
    and "s \<noteq> {}"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1943
    and "continuous_on s f"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1944
    and "f ` s \<subseteq> s"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1945
  obtains x where "x \<in> s" and "f x = x"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1946
proof -
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1947
  have "\<exists>e>0. s \<subseteq> cball 0 e"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1948
    using compact_imp_bounded[OF assms(1)]
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  1949
    unfolding bounded_pos
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1950
    apply (erule_tac exE)
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1951
    apply (rule_tac x=b in exI)
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1952
    apply (auto simp add: dist_norm)
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1953
    done
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1954
  then obtain e where e: "e > 0" "s \<subseteq> cball 0 e"
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1955
    by blast
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1956
  have "\<exists>x\<in> cball 0 e. (f \<circ> closest_point s) x = x"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1957
    apply (rule_tac brouwer_ball[OF e(1), of 0 "f \<circ> closest_point s"])
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1958
    apply (rule continuous_on_compose )
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1959
    apply (rule continuous_on_closest_point[OF assms(2) compact_imp_closed[OF assms(1)] assms(3)])
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1960
    apply (rule continuous_on_subset[OF assms(4)])
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1961
    apply (insert closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)])
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1962
    using assms(5)[unfolded subset_eq]
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1963
    using e(2)[unfolded subset_eq mem_cball]
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1964
    apply (auto simp add: dist_norm)
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1965
    done
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1966
  then obtain x where x: "x \<in> cball 0 e" "(f \<circ> closest_point s) x = x" ..
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1967
  have *: "closest_point s x = x"
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1968
    apply (rule closest_point_self)
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1969
    apply (rule assms(5)[unfolded subset_eq,THEN bspec[where x="x"], unfolded image_iff])
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1970
    apply (rule_tac x="closest_point s x" in bexI)
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1971
    using x
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1972
    unfolding o_def
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1973
    using closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3), of x]
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1974
    apply auto
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1975
    done
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1976
  show thesis
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1977
    apply (rule_tac x="closest_point s x" in that)
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1978
    unfolding x(2)[unfolded o_def]
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1979
    apply (rule closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)])
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1980
    using *
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1981
    apply auto
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1982
    done
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1983
qed
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1984
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
  1985
text \<open>So we get the no-retraction theorem.\<close>
33741
4c414d0835ab Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
hoelzl
parents:
diff changeset
  1986
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1987
lemma no_retraction_cball:
56117
2dbf84ee3deb remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
huffman
parents: 55522
diff changeset
  1988
  fixes a :: "'a::euclidean_space"
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1989
  assumes "e > 0"
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  1990
  shows "\<not> (frontier (cball a e) retract_of (cball a e))"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1991
proof
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1992
  assume *: "frontier (cball a e) retract_of (cball a e)"
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1993
  have **: "\<And>xa. a - (2 *\<^sub>R a - xa) = - (a - xa)"
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  1994
    using scaleR_left_distrib[of 1 1 a] by auto
55522
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1995
  obtain x where x:
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1996
      "x \<in> {x. norm (a - x) = e}"
23d2cbac6dce tuned proofs;
wenzelm
parents: 55493
diff changeset
  1997
      "2 *\<^sub>R a - x = x"
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  1998
    apply (rule retract_fixpoint_property[OF *, of "\<lambda>x. scaleR 2 a - x"])
59765
26d1c71784f1 tweaked a few slow or very ugly proofs
paulson <lp15@cam.ac.uk>
parents: 58877
diff changeset
  1999
    apply (blast intro: brouwer_ball[OF assms])
26d1c71784f1 tweaked a few slow or very ugly proofs
paulson <lp15@cam.ac.uk>
parents: 58877
diff changeset
  2000
    apply (intro continuous_intros)
26d1c71784f1 tweaked a few slow or very ugly proofs
paulson <lp15@cam.ac.uk>
parents: 58877
diff changeset
  2001
    unfolding frontier_cball subset_eq Ball_def image_iff dist_norm
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 60449
diff changeset
  2002
    apply (auto simp add: ** norm_minus_commute)
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  2003
    done
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  2004
  then have "scaleR 2 a = scaleR 1 x + scaleR 1 x"
53248
7a4b4b3b9ecd tuned proofs;
wenzelm
parents: 53186
diff changeset
  2005
    by (auto simp add: algebra_simps)
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  2006
  then have "a = x"
53688
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  2007
    unfolding scaleR_left_distrib[symmetric]
63892cfef47f tuned proofs;
wenzelm
parents: 53674
diff changeset
  2008
    by auto
53674
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  2009
  then show False
7ac7b2eaa5e6 tuned proofs;
wenzelm
parents: 53252
diff changeset
  2010
    using x assms by auto
53185
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  2011
qed
752e05d09708 tuned proofs;
wenzelm
parents: 51478
diff changeset
  2012
34291
4e896680897e finite annotation on cartesian product is now implicit.
hoelzl
parents: 34289
diff changeset
  2013
end