| author | wenzelm | 
| Sun, 20 Oct 2024 21:48:38 +0200 | |
| changeset 81214 | 7c2745efec69 | 
| parent 80095 | 0f9cd1a5edbe | 
| child 82323 | b022c013b04b | 
| permissions | -rw-r--r-- | 
| 53674 | 1 | (* Author: John Harrison | 
| 63305 
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
 paulson <lp15@cam.ac.uk> parents: 
63301diff
changeset | 2 | Author: Robert Himmelmann, TU Muenchen (Translation from HOL light) and LCP | 
| 53674 | 3 | *) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 4 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 5 | (* At the moment this is just Brouwer's fixpoint theorem. The proof is from *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 6 | (* Kuhn: "some combinatorial lemmas in topology", IBM J. v4. (1960) p. 518 *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 7 | (* See "http://www.research.ibm.com/journal/rd/045/ibmrd0405K.pdf". *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 8 | (* *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 9 | (* The script below is quite messy, but at least we avoid formalizing any *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 10 | (* topological machinery; we don't even use barycentric subdivision; this is *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 11 | (* the big advantage of Kuhn's proof over the usual Sperner's lemma one. *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 12 | (* *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 13 | (* (c) Copyright, John Harrison 1998-2008 *) | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 14 | |
| 68617 | 15 | section \<open>Brouwer's Fixed Point 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 | 16 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 17 | theory Brouwer_Fixpoint | 
| 70643 
93a7a85de312
Removal of the redundant ancestor Continuous_Extension
 paulson <lp15@cam.ac.uk> parents: 
70642diff
changeset | 18 | imports Homeomorphism Derivative | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 19 | begin | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 20 | |
| 68617 | 21 | subsection \<open>Retractions\<close> | 
| 22 | ||
| 23 | lemma retract_of_contractible: | |
| 24 | assumes "contractible T" "S retract_of T" | |
| 78457 | 25 | shows "contractible S" | 
| 26 | using assms | |
| 27 | apply (clarsimp simp add: retract_of_def contractible_def retraction_def homotopic_with image_subset_iff_funcset) | |
| 28 | apply (rule_tac x="r a" in exI) | |
| 29 | apply (rule_tac x="r \<circ> h" in exI) | |
| 30 | apply (intro conjI continuous_intros continuous_on_compose) | |
| 31 | apply (erule continuous_on_subset | force)+ | |
| 32 | done | |
| 68617 | 33 | |
| 34 | lemma retract_of_path_connected: | |
| 35 | "\<lbrakk>path_connected T; S retract_of T\<rbrakk> \<Longrightarrow> path_connected S" | |
| 36 | by (metis path_connected_continuous_image retract_of_def retraction) | |
| 37 | ||
| 38 | lemma retract_of_simply_connected: | |
| 80095 | 39 | assumes T: "simply_connected T" and "S retract_of T" | 
| 40 | shows "simply_connected S" | |
| 41 | proof - | |
| 42 | obtain r where r: "retraction T S r" | |
| 43 | using assms by (metis retract_of_def) | |
| 44 | have "S \<subseteq> T" | |
| 45 | by (meson \<open>retraction T S r\<close> retraction) | |
| 46 | then have "(\<lambda>a. a) \<in> S \<rightarrow> T" | |
| 47 | by blast | |
| 48 | then show ?thesis | |
| 49 | using simply_connected_retraction_gen [OF T] | |
| 50 | by (metis (no_types) r retraction retraction_refl) | |
| 51 | qed | |
| 68617 | 52 | |
| 53 | lemma retract_of_homotopically_trivial: | |
| 54 | assumes ts: "T retract_of S" | |
| 78457 | 55 | and hom: "\<And>f g. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S; | 
| 56 | continuous_on U g; g \<in> U \<rightarrow> S\<rbrakk> | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 57 | \<Longrightarrow> homotopic_with_canon (\<lambda>x. True) U S f g" | 
| 78457 | 58 | and "continuous_on U f" "f \<in> U \<rightarrow> T" | 
| 59 | and "continuous_on U g" "g \<in> U \<rightarrow> T" | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 60 | shows "homotopic_with_canon (\<lambda>x. True) U T f g" | 
| 68617 | 61 | proof - | 
| 78457 | 62 | obtain r where "r \<in> S \<rightarrow> S" "continuous_on S r" "\<forall>x\<in>S. r (r x) = r x" "T = r ` S" | 
| 68617 | 63 | using ts by (auto simp: retract_of_def retraction) | 
| 64 | then obtain k where "Retracts S r T k" | |
| 78457 | 65 | unfolding Retracts_def using continuous_on_id by blast | 
| 68617 | 66 | then show ?thesis | 
| 80095 | 67 | by (rule Retracts.homotopically_trivial_retraction_gen) (use assms hom in force)+ | 
| 68617 | 68 | qed | 
| 69 | ||
| 70 | lemma retract_of_homotopically_trivial_null: | |
| 71 | assumes ts: "T retract_of S" | |
| 78457 | 72 | and hom: "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S\<rbrakk> | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 73 | \<Longrightarrow> \<exists>c. homotopic_with_canon (\<lambda>x. True) U S f (\<lambda>x. c)" | 
| 78457 | 74 | and "continuous_on U f" "f \<in> U \<rightarrow> T" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 75 | obtains c where "homotopic_with_canon (\<lambda>x. True) U T f (\<lambda>x. c)" | 
| 68617 | 76 | proof - | 
| 78457 | 77 | obtain r where "r \<in> S \<rightarrow> S" "continuous_on S r" "\<forall>x\<in>S. r (r x) = r x" "T = r ` S" | 
| 68617 | 78 | using ts by (auto simp: retract_of_def retraction) | 
| 79 | then obtain k where "Retracts S r T k" | |
| 78457 | 80 | unfolding Retracts_def by fastforce | 
| 68617 | 81 | then show ?thesis | 
| 80095 | 82 | proof (rule Retracts.homotopically_trivial_retraction_null_gen) | 
| 83 | show "\<And>f. \<lbrakk>continuous_on U f; f \<in> U \<rightarrow> S\<rbrakk> | |
| 84 | \<Longrightarrow> \<exists>c. homotopic_with_canon (\<lambda>a. True) U S f (\<lambda>x. c)" | |
| 85 | using hom by blast | |
| 86 | qed (use assms that in auto) | |
| 68617 | 87 | qed | 
| 88 | ||
| 69945 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 89 | lemma retraction_openin_vimage_iff: | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69738diff
changeset | 90 | "openin (top_of_set S) (S \<inter> r -` U) \<longleftrightarrow> openin (top_of_set T) U" | 
| 76786 | 91 | if "retraction S T r" and "U \<subseteq> T" | 
| 92 | by (simp add: retraction_openin_vimage_iff that) | |
| 68617 | 93 | |
| 94 | lemma retract_of_locally_compact: | |
| 95 |     fixes S :: "'a :: {heine_borel,real_normed_vector} set"
 | |
| 96 | shows "\<lbrakk> locally compact S; T retract_of S\<rbrakk> \<Longrightarrow> locally compact T" | |
| 97 | by (metis locally_compact_closedin closedin_retract) | |
| 98 | ||
| 99 | lemma homotopic_into_retract: | |
| 78457 | 100 | "\<lbrakk>f \<in> S \<rightarrow> T; g \<in> S \<rightarrow> T; T retract_of U; homotopic_with_canon (\<lambda>x. True) S U f g\<rbrakk> | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 101 | \<Longrightarrow> homotopic_with_canon (\<lambda>x. True) S T f g" | 
| 78457 | 102 | apply (subst (asm) homotopic_with_def) | 
| 103 | apply (simp add: homotopic_with retract_of_def retraction_def Pi_iff, clarify) | |
| 104 | apply (rule_tac x="r \<circ> h" in exI) | |
| 105 | by (smt (verit, ccfv_SIG) comp_def continuous_on_compose continuous_on_subset image_subset_iff) | |
| 68617 | 106 | |
| 107 | lemma retract_of_locally_connected: | |
| 108 | assumes "locally connected T" "S retract_of T" | |
| 69661 | 109 | shows "locally connected S" | 
| 68617 | 110 | using assms | 
| 80095 | 111 | by (metis retraction_openin_vimage_iff idempotent_imp_retraction locally_connected_quotient_image retract_ofE) | 
| 68617 | 112 | |
| 113 | lemma retract_of_locally_path_connected: | |
| 114 | assumes "locally path_connected T" "S retract_of T" | |
| 69661 | 115 | shows "locally path_connected S" | 
| 68617 | 116 | using assms | 
| 80095 | 117 | by (metis retraction_openin_vimage_iff idempotent_imp_retraction locally_path_connected_quotient_image retract_ofE) | 
| 68617 | 118 | |
| 119 | text \<open>A few simple lemmas about deformation retracts\<close> | |
| 120 | ||
| 121 | lemma deformation_retract_imp_homotopy_eqv: | |
| 122 | fixes S :: "'a::euclidean_space set" | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 123 | assumes "homotopic_with_canon (\<lambda>x. True) S S id r" and r: "retraction S T r" | 
| 68617 | 124 | shows "S homotopy_eqv T" | 
| 125 | proof - | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 126 | have "homotopic_with_canon (\<lambda>x. True) S S (id \<circ> r) id" | 
| 68617 | 127 | by (simp add: assms(1) homotopic_with_symD) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 128 | moreover have "homotopic_with_canon (\<lambda>x. True) T T (r \<circ> id) id" | 
| 68617 | 129 | using r unfolding retraction_def | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 130 | by (metis eq_id_iff homotopic_with_id2 topspace_euclidean_subtopology) | 
| 68617 | 131 | ultimately | 
| 132 | show ?thesis | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 133 | unfolding homotopy_equivalent_space_def | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 134 | by (smt (verit, del_insts) continuous_map_id continuous_map_subtopology_eu id_def r retraction retraction_comp subset_refl) | 
| 68617 | 135 | qed | 
| 136 | ||
| 137 | lemma deformation_retract: | |
| 138 | fixes S :: "'a::euclidean_space set" | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 139 | shows "(\<exists>r. homotopic_with_canon (\<lambda>x. True) S S id r \<and> retraction S T r) \<longleftrightarrow> | 
| 78457 | 140 | T retract_of S \<and> (\<exists>f. homotopic_with_canon (\<lambda>x. True) S S id f \<and> f \<in> S \<rightarrow> T)" | 
| 68617 | 141 | (is "?lhs = ?rhs") | 
| 142 | proof | |
| 143 | assume ?lhs | |
| 144 | then show ?rhs | |
| 145 | by (auto simp: retract_of_def retraction_def) | |
| 146 | next | |
| 80095 | 147 | assume R: ?rhs | 
| 148 | have "\<And>r f. \<lbrakk>T \<subseteq> S; continuous_on S r; homotopic_with_canon (\<lambda>x. True) S S id f; | |
| 149 | f \<in> S \<rightarrow> T; r \<in> S \<rightarrow> T; \<forall>x\<in>T. r x = x\<rbrakk> | |
| 150 | \<Longrightarrow> homotopic_with_canon (\<lambda>x. True) S S f r" | |
| 151 | apply (rule_tac f = "r \<circ> f" and g="r \<circ> id" in homotopic_with_eq) | |
| 152 | apply (rule_tac Y=S in homotopic_with_compose_continuous_left) | |
| 78457 | 153 | apply (auto simp: homotopic_with_sym Pi_iff) | 
| 68617 | 154 | done | 
| 80095 | 155 | with R homotopic_with_trans show ?lhs | 
| 156 | unfolding retract_of_def retraction_def by blast | |
| 68617 | 157 | qed | 
| 158 | ||
| 159 | lemma deformation_retract_of_contractible_sing: | |
| 160 | fixes S :: "'a::euclidean_space set" | |
| 161 | assumes "contractible S" "a \<in> S" | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 162 |   obtains r where "homotopic_with_canon (\<lambda>x. True) S S id r" "retraction S {a} r"
 | 
| 68617 | 163 | proof - | 
| 164 |   have "{a} retract_of S"
 | |
| 165 | by (simp add: \<open>a \<in> S\<close>) | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 166 | moreover have "homotopic_with_canon (\<lambda>x. True) S S id (\<lambda>x. a)" | 
| 68617 | 167 | using assms | 
| 69738 | 168 | by (auto simp: contractible_def homotopic_into_contractible image_subset_iff) | 
| 78457 | 169 |   moreover have "(\<lambda>x. a) \<in> S \<rightarrow> {a}"
 | 
| 68617 | 170 | by (simp add: image_subsetI) | 
| 171 | ultimately show ?thesis | |
| 78457 | 172 | by (metis that deformation_retract) | 
| 68617 | 173 | qed | 
| 174 | ||
| 175 | ||
| 176 | lemma continuous_on_compact_surface_projection_aux: | |
| 177 | fixes S :: "'a::t2_space set" | |
| 178 | assumes "compact S" "S \<subseteq> T" "image q T \<subseteq> S" | |
| 179 | and contp: "continuous_on T p" | |
| 180 | and "\<And>x. x \<in> S \<Longrightarrow> q x = x" | |
| 181 | and [simp]: "\<And>x. x \<in> T \<Longrightarrow> q(p x) = q x" | |
| 182 | and "\<And>x. x \<in> T \<Longrightarrow> p(q x) = p x" | |
| 183 | shows "continuous_on T q" | |
| 184 | proof - | |
| 185 | have *: "image p T = image p S" | |
| 186 | using assms by auto (metis imageI subset_iff) | |
| 187 | have contp': "continuous_on S p" | |
| 188 | by (rule continuous_on_subset [OF contp \<open>S \<subseteq> T\<close>]) | |
| 189 | have "continuous_on (p ` T) q" | |
| 190 | by (simp add: "*" assms(1) assms(2) assms(5) continuous_on_inv contp' rev_subsetD) | |
| 191 | then have "continuous_on T (q \<circ> p)" | |
| 192 | by (rule continuous_on_compose [OF contp]) | |
| 193 | then show ?thesis | |
| 194 | by (rule continuous_on_eq [of _ "q \<circ> p"]) (simp add: o_def) | |
| 195 | qed | |
| 196 | ||
| 197 | lemma continuous_on_compact_surface_projection: | |
| 198 | fixes S :: "'a::real_normed_vector set" | |
| 199 | assumes "compact S" | |
| 200 |       and S: "S \<subseteq> V - {0}" and "cone V"
 | |
| 201 |       and iff: "\<And>x k. x \<in> V - {0} \<Longrightarrow> 0 < k \<and> (k *\<^sub>R x) \<in> S \<longleftrightarrow> d x = k"
 | |
| 202 |   shows "continuous_on (V - {0}) (\<lambda>x. d x *\<^sub>R x)"
 | |
| 203 | proof (rule continuous_on_compact_surface_projection_aux [OF \<open>compact S\<close> S]) | |
| 204 |   show "(\<lambda>x. d x *\<^sub>R x) ` (V - {0}) \<subseteq> S"
 | |
| 205 | using iff by auto | |
| 206 |   show "continuous_on (V - {0}) (\<lambda>x. inverse(norm x) *\<^sub>R x)"
 | |
| 207 | by (intro continuous_intros) force | |
| 208 | show "\<And>x. x \<in> S \<Longrightarrow> d x *\<^sub>R x = x" | |
| 209 | by (metis S zero_less_one local.iff scaleR_one subset_eq) | |
| 210 |   show "d (x /\<^sub>R norm x) *\<^sub>R (x /\<^sub>R norm x) = d x *\<^sub>R x" if "x \<in> V - {0}" for x
 | |
| 211 | using iff [of "inverse(norm x) *\<^sub>R x" "norm x * d x", symmetric] iff that \<open>cone V\<close> | |
| 212 | by (simp add: field_simps cone_def zero_less_mult_iff) | |
| 213 |   show "d x *\<^sub>R x /\<^sub>R norm (d x *\<^sub>R x) = x /\<^sub>R norm x" if "x \<in> V - {0}" for x
 | |
| 214 | proof - | |
| 215 | have "0 < d x" | |
| 216 | using local.iff that by blast | |
| 217 | then show ?thesis | |
| 218 | by simp | |
| 219 | qed | |
| 220 | qed | |
| 221 | ||
| 222 | subsection \<open>Kuhn Simplices\<close> | |
| 223 | ||
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 224 | lemma bij_betw_singleton_eq: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 225 | assumes f: "bij_betw f A B" and g: "bij_betw g A B" and a: "a \<in> A" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 226 | assumes eq: "(\<And>x. x \<in> A \<Longrightarrow> x \<noteq> a \<Longrightarrow> f x = g x)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 227 | shows "f a = g a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 228 | proof - | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 229 |   have "f ` (A - {a}) = g ` (A - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 230 | by (intro image_cong) (simp_all add: eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 231 |   then have "B - {f a} = B - {g a}"
 | 
| 69286 | 232 | using f g a by (auto simp: bij_betw_def inj_on_image_set_diff set_eq_iff) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 233 | moreover have "f a \<in> B" "g a \<in> B" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 234 | using f g a by (auto simp: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 235 | ultimately show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 236 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 237 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 238 | |
| 63365 | 239 | lemmas swap_apply1 = swap_apply(1) | 
| 240 | lemmas swap_apply2 = swap_apply(2) | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 241 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 242 | lemma pointwise_minimal_pointwise_maximal: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 243 | fixes s :: "(nat \<Rightarrow> nat) set" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 244 | assumes "finite s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 245 |     and "s \<noteq> {}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 246 | and "\<forall>x\<in>s. \<forall>y\<in>s. x \<le> y \<or> y \<le> x" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 247 | shows "\<exists>a\<in>s. \<forall>x\<in>s. a \<le> x" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 248 | and "\<exists>a\<in>s. \<forall>x\<in>s. x \<le> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 249 | using assms | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 250 | proof (induct s rule: finite_ne_induct) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 251 | case (insert b s) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 252 | assume *: "\<forall>x\<in>insert b s. \<forall>y\<in>insert b s. x \<le> y \<or> y \<le> x" | 
| 63540 | 253 | then obtain u l where "l \<in> s" "\<forall>b\<in>s. l \<le> b" "u \<in> s" "\<forall>b\<in>s. b \<le> u" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 254 | using insert by auto | 
| 63540 | 255 | with * show "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. a \<le> x" "\<exists>a\<in>insert b s. \<forall>x\<in>insert b s. x \<le> a" | 
| 76786 | 256 | by (metis insert_iff order.trans)+ | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 257 | qed auto | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 258 | |
| 49555 | 259 | lemma kuhn_labelling_lemma: | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 260 | fixes P Q :: "'a::euclidean_space \<Rightarrow> bool" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 261 | assumes "\<forall>x. P x \<longrightarrow> P (f x)" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 262 | and "\<forall>x. P x \<longrightarrow> (\<forall>i\<in>Basis. Q i \<longrightarrow> 0 \<le> x\<bullet>i \<and> x\<bullet>i \<le> 1)" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 263 | shows "\<exists>l. (\<forall>x.\<forall>i\<in>Basis. l x i \<le> (1::nat)) \<and> | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 264 | (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (x\<bullet>i = 0) \<longrightarrow> (l x i = 0)) \<and> | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 265 | (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (x\<bullet>i = 1) \<longrightarrow> (l x i = 1)) \<and> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 266 | (\<forall>x.\<forall>i\<in>Basis. P x \<and> Q i \<and> (l x i = 0) \<longrightarrow> x\<bullet>i \<le> f x\<bullet>i) \<and> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 267 | (\<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 | 268 | proof - | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 269 |   { fix x i
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 270 | let ?R = "\<lambda>y. (P x \<and> Q i \<and> x \<bullet> i = 0 \<longrightarrow> y = (0::nat)) \<and> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 271 | (P x \<and> Q i \<and> x \<bullet> i = 1 \<longrightarrow> y = 1) \<and> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 272 | (P x \<and> Q i \<and> y = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i) \<and> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 273 | (P x \<and> Q i \<and> y = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 274 |     { assume "P x" "Q i" "i \<in> Basis" with assms have "0 \<le> f x \<bullet> i \<and> f x \<bullet> i \<le> 1" by auto }
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 275 | then have "i \<in> Basis \<Longrightarrow> ?R 0 \<or> ?R 1" by auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 276 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 277 | unfolding all_conj_distrib[symmetric] Ball_def (* FIXME: shouldn't this work by metis? *) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 278 | by (subst choice_iff[symmetric])+ blast | 
| 49374 | 279 | qed | 
| 280 | ||
| 53185 | 281 | |
| 68617 | 282 | subsubsection \<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 | 283 | |
| 53252 | 284 | lemma kuhn_counting_lemma: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 285 | fixes bnd compo compo' face S F | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 286 |   defines "nF s == card {f\<in>F. face f s \<and> compo' f}"
 | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67399diff
changeset | 287 | assumes [simp, intro]: "finite F" \<comment> \<open>faces\<close> and [simp, intro]: "finite S" \<comment> \<open>simplices\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 288 |     and "\<And>f. f \<in> F \<Longrightarrow> bnd f \<Longrightarrow> card {s\<in>S. face f s} = 1"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 289 |     and "\<And>f. f \<in> F \<Longrightarrow> \<not> bnd f \<Longrightarrow> card {s\<in>S. face f s} = 2"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 290 | and "\<And>s. s \<in> S \<Longrightarrow> compo s \<Longrightarrow> nF s = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 291 | and "\<And>s. s \<in> S \<Longrightarrow> \<not> compo s \<Longrightarrow> nF s = 0 \<or> nF s = 2" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 292 |     and "odd (card {f\<in>F. compo' f \<and> bnd f})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 293 |   shows "odd (card {s\<in>S. compo s})"
 | 
| 53185 | 294 | proof - | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 295 | have "(\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) + (\<Sum>s | s \<in> S \<and> compo s. nF s) = (\<Sum>s\<in>S. nF s)" | 
| 64267 | 296 | by (subst sum.union_disjoint[symmetric]) (auto intro!: sum.cong) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 297 |   also have "\<dots> = (\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> bnd f}. face f s}) +
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 298 |                   (\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> \<not> bnd f}. face f s})"
 | 
| 64267 | 299 | unfolding sum.distrib[symmetric] | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 300 | by (subst card_Un_disjoint[symmetric]) | 
| 64267 | 301 | (auto simp: nF_def intro!: sum.cong arg_cong[where f=card]) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 302 |   also have "\<dots> = 1 * card {f\<in>F. compo' f \<and> bnd f} + 2 * card {f\<in>F. compo' f \<and> \<not> bnd f}"
 | 
| 67399 | 303 | using assms(4,5) by (fastforce intro!: arg_cong2[where f="(+)"] sum_multicount) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 304 |   finally have "odd ((\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) + card {s\<in>S. compo s})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 305 | using assms(6,8) by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 306 | moreover have "(\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) = | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 307 | (\<Sum>s | s \<in> S \<and> \<not> compo s \<and> nF s = 0. nF s) + (\<Sum>s | s \<in> S \<and> \<not> compo s \<and> nF s = 2. nF s)" | 
| 64267 | 308 | using assms(7) by (subst sum.union_disjoint[symmetric]) (fastforce intro!: sum.cong)+ | 
| 53688 | 309 | ultimately show ?thesis | 
| 310 | by auto | |
| 53186 | 311 | qed | 
| 312 | ||
| 68617 | 313 | subsubsection \<open>The odd/even result for faces of complete vertices, generalized\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 314 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 315 | lemma kuhn_complete_lemma: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 316 | assumes [simp]: "finite simplices" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 317 |     and face: "\<And>f s. face f s \<longleftrightarrow> (\<exists>a\<in>s. f = s - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 318 | and card_s[simp]: "\<And>s. s \<in> simplices \<Longrightarrow> card s = n + 2" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 319 |     and rl_bd: "\<And>s. s \<in> simplices \<Longrightarrow> rl ` s \<subseteq> {..Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 320 |     and bnd: "\<And>f s. s \<in> simplices \<Longrightarrow> face f s \<Longrightarrow> bnd f \<Longrightarrow> card {s\<in>simplices. face f s} = 1"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 321 |     and nbnd: "\<And>f s. s \<in> simplices \<Longrightarrow> face f s \<Longrightarrow> \<not> bnd f \<Longrightarrow> card {s\<in>simplices. face f s} = 2"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 322 |     and odd_card: "odd (card {f. (\<exists>s\<in>simplices. face f s) \<and> rl ` f = {..n} \<and> bnd f})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 323 |   shows "odd (card {s\<in>simplices. (rl ` s = {..Suc n})})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 324 | proof (rule kuhn_counting_lemma) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 325 | have finite_s[simp]: "\<And>s. s \<in> simplices \<Longrightarrow> finite s" | 
| 72302 
d7d90ed4c74e
fixed some remarkably ugly proofs
 paulson <lp15@cam.ac.uk> parents: 
71745diff
changeset | 326 | by (metis add_is_0 zero_neq_numeral card.infinite assms(3)) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 327 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 328 |   let ?F = "{f. \<exists>s\<in>simplices. face f s}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 329 |   have F_eq: "?F = (\<Union>s\<in>simplices. \<Union>a\<in>s. {s - {a}})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 330 | by (auto simp: face) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 331 | show "finite ?F" | 
| 60420 | 332 | using \<open>finite simplices\<close> unfolding F_eq by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 333 | |
| 60421 | 334 |   show "card {s \<in> simplices. face f s} = 1" if "f \<in> ?F" "bnd f" for f
 | 
| 60449 | 335 | using bnd that by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 336 | |
| 60421 | 337 |   show "card {s \<in> simplices. face f s} = 2" if "f \<in> ?F" "\<not> bnd f" for f
 | 
| 60449 | 338 | using nbnd that by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 339 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 340 |   show "odd (card {f \<in> {f. \<exists>s\<in>simplices. face f s}. rl ` f = {..n} \<and> bnd f})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 341 | using odd_card by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 342 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 343 | fix s assume s[simp]: "s \<in> simplices" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 344 |   let ?S = "{f \<in> {f. \<exists>s\<in>simplices. face f s}. face f s \<and> rl ` f = {..n}}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 345 |   have "?S = (\<lambda>a. s - {a}) ` {a\<in>s. rl ` (s - {a}) = {..n}}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 346 | using s by (fastforce simp: face) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 347 |   then have card_S: "card ?S = card {a\<in>s. rl ` (s - {a}) = {..n}}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 348 | by (auto intro!: card_image inj_onI) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 349 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 350 |   { assume rl: "rl ` s = {..Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 351 | then have inj_rl: "inj_on rl s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 352 | by (intro eq_card_imp_inj_on) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 353 | moreover obtain a where "rl a = Suc n" "a \<in> s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 354 | by (metis atMost_iff image_iff le_Suc_eq rl) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 355 |     ultimately have n: "{..n} = rl ` (s - {a})"
 | 
| 69286 | 356 | by (auto simp: inj_on_image_set_diff rl) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 357 |     have "{a\<in>s. rl ` (s - {a}) = {..n}} = {a}"
 | 
| 69286 | 358 | using inj_rl \<open>a \<in> s\<close> by (auto simp: n inj_on_image_eq_iff[OF inj_rl]) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 359 | then show "card ?S = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 360 | unfolding card_S by simp } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 361 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 362 |   { assume rl: "rl ` s \<noteq> {..Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 363 | show "card ?S = 0 \<or> card ?S = 2" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 364 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 365 |       assume *: "{..n} \<subseteq> rl ` s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 366 |       with rl rl_bd[OF s] have rl_s: "rl ` s = {..n}"
 | 
| 68022 | 367 | by (auto simp: atMost_Suc subset_insert_iff split: if_split_asm) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 368 | then have "\<not> inj_on rl s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 369 | by (intro pigeonhole) simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 370 | then obtain a b where ab: "a \<in> s" "b \<in> s" "rl a = rl b" "a \<noteq> b" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 371 | by (auto simp: inj_on_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 372 |       then have eq: "rl ` (s - {a}) = rl ` s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 373 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 374 |       with ab have inj: "inj_on rl (s - {a})"
 | 
| 68022 | 375 | by (intro eq_card_imp_inj_on) (auto simp: rl_s card_Diff_singleton_if) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 376 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 377 |       { fix x assume "x \<in> s" "x \<notin> {a, b}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 378 |         then have "rl ` s - {rl x} = rl ` ((s - {a}) - {x})"
 | 
| 69286 | 379 | by (auto simp: eq inj_on_image_set_diff[OF inj]) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 380 |         also have "\<dots> = rl ` (s - {x})"
 | 
| 60420 | 381 |           using ab \<open>x \<notin> {a, b}\<close> by auto
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 382 | also assume "\<dots> = rl ` s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 383 | finally have False | 
| 60420 | 384 | using \<open>x\<in>s\<close> by auto } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 385 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 386 |       { fix x assume "x \<in> {a, b}" with ab have "x \<in> s \<and> rl ` (s - {x}) = rl ` s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 387 | by (simp add: set_eq_iff image_iff Bex_def) metis } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 388 |       ultimately have "{a\<in>s. rl ` (s - {a}) = {..n}} = {a, b}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 389 | unfolding rl_s[symmetric] by fastforce | 
| 60420 | 390 | with \<open>a \<noteq> b\<close> show "card ?S = 0 \<or> card ?S = 2" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 391 | unfolding card_S by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 392 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 393 |       assume "\<not> {..n} \<subseteq> rl ` s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 394 |       then have "\<And>x. rl ` (s - {x}) \<noteq> {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 395 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 396 | then show "card ?S = 0 \<or> card ?S = 2" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 397 | unfolding card_S by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 398 | qed } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 399 | qed fact | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 400 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 401 | locale kuhn_simplex = | 
| 78475 | 402 | fixes p n and base upd and S :: "(nat \<Rightarrow> nat) set" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 403 |   assumes base: "base \<in> {..< n} \<rightarrow> {..< p}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 404 | assumes base_out: "\<And>i. n \<le> i \<Longrightarrow> base i = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 405 |   assumes upd: "bij_betw upd {..< n} {..< n}"
 | 
| 78475 | 406 |   assumes s_pre: "S = (\<lambda>i j. if j \<in> upd`{..< i} then Suc (base j) else base j) ` {.. n}"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 407 | begin | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 408 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 409 | definition "enum i j = (if j \<in> upd`{..< i} then Suc (base j) else base j)"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 410 | |
| 78475 | 411 | lemma s_eq: "S = enum ` {.. n}"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 412 | unfolding s_pre enum_def[abs_def] .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 413 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 414 | lemma upd_space: "i < n \<Longrightarrow> upd i < n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 415 | using upd by (auto dest!: bij_betwE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 416 | |
| 78475 | 417 | lemma s_space: "S \<subseteq> {..< n} \<rightarrow> {.. p}"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 418 | proof - | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 419 |   { fix i assume "i \<le> n" then have "enum i \<in> {..< n} \<rightarrow> {.. p}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 420 | proof (induct i) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 421 | case 0 then show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 422 | using base by (auto simp: Pi_iff less_imp_le enum_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 423 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 424 | case (Suc i) with base show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 425 | by (auto simp: Pi_iff Suc_le_eq less_imp_le enum_def intro: upd_space) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 426 | qed } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 427 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 428 | by (auto simp: s_eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 429 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 430 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 431 | lemma inj_upd: "inj_on upd {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 432 | using upd by (simp add: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 433 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 434 | lemma inj_enum: "inj_on enum {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 435 | proof - | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 436 |   { fix x y :: nat assume "x \<noteq> y" "x \<le> n" "y \<le> n"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 437 |     with upd have "upd ` {..< x} \<noteq> upd ` {..< y}"
 | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 438 |       by (subst inj_on_image_eq_iff[where C="{..< n}"]) (auto simp: bij_betw_def)
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 439 | then have "enum x \<noteq> enum y" | 
| 68022 | 440 | by (auto simp: enum_def fun_eq_iff) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 441 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 442 | by (auto simp: inj_on_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 443 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 444 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 445 | lemma enum_0: "enum 0 = base" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 446 | by (simp add: enum_def[abs_def]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 447 | |
| 78475 | 448 | lemma base_in_s: "base \<in> S" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 449 | unfolding s_eq by (subst enum_0[symmetric]) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 450 | |
| 78475 | 451 | lemma enum_in: "i \<le> n \<Longrightarrow> enum i \<in> S" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 452 | unfolding s_eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 453 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 454 | lemma one_step: | 
| 78475 | 455 | assumes a: "a \<in> S" "j < n" | 
| 456 | assumes *: "\<And>a'. a' \<in> S \<Longrightarrow> a' \<noteq> a \<Longrightarrow> a' j = p'" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 457 | shows "a j \<noteq> p'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 458 | proof | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 459 | assume "a j = p'" | 
| 78475 | 460 | with * a have "\<And>a'. a' \<in> S \<Longrightarrow> a' j = p'" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 461 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 462 | then have "\<And>i. i \<le> n \<Longrightarrow> enum i j = p'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 463 | unfolding s_eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 464 |   from this[of 0] this[of n] have "j \<notin> upd ` {..< n}"
 | 
| 62390 | 465 | by (auto simp: enum_def fun_eq_iff split: if_split_asm) | 
| 60420 | 466 | with upd \<open>j < n\<close> show False | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 467 | by (auto simp: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 468 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 469 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 470 | lemma upd_inj: "i < n \<Longrightarrow> j < n \<Longrightarrow> upd i = upd j \<longleftrightarrow> i = j" | 
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61284diff
changeset | 471 | using upd by (auto simp: bij_betw_def inj_on_eq_iff) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 472 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 473 | lemma upd_surj: "upd ` {..< n} = {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 474 | using upd by (auto simp: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 475 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 476 | lemma in_upd_image: "A \<subseteq> {..< n} \<Longrightarrow> i < n \<Longrightarrow> upd i \<in> upd ` A \<longleftrightarrow> i \<in> A"
 | 
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61284diff
changeset | 477 |   using inj_on_image_mem_iff[of upd "{..< n}"] upd
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 478 | by (auto simp: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 479 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 480 | lemma enum_inj: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i = enum j \<longleftrightarrow> i = j" | 
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61284diff
changeset | 481 | using inj_enum by (auto simp: inj_on_eq_iff) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 482 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 483 | lemma in_enum_image: "A \<subseteq> {.. n} \<Longrightarrow> i \<le> n \<Longrightarrow> enum i \<in> enum ` A \<longleftrightarrow> i \<in> A"
 | 
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61284diff
changeset | 484 | using inj_on_image_mem_iff[OF inj_enum] by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 485 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 486 | lemma enum_mono: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i \<le> enum j \<longleftrightarrow> i \<le> j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 487 | by (auto simp: enum_def le_fun_def in_upd_image Ball_def[symmetric]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 488 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 489 | lemma enum_strict_mono: "i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> enum i < enum j \<longleftrightarrow> i < j" | 
| 68022 | 490 | using enum_mono[of i j] enum_inj[of i j] by (auto simp: le_less) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 491 | |
| 78475 | 492 | lemma chain: "a \<in> S \<Longrightarrow> b \<in> S \<Longrightarrow> a \<le> b \<or> b \<le> a" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 493 | by (auto simp: s_eq enum_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 494 | |
| 78475 | 495 | lemma less: "a \<in> S \<Longrightarrow> b \<in> S \<Longrightarrow> a i < b i \<Longrightarrow> a < b" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 496 | using chain[of a b] by (auto simp: less_fun_def le_fun_def not_le[symmetric]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 497 | |
| 78475 | 498 | lemma enum_0_bot: "a \<in> S \<Longrightarrow> a = enum 0 \<longleftrightarrow> (\<forall>a'\<in>S. a \<le> a')" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 499 | unfolding s_eq by (auto simp: enum_mono Ball_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 500 | |
| 78475 | 501 | lemma enum_n_top: "a \<in> S \<Longrightarrow> a = enum n \<longleftrightarrow> (\<forall>a'\<in>S. a' \<le> a)" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 502 | unfolding s_eq by (auto simp: enum_mono Ball_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 503 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 504 | lemma enum_Suc: "i < n \<Longrightarrow> enum (Suc i) = (enum i)(upd i := Suc (enum i (upd i)))" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 505 | by (auto simp: fun_eq_iff enum_def upd_inj) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 506 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 507 | lemma enum_eq_p: "i \<le> n \<Longrightarrow> n \<le> j \<Longrightarrow> enum i j = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 508 | by (induct i) (auto simp: enum_Suc enum_0 base_out upd_space not_less[symmetric]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 509 | |
| 78475 | 510 | lemma out_eq_p: "a \<in> S \<Longrightarrow> n \<le> j \<Longrightarrow> a j = p" | 
| 68022 | 511 | unfolding s_eq by (auto simp: enum_eq_p) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 512 | |
| 78475 | 513 | lemma s_le_p: "a \<in> S \<Longrightarrow> a j \<le> p" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 514 | using out_eq_p[of a j] s_space by (cases "j < n") auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 515 | |
| 78475 | 516 | lemma le_Suc_base: "a \<in> S \<Longrightarrow> a j \<le> Suc (base j)" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 517 | unfolding s_eq by (auto simp: enum_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 518 | |
| 78475 | 519 | lemma base_le: "a \<in> S \<Longrightarrow> base j \<le> a j" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 520 | unfolding s_eq by (auto simp: enum_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 521 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 522 | lemma enum_le_p: "i \<le> n \<Longrightarrow> j < n \<Longrightarrow> enum i j \<le> p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 523 | using enum_in[of i] s_space by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 524 | |
| 78475 | 525 | lemma enum_less: "a \<in> S \<Longrightarrow> i < n \<Longrightarrow> enum i < a \<longleftrightarrow> enum (Suc i) \<le> a" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 526 | unfolding s_eq by (auto simp: enum_strict_mono enum_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 527 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 528 | lemma ksimplex_0: | 
| 78475 | 529 |   "n = 0 \<Longrightarrow> S = {(\<lambda>x. p)}"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 530 | using s_eq enum_def base_out by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 531 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 532 | lemma replace_0: | 
| 78475 | 533 |   assumes "j < n" "a \<in> S" and p: "\<forall>x\<in>S - {a}. x j = 0" and "x \<in> S"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 534 | shows "x \<le> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 535 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 536 | assume "x \<noteq> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 537 | have "a j \<noteq> 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 538 | using assms by (intro one_step[where a=a]) auto | 
| 78475 | 539 | with less[OF \<open>x\<in>S\<close> \<open>a\<in>S\<close>, of j] p[rule_format, of x] \<open>x \<in> S\<close> \<open>x \<noteq> a\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 540 | show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 541 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 542 | qed simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 543 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 544 | lemma replace_1: | 
| 78475 | 545 |   assumes "j < n" "a \<in> S" and p: "\<forall>x\<in>S - {a}. x j = p" and "x \<in> S"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 546 | shows "a \<le> x" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 547 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 548 | assume "x \<noteq> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 549 | have "a j \<noteq> p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 550 | using assms by (intro one_step[where a=a]) auto | 
| 78475 | 551 | with enum_le_p[of _ j] \<open>j < n\<close> \<open>a\<in>S\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 552 | have "a j < p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 553 | by (auto simp: less_le s_eq) | 
| 78475 | 554 | with less[OF \<open>a\<in>S\<close> \<open>x\<in>S\<close>, of j] p[rule_format, of x] \<open>x \<in> S\<close> \<open>x \<noteq> a\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 555 | show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 556 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 557 | qed simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 558 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 559 | end | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 560 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 561 | locale kuhn_simplex_pair = s: kuhn_simplex p n b_s u_s s + t: kuhn_simplex p n b_t u_t t | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 562 | for p n b_s u_s s b_t u_t t | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 563 | begin | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 564 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 565 | lemma enum_eq: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 566 | assumes l: "i \<le> l" "l \<le> j" and "j + d \<le> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 567 |   assumes eq: "s.enum ` {i .. j} = t.enum ` {i + d .. j + d}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 568 | shows "s.enum l = t.enum (l + d)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 569 | using l proof (induct l rule: dec_induct) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 570 | case base | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 571 |   then have s: "s.enum i \<in> t.enum ` {i + d .. j + d}" and t: "t.enum (i + d) \<in> s.enum ` {i .. j}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 572 | using eq by auto | 
| 60420 | 573 | from t \<open>i \<le> j\<close> \<open>j + d \<le> n\<close> have "s.enum i \<le> t.enum (i + d)" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 574 | by (auto simp: s.enum_mono) | 
| 60420 | 575 | moreover from s \<open>i \<le> j\<close> \<open>j + d \<le> n\<close> have "t.enum (i + d) \<le> s.enum i" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 576 | by (auto simp: t.enum_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 577 | ultimately show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 578 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 579 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 580 | case (step l) | 
| 60420 | 581 | moreover from step.prems \<open>j + d \<le> n\<close> have | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 582 | "s.enum l < s.enum (Suc l)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 583 | "t.enum (l + d) < t.enum (Suc l + d)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 584 | by (simp_all add: s.enum_strict_mono t.enum_strict_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 585 | moreover have | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 586 |       "s.enum (Suc l) \<in> t.enum ` {i + d .. j + d}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 587 |       "t.enum (Suc l + d) \<in> s.enum ` {i .. j}"
 | 
| 60420 | 588 | using step \<open>j + d \<le> n\<close> eq by (auto simp: s.enum_inj t.enum_inj) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 589 | ultimately have "s.enum (Suc l) = t.enum (Suc (l + d))" | 
| 60420 | 590 | using \<open>j + d \<le> n\<close> | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 591 | by (intro antisym s.enum_less[THEN iffD1] t.enum_less[THEN iffD1]) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 592 | (auto intro!: s.enum_in t.enum_in) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 593 | then show ?case by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 594 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 595 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 596 | lemma ksimplex_eq_bot: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 597 | assumes a: "a \<in> s" "\<And>a'. a' \<in> s \<Longrightarrow> a \<le> a'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 598 | assumes b: "b \<in> t" "\<And>b'. b' \<in> t \<Longrightarrow> b \<le> b'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 599 |   assumes eq: "s - {a} = t - {b}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 600 | shows "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 601 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 602 | assume "n = 0" with s.ksimplex_0 t.ksimplex_0 show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 603 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 604 | assume "n \<noteq> 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 605 | have "s.enum 0 = (s.enum (Suc 0)) (u_s 0 := s.enum (Suc 0) (u_s 0) - 1)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 606 | "t.enum 0 = (t.enum (Suc 0)) (u_t 0 := t.enum (Suc 0) (u_t 0) - 1)" | 
| 60420 | 607 | using \<open>n \<noteq> 0\<close> by (simp_all add: s.enum_Suc t.enum_Suc) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 608 | moreover have e0: "a = s.enum 0" "b = t.enum 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 609 | using a b by (simp_all add: s.enum_0_bot t.enum_0_bot) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 610 | moreover | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 611 |   { fix j assume "0 < j" "j \<le> n"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 612 |     moreover have "s - {a} = s.enum ` {Suc 0 .. n}" "t - {b} = t.enum ` {Suc 0 .. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 613 | unfolding s.s_eq t.s_eq e0 by (auto simp: s.enum_inj t.enum_inj) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 614 | ultimately have "s.enum j = t.enum j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 615 | using enum_eq[of "1" j n 0] eq by auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 616 | note enum_eq = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 617 | then have "s.enum (Suc 0) = t.enum (Suc 0)" | 
| 60420 | 618 | using \<open>n \<noteq> 0\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 619 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 620 |   { fix j assume "Suc j < n"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 621 | with enum_eq[of "Suc j"] enum_eq[of "Suc (Suc j)"] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 622 | have "u_s (Suc j) = u_t (Suc j)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 623 | using s.enum_Suc[of "Suc j"] t.enum_Suc[of "Suc j"] | 
| 62390 | 624 | by (auto simp: fun_eq_iff split: if_split_asm) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 625 | then have "\<And>j. 0 < j \<Longrightarrow> j < n \<Longrightarrow> u_s j = u_t j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 626 | by (auto simp: gr0_conv_Suc) | 
| 60420 | 627 | with \<open>n \<noteq> 0\<close> have "u_t 0 = u_s 0" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 628 | by (intro bij_betw_singleton_eq[OF t.upd s.upd, of 0]) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 629 | ultimately have "a = b" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 630 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 631 | with assms show "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 632 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 633 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 634 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 635 | lemma ksimplex_eq_top: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 636 | assumes a: "a \<in> s" "\<And>a'. a' \<in> s \<Longrightarrow> a' \<le> a" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 637 | assumes b: "b \<in> t" "\<And>b'. b' \<in> t \<Longrightarrow> b' \<le> b" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 638 |   assumes eq: "s - {a} = t - {b}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 639 | shows "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 640 | proof (cases n) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 641 | assume "n = 0" with s.ksimplex_0 t.ksimplex_0 show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 642 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 643 | case (Suc n') | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 644 | have "s.enum n = (s.enum n') (u_s n' := Suc (s.enum n' (u_s n')))" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 645 | "t.enum n = (t.enum n') (u_t n' := Suc (t.enum n' (u_t n')))" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 646 | using Suc by (simp_all add: s.enum_Suc t.enum_Suc) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 647 | moreover have en: "a = s.enum n" "b = t.enum n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 648 | using a b by (simp_all add: s.enum_n_top t.enum_n_top) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 649 | moreover | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 650 |   { fix j assume "j < n"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 651 |     moreover have "s - {a} = s.enum ` {0 .. n'}" "t - {b} = t.enum ` {0 .. n'}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 652 | unfolding s.s_eq t.s_eq en by (auto simp: s.enum_inj t.enum_inj Suc) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 653 | ultimately have "s.enum j = t.enum j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 654 | using enum_eq[of "0" j n' 0] eq Suc by auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 655 | note enum_eq = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 656 | then have "s.enum n' = t.enum n'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 657 | using Suc by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 658 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 659 |   { fix j assume "j < n'"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 660 | with enum_eq[of j] enum_eq[of "Suc j"] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 661 | have "u_s j = u_t j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 662 | using s.enum_Suc[of j] t.enum_Suc[of j] | 
| 62390 | 663 | by (auto simp: Suc fun_eq_iff split: if_split_asm) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 664 | then have "\<And>j. j < n' \<Longrightarrow> u_s j = u_t j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 665 | by (auto simp: gr0_conv_Suc) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 666 | then have "u_t n' = u_s n'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 667 | by (intro bij_betw_singleton_eq[OF t.upd s.upd, of n']) (auto simp: Suc) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 668 | ultimately have "a = b" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 669 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 670 | with assms show "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 671 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 672 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 673 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 674 | end | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 675 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 676 | inductive ksimplex for p n :: nat where | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 677 | ksimplex: "kuhn_simplex p n base upd s \<Longrightarrow> ksimplex p n s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 678 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 679 | lemma finite_ksimplexes: "finite {s. ksimplex p n s}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 680 | proof (rule finite_subset) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 681 |   { fix a s assume "ksimplex p n s" "a \<in> s"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 682 | then obtain b u where "kuhn_simplex p n b u s" by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 683 | then interpret kuhn_simplex p n b u s . | 
| 60420 | 684 | from s_space \<open>a \<in> s\<close> out_eq_p[OF \<open>a \<in> s\<close>] | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 685 |     have "a \<in> (\<lambda>f x. if n \<le> x then p else f x) ` ({..< n} \<rightarrow>\<^sub>E {.. p})"
 | 
| 62390 | 686 | by (auto simp: image_iff subset_eq Pi_iff split: if_split_asm | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 687 |                intro!: bexI[of _ "restrict a {..< n}"]) }
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 688 |   then show "{s. ksimplex p n s} \<subseteq> Pow ((\<lambda>f x. if n \<le> x then p else f x) ` ({..< n} \<rightarrow>\<^sub>E {.. p}))"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 689 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 690 | qed (simp add: finite_PiE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 691 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 692 | lemma ksimplex_card: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 693 | assumes "ksimplex p n s" shows "card s = Suc n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 694 | using assms proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 695 | case (ksimplex u b) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 696 | then interpret kuhn_simplex p n u b s . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 697 | show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 698 | by (simp add: card_image s_eq inj_enum) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 699 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 700 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 701 | lemma simplex_top_face: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 702 | assumes "0 < p" "\<forall>x\<in>s'. x n = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 703 |   shows "ksimplex p n s' \<longleftrightarrow> (\<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> s' = s - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 704 | using assms | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 705 | proof safe | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 706 |   fix s a assume "ksimplex p (Suc n) s" and a: "a \<in> s" and na: "\<forall>x\<in>s - {a}. x n = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 707 |   then show "ksimplex p n (s - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 708 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 709 | case (ksimplex base upd) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 710 | then interpret kuhn_simplex p "Suc n" base upd "s" . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 711 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 712 | have "a n < p" | 
| 60420 | 713 | using one_step[of a n p] na \<open>a\<in>s\<close> s_space by (auto simp: less_le) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 714 | then have "a = enum 0" | 
| 60420 | 715 | using \<open>a \<in> s\<close> na by (subst enum_0_bot) (auto simp: le_less intro!: less[of a _ n]) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 716 |     then have s_eq: "s - {a} = enum ` Suc ` {.. n}"
 | 
| 71172 | 717 | using s_eq by (simp add: atMost_Suc_eq_insert_0 insert_ident in_enum_image subset_eq) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 718 |     then have "enum 1 \<in> s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 719 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 720 | then have "upd 0 = n" | 
| 60420 | 721 | using \<open>a n < p\<close> \<open>a = enum 0\<close> na[rule_format, of "enum 1"] | 
| 62390 | 722 | by (auto simp: fun_eq_iff enum_Suc split: if_split_asm) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 723 |     then have "bij_betw upd (Suc ` {..< n}) {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 724 | using upd | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 725 | by (subst notIn_Un_bij_betw3[where b=0]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 726 | (auto simp: lessThan_Suc[symmetric] lessThan_Suc_eq_insert_0) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 727 |     then have "bij_betw (upd\<circ>Suc) {..<n} {..<n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 728 | by (rule bij_betw_trans[rotated]) (auto simp: bij_betw_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 729 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 730 | have "a n = p - 1" | 
| 60420 | 731 |       using enum_Suc[of 0] na[rule_format, OF \<open>enum 1 \<in> s - {a}\<close>] \<open>a = enum 0\<close> by (auto simp: \<open>upd 0 = n\<close>)
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 732 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 733 | show ?thesis | 
| 61169 | 734 | proof (rule ksimplex.intros, standard) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 735 |       show "bij_betw (upd\<circ>Suc) {..< n} {..< n}" by fact
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 736 |       show "base(n := p) \<in> {..<n} \<rightarrow> {..<p}" "\<And>i. n\<le>i \<Longrightarrow> (base(n := p)) i = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 737 | using base base_out by (auto simp: Pi_iff) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 738 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 739 |       have "\<And>i. Suc ` {..< i} = {..< Suc i} - {0}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 740 | by (auto simp: image_iff Ball_def) arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 741 |       then have upd_Suc: "\<And>i. i \<le> n \<Longrightarrow> (upd\<circ>Suc) ` {..< i} = upd ` {..< Suc i} - {n}"
 | 
| 69661 | 742 | using \<open>upd 0 = n\<close> upd_inj by (auto simp add: image_iff less_Suc_eq_0_disj) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 743 |       have n_in_upd: "\<And>i. n \<in> upd ` {..< Suc i}"
 | 
| 60420 | 744 | using \<open>upd 0 = n\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 745 | |
| 63040 | 746 | define f' where "f' i j = | 
| 747 |         (if j \<in> (upd\<circ>Suc)`{..< i} then Suc ((base(n := p)) j) else (base(n := p)) j)" for i j
 | |
| 69661 | 748 |       { fix x i
 | 
| 749 | assume i [arith]: "i \<le> n" | |
| 750 |         with upd_Suc have "(upd \<circ> Suc) ` {..<i} = upd ` {..<Suc i} - {n}" .
 | |
| 751 | with \<open>a n < p\<close> \<open>a = enum 0\<close> \<open>upd 0 = n\<close> \<open>a n = p - 1\<close> | |
| 752 | have "enum (Suc i) x = f' i x" | |
| 753 | by (auto simp add: f'_def enum_def) } | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 754 |       then show "s - {a} = f' ` {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 755 | unfolding s_eq image_comp by (intro image_cong) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 756 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 757 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 758 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 759 | assume "ksimplex p n s'" and *: "\<forall>x\<in>s'. x n = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 760 |   then show "\<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> s' = s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 761 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 762 | case (ksimplex base upd) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 763 | then interpret kuhn_simplex p n base upd s' . | 
| 63040 | 764 | define b where "b = base (n := p - 1)" | 
| 765 | define u where "u i = (case i of 0 \<Rightarrow> n | Suc i \<Rightarrow> upd i)" for i | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 766 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 767 |     have "ksimplex p (Suc n) (s' \<union> {b})"
 | 
| 61169 | 768 | proof (rule ksimplex.intros, standard) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 769 |       show "b \<in> {..<Suc n} \<rightarrow> {..<p}"
 | 
| 60420 | 770 | using base \<open>0 < p\<close> unfolding lessThan_Suc b_def by (auto simp: PiE_iff) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 771 | show "\<And>i. Suc n \<le> i \<Longrightarrow> b i = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 772 | using base_out by (auto simp: b_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 773 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 774 |       have "bij_betw u (Suc ` {..< n} \<union> {0}) ({..<n} \<union> {u 0})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 775 | using upd | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 776 | by (intro notIn_Un_bij_betw) (auto simp: u_def bij_betw_def image_comp comp_def inj_on_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 777 |       then show "bij_betw u {..<Suc n} {..<Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 778 | by (simp add: u_def lessThan_Suc[symmetric] lessThan_Suc_eq_insert_0) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 779 | |
| 63040 | 780 |       define f' where "f' i j = (if j \<in> u`{..< i} then Suc (b j) else b j)" for i j
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 781 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 782 |       have u_eq: "\<And>i. i \<le> n \<Longrightarrow> u ` {..< Suc i} = upd ` {..< i} \<union> { n }"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 783 | by (auto simp: u_def image_iff upd_inj Ball_def split: nat.split) arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 784 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 785 |       { fix x have "x \<le> n \<Longrightarrow> n \<notin> upd ` {..<x}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 786 | using upd_space by (simp add: image_iff neq_iff) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 787 | note n_not_upd = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 788 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 789 |       have *: "f' ` {.. Suc n} = f' ` (Suc ` {.. n} \<union> {0})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 790 | unfolding atMost_Suc_eq_insert_0 by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 791 |       also have "\<dots> = (f' \<circ> Suc) ` {.. n} \<union> {b}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 792 | by (auto simp: f'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 793 |       also have "(f' \<circ> Suc) ` {.. n} = s'"
 | 
| 60420 | 794 | using \<open>0 < p\<close> base_out[of n] | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 795 | unfolding s_eq enum_def[abs_def] f'_def[abs_def] upd_space | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 796 | by (intro image_cong) (simp_all add: u_eq b_def fun_eq_iff n_not_upd) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 797 |       finally show "s' \<union> {b} = f' ` {.. Suc n}" ..
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 798 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 799 | moreover have "b \<notin> s'" | 
| 60420 | 800 | using * \<open>0 < p\<close> by (auto simp: b_def) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 801 | ultimately show ?thesis by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 802 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 803 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 804 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 805 | lemma ksimplex_replace_0: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 806 | assumes s: "ksimplex p n s" and a: "a \<in> s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 807 |   assumes j: "j < n" and p: "\<forall>x\<in>s - {a}. x j = 0"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 808 |   shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 1"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 809 | using s | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 810 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 811 | case (ksimplex b_s u_s) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 812 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 813 |   { fix t b assume "ksimplex p n t"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 814 | then obtain b_t u_t where "kuhn_simplex p n b_t u_t t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 815 | by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 816 | interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 817 | by intro_locales fact+ | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 818 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 819 |     assume b: "b \<in> t" "t - {b} = s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 820 | with a j p s.replace_0[of _ a] t.replace_0[of _ b] have "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 821 | by (intro ksimplex_eq_top[of a b]) auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 822 |   then have "{s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = {s}"
 | 
| 60420 | 823 | using s \<open>a \<in> s\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 824 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 825 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 826 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 827 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 828 | lemma ksimplex_replace_1: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 829 | assumes s: "ksimplex p n s" and a: "a \<in> s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 830 |   assumes j: "j < n" and p: "\<forall>x\<in>s - {a}. x j = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 831 |   shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 1"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 832 | using s | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 833 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 834 | case (ksimplex b_s u_s) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 835 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 836 |   { fix t b assume "ksimplex p n t"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 837 | then obtain b_t u_t where "kuhn_simplex p n b_t u_t t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 838 | by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 839 | interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 840 | by intro_locales fact+ | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 841 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 842 |     assume b: "b \<in> t" "t - {b} = s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 843 | with a j p s.replace_1[of _ a] t.replace_1[of _ b] have "s = t" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 844 | by (intro ksimplex_eq_bot[of a b]) auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 845 |   then have "{s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = {s}"
 | 
| 60420 | 846 | using s \<open>a \<in> s\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 847 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 848 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 849 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 850 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 851 | lemma ksimplex_replace_2: | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 852 | assumes s: "ksimplex p n s" and "a \<in> s" and "n \<noteq> 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 853 |     and lb: "\<forall>j<n. \<exists>x\<in>s - {a}. x j \<noteq> 0"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 854 |     and ub: "\<forall>j<n. \<exists>x\<in>s - {a}. x j \<noteq> p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 855 |   shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 2"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 856 | using s | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 857 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 858 | case (ksimplex base upd) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 859 | then interpret kuhn_simplex p n base upd s . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 860 | |
| 60420 | 861 | from \<open>a \<in> s\<close> obtain i where "i \<le> n" "a = enum i" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 862 | unfolding s_eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 863 | |
| 60420 | 864 | from \<open>i \<le> n\<close> have "i = 0 \<or> i = n \<or> (0 < i \<and> i < n)" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 865 | by linarith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 866 |   then have "\<exists>!s'. s' \<noteq> s \<and> ksimplex p n s' \<and> (\<exists>b\<in>s'. s - {a} = s'- {b})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 867 | proof (elim disjE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 868 | assume "i = 0" | 
| 63040 | 869 | define rot where [abs_def]: "rot i = (if i + 1 = n then 0 else i + 1)" for i | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 870 | let ?upd = "upd \<circ> rot" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 871 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 872 |     have rot: "bij_betw rot {..< n} {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 873 | by (auto simp: bij_betw_def inj_on_def image_iff Ball_def rot_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 874 | arith+ | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 875 |     from rot upd have "bij_betw ?upd {..<n} {..<n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 876 | by (rule bij_betw_trans) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 877 | |
| 63040 | 878 | define f' where [abs_def]: "f' i j = | 
| 879 |       (if j \<in> ?upd`{..< i} then Suc (enum (Suc 0) j) else enum (Suc 0) j)" for i j
 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 880 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 881 |     interpret b: kuhn_simplex p n "enum (Suc 0)" "upd \<circ> rot" "f' ` {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 882 | proof | 
| 60420 | 883 | from \<open>a = enum i\<close> ub \<open>n \<noteq> 0\<close> \<open>i = 0\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 884 | obtain i' where "i' \<le> n" "enum i' \<noteq> enum 0" "enum i' (upd 0) \<noteq> p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 885 | unfolding s_eq by (auto intro: upd_space simp: enum_inj) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 886 | then have "enum 1 \<le> enum i'" "enum i' (upd 0) < p" | 
| 68022 | 887 | using enum_le_p[of i' "upd 0"] by (auto simp: enum_inj enum_mono upd_space) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 888 | then have "enum 1 (upd 0) < p" | 
| 68022 | 889 | by (auto simp: le_fun_def intro: le_less_trans) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 890 |       then show "enum (Suc 0) \<in> {..<n} \<rightarrow> {..<p}"
 | 
| 68022 | 891 | using base \<open>n \<noteq> 0\<close> by (auto simp: enum_0 enum_Suc PiE_iff extensional_def upd_space) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 892 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 893 |       { fix i assume "n \<le> i" then show "enum (Suc 0) i = p"
 | 
| 60420 | 894 | using \<open>n \<noteq> 0\<close> by (auto simp: enum_eq_p) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 895 |       show "bij_betw ?upd {..<n} {..<n}" by fact
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 896 | qed (simp add: f'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 897 |     have ks_f': "ksimplex p n (f' ` {.. n})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 898 | by rule unfold_locales | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 899 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 900 | have b_enum: "b.enum = f'" unfolding f'_def b.enum_def[abs_def] .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 901 |     with b.inj_enum have inj_f': "inj_on f' {.. n}" by simp
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 902 | |
| 69661 | 903 | have f'_eq_enum: "f' j = enum (Suc j)" if "j < n" for j | 
| 904 | proof - | |
| 905 |       from that have "rot ` {..< j} = {0 <..< Suc j}"
 | |
| 906 | by (auto simp: rot_def image_Suc_lessThan cong: image_cong_simp) | |
| 907 | with that \<open>n \<noteq> 0\<close> show ?thesis | |
| 908 | by (simp only: f'_def enum_def fun_eq_iff image_comp [symmetric]) | |
| 909 | (auto simp add: upd_inj) | |
| 910 | qed | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 911 |     then have "enum ` Suc ` {..< n} = f' ` {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 912 | by (force simp: enum_inj) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 913 |     also have "Suc ` {..< n} = {.. n} - {0}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 914 | by (auto simp: image_iff Ball_def) arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 915 |     also have "{..< n} = {.. n} - {n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 916 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 917 |     finally have eq: "s - {a} = f' ` {.. n} - {f' n}"
 | 
| 60420 | 918 | unfolding s_eq \<open>a = enum i\<close> \<open>i = 0\<close> | 
| 69286 | 919 | by (simp add: inj_on_image_set_diff[OF inj_enum] inj_on_image_set_diff[OF inj_f']) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 920 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 921 | have "enum 0 < f' 0" | 
| 60420 | 922 | using \<open>n \<noteq> 0\<close> by (simp add: enum_strict_mono f'_eq_enum) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 923 | also have "\<dots> < f' n" | 
| 60420 | 924 | using \<open>n \<noteq> 0\<close> b.enum_strict_mono[of 0 n] unfolding b_enum by simp | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 925 | finally have "a \<noteq> f' n" | 
| 60420 | 926 | using \<open>a = enum i\<close> \<open>i = 0\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 927 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 928 |     { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 929 | obtain b u where "kuhn_simplex p n b u t" | 
| 60420 | 930 | using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 931 | then interpret t: kuhn_simplex p n b u t . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 932 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 933 |       { fix x assume "x \<in> s" "x \<noteq> a"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 934 | then have "x (upd 0) = enum (Suc 0) (upd 0)" | 
| 60420 | 935 | by (auto simp: \<open>a = enum i\<close> \<open>i = 0\<close> s_eq enum_def enum_inj) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 936 |       then have eq_upd0: "\<forall>x\<in>t-{c}. x (upd 0) = enum (Suc 0) (upd 0)"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 937 | unfolding eq_sma[symmetric] by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 938 | then have "c (upd 0) \<noteq> enum (Suc 0) (upd 0)" | 
| 60420 | 939 | using \<open>n \<noteq> 0\<close> by (intro t.one_step[OF \<open>c\<in>t\<close> ]) (auto simp: upd_space) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 940 | then have "c (upd 0) < enum (Suc 0) (upd 0) \<or> c (upd 0) > enum (Suc 0) (upd 0)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 941 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 942 |       then have "t = s \<or> t = f' ` {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 943 | proof (elim disjE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 944 | assume *: "c (upd 0) < enum (Suc 0) (upd 0)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 945 | interpret st: kuhn_simplex_pair p n base upd s b u t .. | 
| 60420 | 946 |         { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "c \<le> x"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 947 | by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 948 | note top = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 949 | have "s = t" | 
| 60420 | 950 | using \<open>a = enum i\<close> \<open>i = 0\<close> \<open>c \<in> t\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 951 | by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq_sma]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 952 | (auto simp: s_eq enum_mono t.s_eq t.enum_mono top) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 953 | then show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 954 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 955 | assume *: "c (upd 0) > enum (Suc 0) (upd 0)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 956 |         interpret st: kuhn_simplex_pair p n "enum (Suc 0)" "upd \<circ> rot" "f' ` {.. n}" b u t ..
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 957 |         have eq: "f' ` {..n} - {f' n} = t - {c}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 958 | using eq_sma eq by simp | 
| 60420 | 959 |         { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "x \<le> c"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 960 | by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 961 | note top = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 962 |         have "f' ` {..n} = t"
 | 
| 60420 | 963 | using \<open>a = enum i\<close> \<open>i = 0\<close> \<open>c \<in> t\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 964 | by (intro st.ksimplex_eq_top[OF _ _ _ _ eq]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 965 | (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono b_enum[symmetric] top) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 966 | then show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 967 | qed } | 
| 60420 | 968 | with ks_f' eq \<open>a \<noteq> f' n\<close> \<open>n \<noteq> 0\<close> show ?thesis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 969 |       apply (intro ex1I[of _ "f' ` {.. n}"])
 | 
| 80095 | 970 | apply auto [] | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 971 | apply metis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 972 | done | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 973 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 974 | assume "i = n" | 
| 60420 | 975 | from \<open>n \<noteq> 0\<close> obtain n' where n': "n = Suc n'" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 976 | by (cases n) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 977 | |
| 63040 | 978 | define rot where "rot i = (case i of 0 \<Rightarrow> n' | Suc i \<Rightarrow> i)" for i | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 979 | let ?upd = "upd \<circ> rot" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 980 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 981 |     have rot: "bij_betw rot {..< n} {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 982 | by (auto simp: bij_betw_def inj_on_def image_iff Bex_def rot_def n' split: nat.splits) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 983 | arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 984 |     from rot upd have "bij_betw ?upd {..<n} {..<n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 985 | by (rule bij_betw_trans) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 986 | |
| 63040 | 987 | define b where "b = base (upd n' := base (upd n') - 1)" | 
| 988 |     define f' where [abs_def]: "f' i j = (if j \<in> ?upd`{..< i} then Suc (b j) else b j)" for i j
 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 989 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 990 |     interpret b: kuhn_simplex p n b "upd \<circ> rot" "f' ` {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 991 | proof | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 992 |       { fix i assume "n \<le> i" then show "b i = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 993 | using base_out[of i] upd_space[of n'] by (auto simp: b_def n') } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 994 |       show "b \<in> {..<n} \<rightarrow> {..<p}"
 | 
| 60420 | 995 | using base \<open>n \<noteq> 0\<close> upd_space[of n'] | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 996 | by (auto simp: b_def PiE_def Pi_iff Ball_def upd_space extensional_def n') | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 997 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 998 |       show "bij_betw ?upd {..<n} {..<n}" by fact
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 999 | qed (simp add: f'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1000 | have f': "b.enum = f'" unfolding f'_def b.enum_def[abs_def] .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1001 |     have ks_f': "ksimplex p n (b.enum ` {.. n})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1002 | unfolding f' by rule unfold_locales | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1003 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 1004 | have "0 < n" | 
| 60420 | 1005 | using \<open>n \<noteq> 0\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1006 | |
| 60420 | 1007 |     { from \<open>a = enum i\<close> \<open>n \<noteq> 0\<close> \<open>i = n\<close> lb upd_space[of n']
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1008 | obtain i' where "i' \<le> n" "enum i' \<noteq> enum n" "0 < enum i' (upd n')" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1009 | unfolding s_eq by (auto simp: enum_inj n') | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1010 | moreover have "enum i' (upd n') = base (upd n')" | 
| 60420 | 1011 | unfolding enum_def using \<open>i' \<le> n\<close> \<open>enum i' \<noteq> enum n\<close> by (auto simp: n' upd_inj enum_inj) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1012 | ultimately have "0 < base (upd n')" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1013 | by auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1014 | then have benum1: "b.enum (Suc 0) = base" | 
| 60420 | 1015 | unfolding b.enum_Suc[OF \<open>0<n\<close>] b.enum_0 by (auto simp: b_def rot_def) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1016 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1017 |     have [simp]: "\<And>j. Suc j < n \<Longrightarrow> rot ` {..< Suc j} = {n'} \<union> {..< j}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1018 | by (auto simp: rot_def image_iff Ball_def split: nat.splits) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1019 | have rot_simps: "\<And>j. rot (Suc j) = j" "rot 0 = n'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1020 | by (simp_all add: rot_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1021 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1022 |     { fix j assume j: "Suc j \<le> n" then have "b.enum (Suc j) = enum j"
 | 
| 68022 | 1023 | by (induct j) (auto simp: benum1 enum_0 b.enum_Suc enum_Suc rot_simps) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1024 | note b_enum_eq_enum = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1025 |     then have "enum ` {..< n} = b.enum ` Suc ` {..< n}"
 | 
| 68022 | 1026 | by (auto simp: image_comp intro!: image_cong) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1027 |     also have "Suc ` {..< n} = {.. n} - {0}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1028 | by (auto simp: image_iff Ball_def) arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1029 |     also have "{..< n} = {.. n} - {n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1030 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1031 |     finally have eq: "s - {a} = b.enum ` {.. n} - {b.enum 0}"
 | 
| 60420 | 1032 | unfolding s_eq \<open>a = enum i\<close> \<open>i = n\<close> | 
| 60303 | 1033 |       using inj_on_image_set_diff[OF inj_enum Diff_subset, of "{n}"]
 | 
| 1034 |             inj_on_image_set_diff[OF b.inj_enum Diff_subset, of "{0}"]
 | |
| 68022 | 1035 | by (simp add: comp_def) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1036 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1037 | have "b.enum 0 \<le> b.enum n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1038 | by (simp add: b.enum_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1039 | also have "b.enum n < enum n" | 
| 60420 | 1040 | using \<open>n \<noteq> 0\<close> by (simp add: enum_strict_mono b_enum_eq_enum n') | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1041 | finally have "a \<noteq> b.enum 0" | 
| 60420 | 1042 | using \<open>a = enum i\<close> \<open>i = n\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1043 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1044 |     { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1045 | obtain b' u where "kuhn_simplex p n b' u t" | 
| 60420 | 1046 | using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1047 | then interpret t: kuhn_simplex p n b' u t . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1048 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1049 |       { fix x assume "x \<in> s" "x \<noteq> a"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1050 | then have "x (upd n') = enum n' (upd n')" | 
| 60420 | 1051 | by (auto simp: \<open>a = enum i\<close> n' \<open>i = n\<close> s_eq enum_def enum_inj in_upd_image) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1052 |       then have eq_upd0: "\<forall>x\<in>t-{c}. x (upd n') = enum n' (upd n')"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1053 | unfolding eq_sma[symmetric] by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1054 | then have "c (upd n') \<noteq> enum n' (upd n')" | 
| 60420 | 1055 | using \<open>n \<noteq> 0\<close> by (intro t.one_step[OF \<open>c\<in>t\<close> ]) (auto simp: n' upd_space[unfolded n']) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1056 | then have "c (upd n') < enum n' (upd n') \<or> c (upd n') > enum n' (upd n')" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1057 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1058 |       then have "t = s \<or> t = b.enum ` {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1059 | proof (elim disjE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1060 | assume *: "c (upd n') > enum n' (upd n')" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1061 | interpret st: kuhn_simplex_pair p n base upd s b' u t .. | 
| 60420 | 1062 |         { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "x \<le> c"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1063 | by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1064 | note top = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1065 | have "s = t" | 
| 60420 | 1066 | using \<open>a = enum i\<close> \<open>i = n\<close> \<open>c \<in> t\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1067 | by (intro st.ksimplex_eq_top[OF _ _ _ _ eq_sma]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1068 | (auto simp: s_eq enum_mono t.s_eq t.enum_mono top) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1069 | then show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1070 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1071 | assume *: "c (upd n') < enum n' (upd n')" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1072 |         interpret st: kuhn_simplex_pair p n b "upd \<circ> rot" "f' ` {.. n}" b' u t ..
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1073 |         have eq: "f' ` {..n} - {b.enum 0} = t - {c}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1074 | using eq_sma eq f' by simp | 
| 60420 | 1075 |         { fix x assume "x \<in> t" with * \<open>c\<in>t\<close> eq_upd0[rule_format, of x] have "c \<le> x"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1076 | by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1077 | note bot = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1078 |         have "f' ` {..n} = t"
 | 
| 60420 | 1079 | using \<open>a = enum i\<close> \<open>i = n\<close> \<open>c \<in> t\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1080 | by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq]) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1081 | (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono bot) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1082 | with f' show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1083 | qed } | 
| 60420 | 1084 | with ks_f' eq \<open>a \<noteq> b.enum 0\<close> \<open>n \<noteq> 0\<close> show ?thesis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1085 |       apply (intro ex1I[of _ "b.enum ` {.. n}"])
 | 
| 76786 | 1086 | apply fastforce | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1087 | apply metis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1088 | done | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1089 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1090 | assume i: "0 < i" "i < n" | 
| 63040 | 1091 | define i' where "i' = i - 1" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1092 | with i have "Suc i' < n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1093 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1094 | with i have Suc_i': "Suc i' = i" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1095 | by (simp add: i'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1096 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1097 | let ?upd = "Fun.swap i' i upd" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1098 |     from i upd have "bij_betw ?upd {..< n} {..< n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1099 | by (subst bij_betw_swap_iff) (auto simp: i'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1100 | |
| 63040 | 1101 |     define f' where [abs_def]: "f' i j = (if j \<in> ?upd`{..< i} then Suc (base j) else base j)"
 | 
| 1102 | for i j | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1103 |     interpret b: kuhn_simplex p n base ?upd "f' ` {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1104 | proof | 
| 67682 
00c436488398
tuned proofs -- prefer explicit names for facts from 'interpret';
 wenzelm parents: 
67673diff
changeset | 1105 |       show "base \<in> {..<n} \<rightarrow> {..<p}" by (rule base)
 | 
| 
00c436488398
tuned proofs -- prefer explicit names for facts from 'interpret';
 wenzelm parents: 
67673diff
changeset | 1106 |       { fix i assume "n \<le> i" then show "base i = p" by (rule base_out) }
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1107 |       show "bij_betw ?upd {..<n} {..<n}" by fact
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1108 | qed (simp add: f'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1109 | have f': "b.enum = f'" unfolding f'_def b.enum_def[abs_def] .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1110 |     have ks_f': "ksimplex p n (b.enum ` {.. n})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1111 | unfolding f' by rule unfold_locales | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1112 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1113 |     have "{i} \<subseteq> {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1114 | using i by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1115 |     { fix j assume "j \<le> n"
 | 
| 76786 | 1116 | with i Suc_i' have "enum j = b.enum j \<longleftrightarrow> j \<noteq> i" | 
| 1117 | unfolding fun_eq_iff enum_def b.enum_def image_comp [symmetric] | |
| 73648 | 1118 | apply (cases \<open>i = j\<close>) | 
| 76786 | 1119 | apply (metis imageI in_upd_image lessI lessThan_iff lessThan_subset_iff order_less_le transpose_apply_first) | 
| 1120 | by (metis lessThan_iff linorder_not_less not_less_eq_eq order_less_le transpose_image_eq) | |
| 73648 | 1121 | } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1122 | note enum_eq_benum = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1123 |     then have "enum ` ({.. n} - {i}) = b.enum ` ({.. n} - {i})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1124 | by (intro image_cong) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1125 |     then have eq: "s - {a} = b.enum ` {.. n} - {b.enum i}"
 | 
| 60420 | 1126 | unfolding s_eq \<open>a = enum i\<close> | 
| 1127 |       using inj_on_image_set_diff[OF inj_enum Diff_subset \<open>{i} \<subseteq> {..n}\<close>]
 | |
| 1128 |             inj_on_image_set_diff[OF b.inj_enum Diff_subset \<open>{i} \<subseteq> {..n}\<close>]
 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1129 | by (simp add: comp_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1130 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1131 | have "a \<noteq> b.enum i" | 
| 60420 | 1132 | using \<open>a = enum i\<close> enum_eq_benum i by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1133 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1134 |     { fix t c assume "ksimplex p n t" "c \<in> t" and eq_sma: "s - {a} = t - {c}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1135 | obtain b' u where "kuhn_simplex p n b' u t" | 
| 60420 | 1136 | using \<open>ksimplex p n t\<close> by (auto elim: ksimplex.cases) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1137 | then interpret t: kuhn_simplex p n b' u t . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1138 |       have "enum i' \<in> s - {a}" "enum (i + 1) \<in> s - {a}"
 | 
| 60420 | 1139 | using \<open>a = enum i\<close> i enum_in by (auto simp: enum_inj i'_def) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1140 | then obtain l k where | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1141 | l: "t.enum l = enum i'" "l \<le> n" "t.enum l \<noteq> c" and | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1142 | k: "t.enum k = enum (i + 1)" "k \<le> n" "t.enum k \<noteq> c" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1143 | unfolding eq_sma by (auto simp: t.s_eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1144 | with i have "t.enum l < t.enum k" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1145 | by (simp add: enum_strict_mono i'_def) | 
| 60420 | 1146 | with \<open>l \<le> n\<close> \<open>k \<le> n\<close> have "l < k" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1147 | by (simp add: t.enum_strict_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1148 |       { assume "Suc l = k"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1149 | have "enum (Suc (Suc i')) = t.enum (Suc l)" | 
| 60420 | 1150 | using i by (simp add: k \<open>Suc l = k\<close> i'_def) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1151 | then have False | 
| 60420 | 1152 | using \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close> | 
| 62390 | 1153 | by (auto simp: t.enum_Suc enum_Suc l upd_inj fun_eq_iff split: if_split_asm) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1154 | (metis Suc_lessD n_not_Suc_n upd_inj) } | 
| 60420 | 1155 | with \<open>l < k\<close> have "Suc l < k" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1156 | by arith | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1157 | have c_eq: "c = t.enum (Suc l)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1158 | proof (rule ccontr) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1159 | assume "c \<noteq> t.enum (Suc l)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1160 |         then have "t.enum (Suc l) \<in> s - {a}"
 | 
| 60420 | 1161 | using \<open>l < k\<close> \<open>k \<le> n\<close> by (simp add: t.s_eq eq_sma) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1162 | then obtain j where "t.enum (Suc l) = enum j" "j \<le> n" "enum j \<noteq> enum i" | 
| 60420 | 1163 | unfolding s_eq \<open>a = enum i\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1164 | with i have "t.enum (Suc l) \<le> t.enum l \<or> t.enum k \<le> t.enum (Suc l)" | 
| 68022 | 1165 | by (auto simp: i'_def enum_mono enum_inj l k) | 
| 60420 | 1166 | with \<open>Suc l < k\<close> \<open>k \<le> n\<close> show False | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1167 | by (simp add: t.enum_mono) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1168 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1169 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1170 |       { have "t.enum (Suc (Suc l)) \<in> s - {a}"
 | 
| 60420 | 1171 | unfolding eq_sma c_eq t.s_eq using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (auto simp: t.enum_inj) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1172 | then obtain j where eq: "t.enum (Suc (Suc l)) = enum j" and "j \<le> n" "j \<noteq> i" | 
| 60420 | 1173 | by (auto simp: s_eq \<open>a = enum i\<close>) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1174 | moreover have "enum i' < t.enum (Suc (Suc l))" | 
| 60420 | 1175 | unfolding l(1)[symmetric] using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (auto simp: t.enum_strict_mono) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1176 | ultimately have "i' < j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1177 | using i by (simp add: enum_strict_mono i'_def) | 
| 60420 | 1178 | with \<open>j \<noteq> i\<close> \<open>j \<le> n\<close> have "t.enum k \<le> t.enum (Suc (Suc l))" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1179 | unfolding i'_def by (simp add: enum_mono k eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1180 | then have "k \<le> Suc (Suc l)" | 
| 60420 | 1181 | using \<open>k \<le> n\<close> \<open>Suc l < k\<close> by (simp add: t.enum_mono) } | 
| 1182 | with \<open>Suc l < k\<close> have "Suc (Suc l) = k" by simp | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1183 | then have "enum (Suc (Suc i')) = t.enum (Suc (Suc l))" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1184 | using i by (simp add: k i'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1185 | also have "\<dots> = (enum i') (u l := Suc (enum i' (u l)), u (Suc l) := Suc (enum i' (u (Suc l))))" | 
| 60420 | 1186 | using \<open>Suc l < k\<close> \<open>k \<le> n\<close> by (simp add: t.enum_Suc l t.upd_inj) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 1187 | finally have "(u l = upd i' \<and> u (Suc l) = upd (Suc i')) \<or> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1188 | (u l = upd (Suc i') \<and> u (Suc l) = upd i')" | 
| 62390 | 1189 | using \<open>Suc i' < n\<close> by (auto simp: enum_Suc fun_eq_iff split: if_split_asm) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1190 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1191 |       then have "t = s \<or> t = b.enum ` {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1192 | proof (elim disjE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1193 | assume u: "u l = upd i'" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1194 | have "c = t.enum (Suc l)" unfolding c_eq .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1195 | also have "t.enum (Suc l) = enum (Suc i')" | 
| 60420 | 1196 | using u \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close> by (simp add: enum_Suc t.enum_Suc l) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1197 | also have "\<dots> = a" | 
| 60420 | 1198 | using \<open>a = enum i\<close> i by (simp add: i'_def) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1199 | finally show ?thesis | 
| 60420 | 1200 | using eq_sma \<open>a \<in> s\<close> \<open>c \<in> t\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1201 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1202 | assume u: "u l = upd (Suc i')" | 
| 63040 | 1203 |         define B where "B = b.enum ` {..n}"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1204 | have "b.enum i' = enum i'" | 
| 68022 | 1205 | using enum_eq_benum[of i'] i by (auto simp: i'_def gr0_conv_Suc) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1206 | have "c = t.enum (Suc l)" unfolding c_eq .. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1207 | also have "t.enum (Suc l) = b.enum (Suc i')" | 
| 60420 | 1208 | using u \<open>l < k\<close> \<open>k \<le> n\<close> \<open>Suc i' < n\<close> | 
| 71172 | 1209 | by (simp_all add: enum_Suc t.enum_Suc l b.enum_Suc \<open>b.enum i' = enum i'\<close>) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1210 | (simp add: Suc_i') | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1211 | also have "\<dots> = b.enum i" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1212 | using i by (simp add: i'_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1213 | finally have "c = b.enum i" . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1214 |         then have "t - {c} = B - {c}" "c \<in> B"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1215 | unfolding eq_sma[symmetric] eq B_def using i by auto | 
| 60420 | 1216 | with \<open>c \<in> t\<close> have "t = B" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1217 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1218 | then show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1219 | by (simp add: B_def) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1220 | qed } | 
| 60420 | 1221 | with ks_f' eq \<open>a \<noteq> b.enum i\<close> \<open>n \<noteq> 0\<close> \<open>i \<le> n\<close> show ?thesis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1222 |       apply (intro ex1I[of _ "b.enum ` {.. n}"])
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1223 | apply auto [] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1224 | apply metis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1225 | done | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1226 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1227 | then show ?thesis | 
| 71449 
3cf130a896a3
lemmas about "card A = 2"; prefer iff to implications
 nipkow parents: 
71172diff
changeset | 1228 | using s \<open>a \<in> s\<close> by (simp add: card_2_iff' Ex1_def) metis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1229 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1230 | |
| 60420 | 1231 | 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 | 1232 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1233 | lemma kuhn_simplex_lemma: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1234 |   assumes "\<forall>s. ksimplex p (Suc n) s \<longrightarrow> rl ` s \<subseteq> {.. Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1235 |     and "odd (card {f. \<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> (f = s - {a}) \<and>
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1236 |       rl ` f = {..n} \<and> ((\<exists>j\<le>n. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>f. x j = p))})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1237 |   shows "odd (card {s. ksimplex p (Suc n) s \<and> rl ` s = {..Suc n}})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1238 | proof (rule kuhn_complete_lemma[OF finite_ksimplexes refl, unfolded mem_Collect_eq, | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1239 |     where bnd="\<lambda>f. (\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{..n}. \<forall>x\<in>f. x j = p)"],
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1240 | safe del: notI) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1241 | |
| 53186 | 1242 | have *: "\<And>x y. x = y \<Longrightarrow> odd (card x) \<Longrightarrow> odd (card y)" | 
| 1243 | by auto | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1244 |   show "odd (card {f. (\<exists>s\<in>{s. ksimplex p (Suc n) s}. \<exists>a\<in>s. f = s - {a}) \<and>
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1245 |     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 | 1246 | apply (rule *[OF _ assms(2)]) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1247 | apply (auto simp: atLeast0AtMost) | 
| 53186 | 1248 | done | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1249 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1250 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1251 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1252 | fix s assume s: "ksimplex p (Suc n) s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1253 | then show "card s = n + 2" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1254 | by (simp add: ksimplex_card) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1255 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1256 | fix a assume a: "a \<in> s" then show "rl a \<le> Suc n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1257 | using assms(1) s by (auto simp: subset_eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1258 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1259 |   let ?S = "{t. ksimplex p (Suc n) t \<and> (\<exists>b\<in>t. s - {a} = t - {b})}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1260 |   { fix j assume j: "j \<le> n" "\<forall>x\<in>s - {a}. x j = 0"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1261 | with s a show "card ?S = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1262 | using ksimplex_replace_0[of p "n + 1" s a j] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1263 | by (subst eq_commute) simp } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1264 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1265 |   { fix j assume j: "j \<le> n" "\<forall>x\<in>s - {a}. x j = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1266 | with s a show "card ?S = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1267 | using ksimplex_replace_1[of p "n + 1" s a j] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1268 | by (subst eq_commute) simp } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1269 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1270 |   { assume "card ?S \<noteq> 2" "\<not> (\<exists>j\<in>{..n}. \<forall>x\<in>s - {a}. x j = p)"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1271 |     with s a show "\<exists>j\<in>{..n}. \<forall>x\<in>s - {a}. x j = 0"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1272 | using ksimplex_replace_2[of p "n + 1" s a] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1273 | by (subst (asm) eq_commute) auto } | 
| 53186 | 1274 | qed | 
| 1275 | ||
| 68617 | 1276 | subsubsection \<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 | 1277 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1278 | 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 | 1279 | |
| 53186 | 1280 | lemma reduced_labelling: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1281 | shows "reduced n x \<le> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1282 | and "\<forall>i<reduced n x. x i = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1283 | and "reduced n x = n \<or> x (reduced n x) \<noteq> 0" | 
| 53186 | 1284 | proof - | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1285 | show "reduced n x \<le> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1286 | unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1287 | show "\<forall>i<reduced n x. x i = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1288 | unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) fastforce+ | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1289 | show "reduced n x = n \<or> x (reduced n x) \<noteq> 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1290 | unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) fastforce+ | 
| 53186 | 1291 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1292 | |
| 53186 | 1293 | lemma reduced_labelling_unique: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1294 | "r \<le> n \<Longrightarrow> \<forall>i<r. x i = 0 \<Longrightarrow> r = n \<or> x r \<noteq> 0 \<Longrightarrow> reduced n x = r" | 
| 76786 | 1295 | by (metis linorder_less_linear linorder_not_le reduced_labelling) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1296 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1297 | lemma reduced_labelling_zero: "j < n \<Longrightarrow> x j = 0 \<Longrightarrow> reduced n x \<noteq> j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1298 | 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 | 1299 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1300 | lemma reduce_labelling_zero[simp]: "reduced 0 x = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1301 | 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 | 1302 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1303 | lemma reduced_labelling_nonzero: "j < n \<Longrightarrow> x j \<noteq> 0 \<Longrightarrow> reduced n x \<le> j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1304 | 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 | 1305 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1306 | lemma reduced_labelling_Suc: "reduced (Suc n) x \<noteq> Suc n \<Longrightarrow> reduced (Suc n) x = reduced n x" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1307 | using reduced_labelling[of "Suc n" x] | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1308 | 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 | 1309 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1310 | lemma complete_face_top: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1311 | assumes "\<forall>x\<in>f. \<forall>j\<le>n. x j = 0 \<longrightarrow> lab x j = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1312 | and "\<forall>x\<in>f. \<forall>j\<le>n. x j = p \<longrightarrow> lab x j = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1313 |     and eq: "(reduced (Suc n) \<circ> lab) ` f = {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1314 | shows "((\<exists>j\<le>n. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>f. x j = p)) \<longleftrightarrow> (\<forall>x\<in>f. x n = p)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1315 | proof (safe del: disjCI) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1316 | fix x j assume j: "j \<le> n" "\<forall>x\<in>f. x j = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1317 |   { fix x assume "x \<in> f" with assms j have "reduced (Suc n) (lab x) \<noteq> j"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1318 | by (intro reduced_labelling_zero) auto } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1319 | moreover have "j \<in> (reduced (Suc n) \<circ> lab) ` f" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1320 | using j eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1321 | ultimately show "x n = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1322 | by force | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1323 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1324 | fix x j assume j: "j \<le> n" "\<forall>x\<in>f. x j = p" and x: "x \<in> f" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1325 | have "j = n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1326 | proof (rule ccontr) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1327 | assume "\<not> ?thesis" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1328 |     { fix x assume "x \<in> f"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1329 | with assms j have "reduced (Suc n) (lab x) \<le> j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1330 | by (intro reduced_labelling_nonzero) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1331 | then have "reduced (Suc n) (lab x) \<noteq> n" | 
| 60420 | 1332 | using \<open>j \<noteq> n\<close> \<open>j \<le> n\<close> by simp } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1333 | moreover | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 1334 | have "n \<in> (reduced (Suc n) \<circ> lab) ` f" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1335 | using eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1336 | ultimately show False | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1337 | by force | 
| 53186 | 1338 | qed | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1339 | moreover have "j \<in> (reduced (Suc n) \<circ> lab) ` f" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1340 | using j eq by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1341 | ultimately show "x n = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1342 | using j x by auto | 
| 53688 | 1343 | 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 | 1344 | |
| 60420 | 1345 | 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 | 1346 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1347 | lemma kuhn_induction: | 
| 53688 | 1348 | assumes "0 < p" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1349 | and lab_0: "\<forall>x. \<forall>j\<le>n. (\<forall>j. x j \<le> p) \<and> x j = 0 \<longrightarrow> lab x j = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1350 | and lab_1: "\<forall>x. \<forall>j\<le>n. (\<forall>j. x j \<le> p) \<and> x j = p \<longrightarrow> lab x j = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1351 |     and odd: "odd (card {s. ksimplex p n s \<and> (reduced n\<circ>lab) ` s = {..n}})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1352 |   shows "odd (card {s. ksimplex p (Suc n) s \<and> (reduced (Suc n)\<circ>lab) ` s = {..Suc n}})"
 | 
| 53248 | 1353 | proof - | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1354 | let ?rl = "reduced (Suc n) \<circ> lab" and ?ext = "\<lambda>f v. \<exists>j\<le>n. \<forall>x\<in>f. x j = v" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1355 | let ?ext = "\<lambda>s. (\<exists>j\<le>n. \<forall>x\<in>s. x j = 0) \<or> (\<exists>j\<le>n. \<forall>x\<in>s. x j = p)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1356 |   have "\<forall>s. ksimplex p (Suc n) s \<longrightarrow> ?rl ` s \<subseteq> {..Suc n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1357 | by (simp add: reduced_labelling subset_eq) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1358 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1359 |   have "{s. ksimplex p n s \<and> (reduced n \<circ> lab) ` s = {..n}} =
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1360 |         {f. \<exists>s a. ksimplex p (Suc n) s \<and> a \<in> s \<and> f = s - {a} \<and> ?rl ` f = {..n} \<and> ?ext f}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1361 | proof (intro set_eqI, safe del: disjCI equalityI disjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1362 |     fix s assume s: "ksimplex p n s" and rl: "(reduced n \<circ> lab) ` s = {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1363 | from s obtain u b where "kuhn_simplex p n u b s" by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1364 | then interpret kuhn_simplex p n u b s . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1365 | have all_eq_p: "\<forall>x\<in>s. x n = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1366 | by (auto simp: out_eq_p) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1367 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1368 |     { fix x assume "x \<in> s"
 | 
| 76786 | 1369 | with lab_1[rule_format, of n x] all_eq_p s_le_p[of x] | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1370 | have "?rl x \<le> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1371 | by (auto intro!: reduced_labelling_nonzero) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1372 | then have "?rl x = reduced n (lab x)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1373 | by (auto intro!: reduced_labelling_Suc) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1374 |     then have "?rl ` s = {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1375 | using rl by (simp cong: image_cong) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1376 | moreover | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1377 |     obtain t a where "ksimplex p (Suc n) t" "a \<in> t" "s = t - {a}"
 | 
| 60420 | 1378 | using s unfolding simplex_top_face[OF \<open>0 < p\<close> all_eq_p] by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1379 | ultimately | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1380 |     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 | 1381 | by auto | 
| 53248 | 1382 | next | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1383 |     fix x s a assume s: "ksimplex p (Suc n) s" and rl: "?rl ` (s - {a}) = {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1384 |       and a: "a \<in> s" and "?ext (s - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1385 | from s obtain u b where "kuhn_simplex p (Suc n) u b s" by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1386 | then interpret kuhn_simplex p "Suc n" u b s . | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1387 | have all_eq_p: "\<forall>x\<in>s. x (Suc n) = p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1388 | by (auto simp: out_eq_p) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1389 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1390 |     { fix x assume "x \<in> s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1391 |       then have "?rl x \<in> ?rl ` (s - {a})"
 | 
| 53248 | 1392 | by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1393 | then have "?rl x \<le> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1394 | unfolding rl by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1395 | then have "?rl x = reduced n (lab x)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1396 | by (auto intro!: reduced_labelling_Suc) } | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1397 |     then show rl': "(reduced n\<circ>lab) ` (s - {a}) = {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1398 | unfolding rl[symmetric] by (intro image_cong) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1399 | |
| 60420 | 1400 |     from \<open>?ext (s - {a})\<close>
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1401 |     have all_eq_p: "\<forall>x\<in>s - {a}. x n = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1402 | proof (elim disjE exE conjE) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1403 |       fix j assume "j \<le> n" "\<forall>x\<in>s - {a}. x j = 0"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1404 | with lab_0[rule_format, of j] all_eq_p s_le_p | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1405 |       have "\<And>x. x \<in> s - {a} \<Longrightarrow> reduced (Suc n) (lab x) \<noteq> j"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1406 | by (intro reduced_labelling_zero) auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1407 |       moreover have "j \<in> ?rl ` (s - {a})"
 | 
| 60420 | 1408 | using \<open>j \<le> n\<close> unfolding rl by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1409 | ultimately show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1410 | by force | 
| 53248 | 1411 | next | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1412 |       fix j assume "j \<le> n" and eq_p: "\<forall>x\<in>s - {a}. x j = p"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1413 | show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1414 | proof cases | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1415 | assume "j = n" with eq_p show ?thesis by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1416 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1417 | assume "j \<noteq> n" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1418 |         { fix x assume x: "x \<in> s - {a}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1419 | have "reduced n (lab x) \<le> j" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1420 | proof (rule reduced_labelling_nonzero) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1421 | show "lab x j \<noteq> 0" | 
| 60420 | 1422 | using lab_1[rule_format, of j x] x s_le_p[of x] eq_p \<open>j \<le> n\<close> by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1423 | show "j < n" | 
| 60420 | 1424 | using \<open>j \<le> n\<close> \<open>j \<noteq> n\<close> by simp | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1425 | qed | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1426 | then have "reduced n (lab x) \<noteq> n" | 
| 60420 | 1427 | using \<open>j \<le> n\<close> \<open>j \<noteq> n\<close> by simp } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1428 |         moreover have "n \<in> (reduced n\<circ>lab) ` (s - {a})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1429 | unfolding rl' by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1430 | ultimately show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1431 | by force | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1432 | qed | 
| 53248 | 1433 | qed | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1434 |     show "ksimplex p n (s - {a})"
 | 
| 60420 | 1435 | unfolding simplex_top_face[OF \<open>0 < p\<close> all_eq_p] using s a by auto | 
| 53248 | 1436 | qed | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1437 | ultimately show ?thesis | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1438 | using assms by (intro kuhn_simplex_lemma) auto | 
| 53248 | 1439 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1440 | |
| 60420 | 1441 | 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 | 1442 | |
| 53688 | 1443 | lemma ksimplex_0: "ksimplex p 0 s \<longleftrightarrow> s = {(\<lambda>x. p)}"
 | 
| 53248 | 1444 | proof | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1445 |   assume "ksimplex p 0 s" then show "s = {(\<lambda>x. p)}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1446 | by (blast dest: kuhn_simplex.ksimplex_0 elim: ksimplex.cases) | 
| 53248 | 1447 | next | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1448 |   assume s: "s = {(\<lambda>x. p)}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1449 | show "ksimplex p 0 s" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1450 | proof (intro ksimplex, unfold_locales) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1451 |     show "(\<lambda>_. p) \<in> {..<0::nat} \<rightarrow> {..<p}" by auto
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1452 |     show "bij_betw id {..<0} {..<0}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1453 | by simp | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1454 | qed (auto simp: s) | 
| 53248 | 1455 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1456 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1457 | lemma kuhn_combinatorial: | 
| 53688 | 1458 | assumes "0 < p" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1459 | and "\<forall>x j. (\<forall>j. x j \<le> p) \<and> j < n \<and> x j = 0 \<longrightarrow> lab x j = 0" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1460 | and "\<forall>x j. (\<forall>j. x j \<le> p) \<and> j < n \<and> x j = p \<longrightarrow> lab x j = 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1461 |   shows "odd (card {s. ksimplex p n s \<and> (reduced n\<circ>lab) ` s = {..n}})"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1462 | (is "odd (card (?M n))") | 
| 53248 | 1463 | using assms | 
| 1464 | proof (induct n) | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1465 | case 0 then show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1466 | by (simp add: ksimplex_0 cong: conj_cong) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1467 | next | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1468 | case (Suc n) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1469 | then have "odd (card (?M n))" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1470 | by force | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1471 | with Suc show ?case | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1472 | using kuhn_induction[of p n] by (auto simp: comp_def) | 
| 53248 | 1473 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1474 | |
| 53248 | 1475 | lemma kuhn_lemma: | 
| 53688 | 1476 | fixes n p :: nat | 
| 1477 | assumes "0 < p" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1478 | and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. label x i = (0::nat) \<or> label x i = 1)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1479 | and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = 0 \<longrightarrow> label x i = 0)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1480 | and "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = p \<longrightarrow> label x i = 1)" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1481 | obtains q where "\<forall>i<n. q i < p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1482 | 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 | 1483 | proof - | 
| 60580 | 1484 | let ?rl = "reduced n \<circ> label" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1485 |   let ?A = "{s. ksimplex p n s \<and> ?rl ` s = {..n}}"
 | 
| 53248 | 1486 | have "odd (card ?A)" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1487 | using assms by (intro kuhn_combinatorial[of p n label]) auto | 
| 53688 | 1488 |   then have "?A \<noteq> {}"
 | 
| 69661 | 1489 | by (rule odd_card_imp_not_empty) | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1490 |   then obtain s b u where "kuhn_simplex p n b u s" and rl: "?rl ` s = {..n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1491 | by (auto elim: ksimplex.cases) | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1492 | interpret kuhn_simplex p n b u s by fact | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1493 | |
| 53248 | 1494 | show ?thesis | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1495 | proof (intro that[of b] allI impI) | 
| 60580 | 1496 | fix i | 
| 1497 | assume "i < n" | |
| 1498 | then show "b i < p" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1499 | using base by auto | 
| 53248 | 1500 | next | 
| 60580 | 1501 | fix i | 
| 1502 | assume "i < n" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1503 |     then have "i \<in> {.. n}" "Suc i \<in> {.. n}"
 | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1504 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1505 | then obtain u v where u: "u \<in> s" "Suc i = ?rl u" and v: "v \<in> s" "i = ?rl v" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1506 | unfolding rl[symmetric] by blast | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1507 | |
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1508 | have "label u i \<noteq> label v i" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1509 | using reduced_labelling [of n "label u"] reduced_labelling [of n "label v"] | 
| 60420 | 1510 | u(2)[symmetric] v(2)[symmetric] \<open>i < n\<close> | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1511 | by auto | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1512 | moreover | 
| 60580 | 1513 | 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 | 
| 1514 | 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>] | |
| 1515 | by auto | |
| 1516 | ultimately show "\<exists>r s. (\<forall>j<n. b j \<le> r j \<and> r j \<le> b j + 1) \<and> | |
| 1517 | (\<forall>j<n. b j \<le> s j \<and> s j \<le> b j + 1) \<and> label r i \<noteq> label s i" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1518 | by blast | 
| 53248 | 1519 | qed | 
| 1520 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1521 | |
| 68617 | 1522 | subsubsection \<open>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 | 1523 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1524 | lemma kuhn_labelling_lemma': | 
| 53688 | 1525 | assumes "(\<forall>x::nat\<Rightarrow>real. P x \<longrightarrow> P (f x))" | 
| 1526 | 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 | 1527 | shows "\<exists>l. (\<forall>x i. l x i \<le> (1::nat)) \<and> | 
| 53688 | 1528 | (\<forall>x i. P x \<and> Q i \<and> x i = 0 \<longrightarrow> l x i = 0) \<and> | 
| 1529 | (\<forall>x i. P x \<and> Q i \<and> x i = 1 \<longrightarrow> l x i = 1) \<and> | |
| 1530 | (\<forall>x i. P x \<and> Q i \<and> l x i = 0 \<longrightarrow> x i \<le> f x i) \<and> | |
| 1531 | (\<forall>x i. P x \<and> Q i \<and> l x i = 1 \<longrightarrow> f x i \<le> x i)" | |
| 76786 | 1532 | unfolding all_conj_distrib [symmetric] | 
| 1533 | apply (subst choice_iff[symmetric])+ | |
| 80095 | 1534 | by (metis assms choice_iff bot_nat_0.extremum nle_le zero_neq_one) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1535 | |
| 68617 | 1536 | subsection \<open>Brouwer's fixed point theorem\<close> | 
| 1537 | ||
| 68621 | 1538 | text \<open>We start proving Brouwer's fixed point theorem for the unit cube = \<open>cbox 0 One\<close>.\<close> | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1539 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1540 | lemma brouwer_cube: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1541 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | 
| 68621 | 1542 | assumes "continuous_on (cbox 0 One) f" | 
| 1543 | and "f ` cbox 0 One \<subseteq> cbox 0 One" | |
| 1544 | shows "\<exists>x\<in>cbox 0 One. f x = x" | |
| 53185 | 1545 | proof (rule ccontr) | 
| 63040 | 1546 |   define n where "n = DIM('a)"
 | 
| 53185 | 1547 | have n: "1 \<le> n" "0 < n" "n \<noteq> 0" | 
| 71172 | 1548 | unfolding n_def by (auto simp: Suc_le_eq) | 
| 53674 | 1549 | assume "\<not> ?thesis" | 
| 68621 | 1550 | then have *: "\<not> (\<exists>x\<in>cbox 0 One. f x - x = 0)" | 
| 53674 | 1551 | by auto | 
| 55522 | 1552 | obtain d where | 
| 68621 | 1553 | d: "d > 0" "\<And>x. x \<in> cbox 0 One \<Longrightarrow> d \<le> norm (f x - x)" | 
| 80095 | 1554 | using brouwer_compactness_lemma[OF compact_cbox _ *] assms | 
| 1555 | by (metis (no_types, lifting) continuous_on_cong continuous_on_diff continuous_on_id) | |
| 68621 | 1556 | have *: "\<forall>x. x \<in> cbox 0 One \<longrightarrow> f x \<in> cbox 0 One" | 
| 1557 | "\<forall>x. x \<in> (cbox 0 One::'a set) \<longrightarrow> (\<forall>i\<in>Basis. True \<longrightarrow> 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)" | |
| 53185 | 1558 | using assms(2)[unfolded image_subset_iff Ball_def] | 
| 68621 | 1559 | unfolding cbox_def | 
| 55522 | 1560 | by auto | 
| 68022 | 1561 | obtain label :: "'a \<Rightarrow> 'a \<Rightarrow> nat" where label [rule_format]: | 
| 55522 | 1562 | "\<forall>x. \<forall>i\<in>Basis. label x i \<le> 1" | 
| 68621 | 1563 | "\<forall>x. \<forall>i\<in>Basis. x \<in> cbox 0 One \<and> x \<bullet> i = 0 \<longrightarrow> label x i = 0" | 
| 1564 | "\<forall>x. \<forall>i\<in>Basis. x \<in> cbox 0 One \<and> x \<bullet> i = 1 \<longrightarrow> label x i = 1" | |
| 1565 | "\<forall>x. \<forall>i\<in>Basis. x \<in> cbox 0 One \<and> label x i = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i" | |
| 1566 | "\<forall>x. \<forall>i\<in>Basis. x \<in> cbox 0 One \<and> label x i = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i" | |
| 68022 | 1567 | using kuhn_labelling_lemma[OF *] by auto | 
| 53185 | 1568 | note label = this [rule_format] | 
| 68621 | 1569 | have lem1: "\<forall>x\<in>cbox 0 One. \<forall>y\<in>cbox 0 One. \<forall>i\<in>Basis. label x i \<noteq> label y i \<longrightarrow> | 
| 61945 | 1570 | \<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)" | 
| 53185 | 1571 | proof safe | 
| 1572 | fix x y :: 'a | |
| 68621 | 1573 | assume x: "x \<in> cbox 0 One" and y: "y \<in> cbox 0 One" | 
| 53185 | 1574 | fix i | 
| 1575 | assume i: "label x i \<noteq> label y i" "i \<in> Basis" | |
| 1576 | have *: "\<And>x y fx fy :: real. x \<le> fx \<and> fy \<le> y \<or> fx \<le> x \<and> y \<le> fy \<Longrightarrow> | |
| 61945 | 1577 | \<bar>fx - x\<bar> \<le> \<bar>fy - fx\<bar> + \<bar>y - x\<bar>" by auto | 
| 1578 | have "\<bar>(f x - x) \<bullet> i\<bar> \<le> \<bar>(f y - f x)\<bullet>i\<bar> + \<bar>(y - x)\<bullet>i\<bar>" | |
| 68022 | 1579 | proof (cases "label x i = 0") | 
| 1580 | case True | |
| 1581 | then have fxy: "\<not> f y \<bullet> i \<le> y \<bullet> i \<Longrightarrow> f x \<bullet> i \<le> x \<bullet> i" | |
| 1582 | by (metis True i label(1) label(5) le_antisym less_one not_le_imp_less y) | |
| 1583 | show ?thesis | |
| 1584 | unfolding inner_simps | |
| 1585 | by (rule *) (auto simp: True i label x y fxy) | |
| 53185 | 1586 | next | 
| 68022 | 1587 | case False | 
| 1588 | then show ?thesis | |
| 1589 | using label [OF \<open>i \<in> Basis\<close>] i(1) x y | |
| 80095 | 1590 | by (smt (verit, ccfv_threshold) inner_diff_left less_one order_le_less) | 
| 53185 | 1591 | qed | 
| 1592 | also have "\<dots> \<le> norm (f y - f x) + norm (y - x)" | |
| 68022 | 1593 | by (simp add: add_mono i(2) norm_bound_Basis_le) | 
| 53185 | 1594 | finally show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)" | 
| 1595 | unfolding inner_simps . | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1596 | qed | 
| 68621 | 1597 | have "\<exists>e>0. \<forall>x\<in>cbox 0 One. \<forall>y\<in>cbox 0 One. \<forall>z\<in>cbox 0 One. \<forall>i\<in>Basis. | 
| 68022 | 1598 | norm (x - z) < e \<longrightarrow> norm (y - z) < e \<longrightarrow> label x i \<noteq> label y i \<longrightarrow> | 
| 61945 | 1599 | \<bar>(f(z) - z)\<bullet>i\<bar> < d / (real n)" | 
| 53185 | 1600 | proof - | 
| 53688 | 1601 | have d': "d / real n / 8 > 0" | 
| 71172 | 1602 | using d(1) by (simp add: n_def) | 
| 68621 | 1603 | have *: "uniformly_continuous_on (cbox 0 One) f" | 
| 1604 | by (rule compact_uniformly_continuous[OF assms(1) compact_cbox]) | |
| 55522 | 1605 | obtain e where e: | 
| 1606 | "e > 0" | |
| 68621 | 1607 | "\<And>x x'. x \<in> cbox 0 One \<Longrightarrow> | 
| 1608 | x' \<in> cbox 0 One \<Longrightarrow> | |
| 55522 | 1609 | norm (x' - x) < e \<Longrightarrow> | 
| 1610 | norm (f x' - f x) < d / real n / 8" | |
| 1611 | using *[unfolded uniformly_continuous_on_def,rule_format,OF d'] | |
| 1612 | unfolding dist_norm | |
| 1613 | by blast | |
| 53185 | 1614 | show ?thesis | 
| 68022 | 1615 | proof (intro exI conjI ballI impI) | 
| 53185 | 1616 | show "0 < min (e / 2) (d / real n / 8)" | 
| 1617 | using d' e by auto | |
| 1618 | fix x y z i | |
| 53688 | 1619 | assume as: | 
| 68621 | 1620 | "x \<in> cbox 0 One" "y \<in> cbox 0 One" "z \<in> cbox 0 One" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36587diff
changeset | 1621 | "norm (x - z) < min (e / 2) (d / real n / 8)" | 
| 53688 | 1622 | "norm (y - z) < min (e / 2) (d / real n / 8)" | 
| 1623 | "label x i \<noteq> label y i" | |
| 1624 | assume i: "i \<in> Basis" | |
| 61945 | 1625 | have *: "\<And>z fz x fx n1 n2 n3 n4 d4 d :: real. \<bar>fx - x\<bar> \<le> n1 + n2 \<Longrightarrow> | 
| 1626 | \<bar>fx - fz\<bar> \<le> n3 \<Longrightarrow> \<bar>x - z\<bar> \<le> n4 \<Longrightarrow> | |
| 53185 | 1627 | n1 < d4 \<Longrightarrow> n2 < 2 * d4 \<Longrightarrow> n3 < d4 \<Longrightarrow> n4 < d4 \<Longrightarrow> | 
| 61945 | 1628 | (8 * d4 = d) \<Longrightarrow> \<bar>fz - z\<bar> < d" | 
| 53688 | 1629 | by auto | 
| 1630 | show "\<bar>(f z - z) \<bullet> i\<bar> < d / real n" | |
| 1631 | unfolding inner_simps | |
| 53185 | 1632 | proof (rule *) | 
| 1633 | show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y -f x) + norm (y - x)" | |
| 68022 | 1634 | using as(1) as(2) as(6) i lem1 by blast | 
| 1635 | show "norm (f x - f z) < d / real n / 8" | |
| 1636 | using d' e as by auto | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1637 | 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 | 1638 | unfolding inner_diff_left[symmetric] | 
| 53688 | 1639 | by (rule Basis_le_norm[OF i])+ | 
| 1640 | have tria: "norm (y - x) \<le> norm (y - z) + norm (x - z)" | |
| 53185 | 1641 | using dist_triangle[of y x z, unfolded dist_norm] | 
| 53688 | 1642 | unfolding norm_minus_commute | 
| 1643 | by auto | |
| 53185 | 1644 | also have "\<dots> < e / 2 + e / 2" | 
| 68022 | 1645 | using as(4) as(5) by auto | 
| 53185 | 1646 | finally show "norm (f y - f x) < d / real n / 8" | 
| 68022 | 1647 | using as(1) as(2) e(2) by auto | 
| 53185 | 1648 | have "norm (y - z) + norm (x - z) < d / real n / 8 + d / real n / 8" | 
| 68022 | 1649 | using as(4) as(5) by auto | 
| 1650 | with tria show "norm (y - x) < 2 * (d / real n / 8)" | |
| 53688 | 1651 | by auto | 
| 68022 | 1652 | qed (use as in auto) | 
| 53185 | 1653 | qed | 
| 1654 | qed | |
| 55522 | 1655 | then | 
| 1656 | obtain e where e: | |
| 1657 | "e > 0" | |
| 68621 | 1658 | "\<And>x y z i. x \<in> cbox 0 One \<Longrightarrow> | 
| 1659 | y \<in> cbox 0 One \<Longrightarrow> | |
| 1660 | z \<in> cbox 0 One \<Longrightarrow> | |
| 55522 | 1661 | i \<in> Basis \<Longrightarrow> | 
| 1662 | norm (x - z) < e \<and> norm (y - z) < e \<and> label x i \<noteq> label y i \<Longrightarrow> | |
| 1663 | \<bar>(f z - z) \<bullet> i\<bar> < d / real n" | |
| 1664 | by blast | |
| 1665 | obtain p :: nat where p: "1 + real n / e \<le> real p" | |
| 1666 | using real_arch_simple .. | |
| 53185 | 1667 | have "1 + real n / e > 0" | 
| 56541 | 1668 | using e(1) n by (simp add: add_pos_pos) | 
| 53688 | 1669 | then have "p > 0" | 
| 1670 | using p by auto | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1671 | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1672 |   obtain b :: "nat \<Rightarrow> 'a" where b: "bij_betw b {..< n} Basis"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1673 | by atomize_elim (auto simp: n_def intro!: finite_same_card_bij) | 
| 63040 | 1674 |   define b' where "b' = inv_into {..< n} b"
 | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1675 |   then have b': "bij_betw b' Basis {..< n}"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1676 | using bij_betw_inv_into[OF b] by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1677 |   then have b'_Basis: "\<And>i. i \<in> Basis \<Longrightarrow> b' i \<in> {..< n}"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1678 | unfolding bij_betw_def by (auto simp: set_eq_iff) | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1679 | have bb'[simp]:"\<And>i. i \<in> Basis \<Longrightarrow> b (b' i) = i" | 
| 53688 | 1680 | unfolding b'_def | 
| 1681 | using b | |
| 1682 | by (auto simp: f_inv_into_f bij_betw_def) | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1683 | have b'b[simp]:"\<And>i. i < n \<Longrightarrow> b' (b i) = i" | 
| 53688 | 1684 | unfolding b'_def | 
| 1685 | using b | |
| 1686 | by (auto simp: inv_into_f_eq bij_betw_def) | |
| 1687 | have *: "\<And>x :: nat. x = 0 \<or> x = 1 \<longleftrightarrow> x \<le> 1" | |
| 1688 | by auto | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1689 | have b'': "\<And>j. j < n \<Longrightarrow> b j \<in> Basis" | 
| 53185 | 1690 | using b unfolding bij_betw_def by auto | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1691 | have q1: "0 < p" "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1692 | (\<forall>i<n. (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0 \<or> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1693 | (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)" | 
| 53688 | 1694 | unfolding * | 
| 60420 | 1695 | using \<open>p > 0\<close> \<open>n > 0\<close> | 
| 53688 | 1696 | using label(1)[OF b''] | 
| 1697 | by auto | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1698 |   { fix x :: "nat \<Rightarrow> nat" and i assume "\<forall>i<n. x i \<le> p" "i < n" "x i = p \<or> x i = 0"
 | 
| 68621 | 1699 | then have "(\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<in> (cbox 0 One::'a set)" | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1700 | using b'_Basis | 
| 71172 | 1701 | by (auto simp: cbox_def bij_betw_def zero_le_divide_iff divide_le_eq_1) } | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1702 | note cube = this | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1703 | have q2: "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = 0 \<longrightarrow> | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1704 | (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 0)" | 
| 68022 | 1705 | unfolding o_def using cube \<open>p > 0\<close> by (intro allI impI label(2)) (auto simp: b'') | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61520diff
changeset | 1706 | have q3: "\<forall>x. (\<forall>i<n. x i \<le> p) \<longrightarrow> (\<forall>i<n. x i = p \<longrightarrow> | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1707 | (label (\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<circ> b) i = 1)" | 
| 68022 | 1708 | using cube \<open>p > 0\<close> unfolding o_def by (intro allI impI label(3)) (auto simp: b'') | 
| 55522 | 1709 | obtain q where q: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1710 | "\<forall>i<n. q i < p" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1711 | "\<forall>i<n. | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1712 | \<exists>r s. (\<forall>j<n. q j \<le> r j \<and> r j \<le> q j + 1) \<and> | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1713 | (\<forall>j<n. q j \<le> s j \<and> s j \<le> q j + 1) \<and> | 
| 55522 | 1714 | (label (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i) \<circ> b) i \<noteq> | 
| 1715 | (label (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i) \<circ> b) i" | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1716 | by (rule kuhn_lemma[OF q1 q2 q3]) | 
| 63040 | 1717 | define z :: 'a where "z = (\<Sum>i\<in>Basis. (real (q (b' i)) / real p) *\<^sub>R i)" | 
| 61945 | 1718 | have "\<exists>i\<in>Basis. d / real n \<le> \<bar>(f z - z)\<bullet>i\<bar>" | 
| 53185 | 1719 | proof (rule ccontr) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1720 |     have "\<forall>i\<in>Basis. q (b' i) \<in> {0..p}"
 | 
| 53688 | 1721 | using q(1) b' | 
| 1722 | by (auto intro: less_imp_le simp: bij_betw_def) | |
| 68621 | 1723 | then have "z \<in> cbox 0 One" | 
| 1724 | unfolding z_def cbox_def | |
| 53688 | 1725 | using b'_Basis | 
| 68022 | 1726 | by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) | 
| 53688 | 1727 | then have d_fz_z: "d \<le> norm (f z - z)" | 
| 1728 | by (rule d) | |
| 1729 | assume "\<not> ?thesis" | |
| 53674 | 1730 | then have as: "\<forall>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar> < d / real n" | 
| 60420 | 1731 | using \<open>n > 0\<close> | 
| 68022 | 1732 | by (auto simp: not_le inner_diff) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1733 | have "norm (f z - z) \<le> (\<Sum>i\<in>Basis. \<bar>f z \<bullet> i - z \<bullet> i\<bar>)" | 
| 53688 | 1734 | unfolding inner_diff_left[symmetric] | 
| 1735 | by (rule norm_le_l1) | |
| 53185 | 1736 | also have "\<dots> < (\<Sum>(i::'a) \<in> Basis. d / real n)" | 
| 68022 | 1737 | by (meson as finite_Basis nonempty_Basis sum_strict_mono) | 
| 53185 | 1738 | also have "\<dots> = d" | 
| 68022 | 1739 | using DIM_positive[where 'a='a] by (auto simp: n_def) | 
| 53688 | 1740 | finally show False | 
| 1741 | using d_fz_z by auto | |
| 53185 | 1742 | qed | 
| 55522 | 1743 | then obtain i where i: "i \<in> Basis" "d / real n \<le> \<bar>(f z - z) \<bullet> i\<bar>" .. | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1744 | have *: "b' i < n" | 
| 55522 | 1745 | using i and b'[unfolded bij_betw_def] | 
| 53688 | 1746 | by auto | 
| 55522 | 1747 | obtain r s where rs: | 
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1748 | "\<And>j. j < n \<Longrightarrow> q j \<le> r j \<and> r j \<le> q j + 1" | 
| 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1749 | "\<And>j. j < n \<Longrightarrow> q j \<le> s j \<and> s j \<le> q j + 1" | 
| 55522 | 1750 | "(label (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i) \<circ> b) (b' i) \<noteq> | 
| 1751 | (label (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i) \<circ> b) (b' i)" | |
| 1752 | using q(2)[rule_format,OF *] by blast | |
| 56273 
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
 hoelzl parents: 
56226diff
changeset | 1753 | have b'_im: "\<And>i. i \<in> Basis \<Longrightarrow> b' i < n" | 
| 53185 | 1754 | using b' unfolding bij_betw_def by auto | 
| 63040 | 1755 | define r' ::'a where "r' = (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i)" | 
| 53185 | 1756 | have "\<And>i. i \<in> Basis \<Longrightarrow> r (b' i) \<le> p" | 
| 76786 | 1757 | using b'_im q(1) rs(1) by fastforce | 
| 68621 | 1758 | then have "r' \<in> cbox 0 One" | 
| 1759 | unfolding r'_def cbox_def | |
| 53688 | 1760 | using b'_Basis | 
| 68022 | 1761 | by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) | 
| 63040 | 1762 | define s' :: 'a where "s' = (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i)" | 
| 53688 | 1763 | have "\<And>i. i \<in> Basis \<Longrightarrow> s (b' i) \<le> p" | 
| 68022 | 1764 | using b'_im q(1) rs(2) by fastforce | 
| 68621 | 1765 | then have "s' \<in> cbox 0 One" | 
| 1766 | unfolding s'_def cbox_def | |
| 68022 | 1767 | using b'_Basis by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) | 
| 68621 | 1768 | have "z \<in> cbox 0 One" | 
| 1769 | unfolding z_def cbox_def | |
| 60420 | 1770 | using b'_Basis q(1)[rule_format,OF b'_im] \<open>p > 0\<close> | 
| 68022 | 1771 | by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le) | 
| 53688 | 1772 |   {
 | 
| 1773 | have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)" | |
| 68022 | 1774 | by (rule sum_mono) (use rs(1)[OF b'_im] in force) | 
| 53688 | 1775 | also have "\<dots> < e * real p" | 
| 60420 | 1776 | using p \<open>e > 0\<close> \<open>p > 0\<close> | 
| 68022 | 1777 | by (auto simp: field_simps n_def) | 
| 53185 | 1778 | finally have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) < e * real p" . | 
| 1779 | } | |
| 1780 | moreover | |
| 53688 | 1781 |   {
 | 
| 1782 | have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)" | |
| 68022 | 1783 | by (rule sum_mono) (use rs(2)[OF b'_im] in force) | 
| 53688 | 1784 | also have "\<dots> < e * real p" | 
| 60420 | 1785 | using p \<open>e > 0\<close> \<open>p > 0\<close> | 
| 68022 | 1786 | by (auto simp: field_simps n_def) | 
| 53185 | 1787 | finally have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) < e * real p" . | 
| 1788 | } | |
| 1789 | ultimately | |
| 53688 | 1790 | have "norm (r' - z) < e" and "norm (s' - z) < e" | 
| 53185 | 1791 | unfolding r'_def s'_def z_def | 
| 60420 | 1792 | using \<open>p > 0\<close> | 
| 53185 | 1793 | apply (rule_tac[!] le_less_trans[OF norm_le_l1]) | 
| 68022 | 1794 | apply (auto simp: field_simps sum_divide_distrib[symmetric] inner_diff_left) | 
| 53185 | 1795 | done | 
| 53674 | 1796 | then have "\<bar>(f z - z) \<bullet> i\<bar> < d / real n" | 
| 53688 | 1797 | using rs(3) i | 
| 1798 | unfolding r'_def[symmetric] s'_def[symmetric] o_def bb' | |
| 68621 | 1799 | by (intro e(2)[OF \<open>r'\<in>cbox 0 One\<close> \<open>s'\<in>cbox 0 One\<close> \<open>z\<in>cbox 0 One\<close>]) auto | 
| 53688 | 1800 | then show False | 
| 1801 | using i by auto | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1802 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1803 | |
| 68617 | 1804 | text \<open>Next step is to prove it for nonempty interiors.\<close> | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1805 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50514diff
changeset | 1806 | lemma brouwer_weak: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1807 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | 
| 68022 | 1808 | assumes "compact S" | 
| 1809 | and "convex S" | |
| 1810 |     and "interior S \<noteq> {}"
 | |
| 1811 | and "continuous_on S f" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 1812 | and "f \<in> S \<rightarrow> S" | 
| 68022 | 1813 | obtains x where "x \<in> S" and "f x = x" | 
| 53185 | 1814 | proof - | 
| 68621 | 1815 | let ?U = "cbox 0 One :: 'a set" | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1816 | have "\<Sum>Basis /\<^sub>R 2 \<in> interior ?U" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1817 | proof (rule interiorI) | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1818 |     let ?I = "(\<Inter>i\<in>Basis. {x::'a. 0 < x \<bullet> i} \<inter> {x. x \<bullet> i < 1})"
 | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1819 | show "open ?I" | 
| 71172 | 1820 | by (intro open_INT finite_Basis ballI open_Int, auto intro: open_Collect_less simp: continuous_on_inner) | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1821 | show "\<Sum>Basis /\<^sub>R 2 \<in> ?I" | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1822 | by simp | 
| 68621 | 1823 | show "?I \<subseteq> cbox 0 One" | 
| 1824 | unfolding cbox_def by force | |
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1825 | qed | 
| 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1826 |   then have *: "interior ?U \<noteq> {}" by fast
 | 
| 68022 | 1827 | have *: "?U homeomorphic S" | 
| 68621 | 1828 | using homeomorphic_convex_compact[OF convex_box(1) compact_cbox * assms(2,1,3)] . | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 1829 | have "\<forall>f. continuous_on ?U f \<and> f \<in> ?U \<rightarrow> ?U \<longrightarrow> (\<exists>x\<in>?U. f x = x)" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36587diff
changeset | 1830 | using brouwer_cube by auto | 
| 53674 | 1831 | then show ?thesis | 
| 53185 | 1832 | unfolding homeomorphic_fixpoint_property[OF *] | 
| 53252 | 1833 | using assms | 
| 68022 | 1834 | by (auto intro: that) | 
| 53185 | 1835 | qed | 
| 1836 | ||
| 68617 | 1837 | text \<open>Then the particular case for closed balls.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1838 | |
| 53185 | 1839 | lemma brouwer_ball: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1840 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | 
| 53674 | 1841 | assumes "e > 0" | 
| 1842 | and "continuous_on (cball a e) f" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 1843 | and "f \<in> cball a e \<rightarrow> cball a e" | 
| 53674 | 1844 | obtains x where "x \<in> cball a e" and "f x = x" | 
| 53185 | 1845 | using brouwer_weak[OF compact_cball convex_cball, of a e f] | 
| 1846 | 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 | 1847 | 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 | 1848 | |
| 68617 | 1849 | text \<open>And finally we prove Brouwer's fixed point theorem in its general version.\<close> | 
| 1850 | ||
| 1851 | theorem brouwer: | |
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1852 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | 
| 68022 | 1853 |   assumes S: "compact S" "convex S" "S \<noteq> {}"
 | 
| 1854 | and contf: "continuous_on S f" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 1855 | and fim: "f \<in> S \<rightarrow> S" | 
| 68022 | 1856 | obtains x where "x \<in> S" and "f x = x" | 
| 53185 | 1857 | proof - | 
| 68022 | 1858 | have "\<exists>e>0. S \<subseteq> cball 0 e" | 
| 1859 | using compact_imp_bounded[OF \<open>compact S\<close>] unfolding bounded_pos | |
| 1860 | by auto | |
| 1861 | then obtain e where e: "e > 0" "S \<subseteq> cball 0 e" | |
| 55522 | 1862 | by blast | 
| 68022 | 1863 | have "\<exists>x\<in> cball 0 e. (f \<circ> closest_point S) x = x" | 
| 1864 | proof (rule_tac brouwer_ball[OF e(1)]) | |
| 1865 | show "continuous_on (cball 0 e) (f \<circ> closest_point S)" | |
| 76786 | 1866 | by (meson assms closest_point_in_set compact_eq_bounded_closed contf continuous_on_closest_point | 
| 1867 | continuous_on_compose continuous_on_subset image_subsetI) | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 1868 | show "f \<circ> closest_point S \<in> cball 0 e \<rightarrow> cball 0 e" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 1869 | by (smt (verit) Pi_iff assms(1) assms(3) closest_point_in_set comp_apply compact_eq_bounded_closed e(2) fim subset_eq) | 
| 68022 | 1870 | qed (use assms in auto) | 
| 1871 | then obtain x where x: "x \<in> cball 0 e" "(f \<circ> closest_point S) x = x" .. | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 1872 | with S have "x \<in> S" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 1873 | by (metis PiE closest_point_in_set comp_apply compact_imp_closed fim) | 
| 68022 | 1874 | then have *: "closest_point S x = x" | 
| 1875 | by (rule closest_point_self) | |
| 53185 | 1876 | show thesis | 
| 68022 | 1877 | proof | 
| 1878 | show "closest_point S x \<in> S" | |
| 1879 | by (simp add: "*" \<open>x \<in> S\<close>) | |
| 1880 | show "f (closest_point S x) = closest_point S x" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 1881 | using "*" x by auto | 
| 68022 | 1882 | qed | 
| 53185 | 1883 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1884 | |
| 68617 | 1885 | subsection \<open>Applications\<close> | 
| 1886 | ||
| 60420 | 1887 | 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 | 1888 | |
| 68617 | 1889 | corollary no_retraction_cball: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55522diff
changeset | 1890 | fixes a :: "'a::euclidean_space" | 
| 53674 | 1891 | assumes "e > 0" | 
| 1892 | shows "\<not> (frontier (cball a e) retract_of (cball a e))" | |
| 53185 | 1893 | proof | 
| 60580 | 1894 | assume *: "frontier (cball a e) retract_of (cball a e)" | 
| 1895 | have **: "\<And>xa. a - (2 *\<^sub>R a - xa) = - (a - xa)" | |
| 53185 | 1896 | using scaleR_left_distrib[of 1 1 a] by auto | 
| 68022 | 1897 |   obtain x where x: "x \<in> {x. norm (a - x) = e}" "2 *\<^sub>R a - x = x"
 | 
| 1898 | proof (rule retract_fixpoint_property[OF *, of "\<lambda>x. scaleR 2 a - x"]) | |
| 1899 | show "continuous_on (frontier (cball a e)) ((-) (2 *\<^sub>R a))" | |
| 1900 | by (intro continuous_intros) | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 1901 | show "(-) (2 *\<^sub>R a) \<in> frontier (cball a e) \<rightarrow> frontier (cball a e)" | 
| 68022 | 1902 | by clarsimp (metis "**" dist_norm norm_minus_cancel) | 
| 1903 | qed (auto simp: dist_norm intro: brouwer_ball[OF assms]) | |
| 53674 | 1904 | then have "scaleR 2 a = scaleR 1 x + scaleR 1 x" | 
| 68022 | 1905 | by (auto simp: algebra_simps) | 
| 53674 | 1906 | then have "a = x" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 1907 | unfolding scaleR_left_distrib[symmetric] by auto | 
| 53674 | 1908 | then show False | 
| 1909 | using x assms by auto | |
| 53185 | 1910 | qed | 
| 1911 | ||
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1912 | corollary contractible_sphere: | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1913 | fixes a :: "'a::euclidean_space" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1914 | shows "contractible(sphere a r) \<longleftrightarrow> r \<le> 0" | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1915 | proof (cases "0 < r") | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1916 | case True | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1917 | then show ?thesis | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1918 | unfolding contractible_def nullhomotopic_from_sphere_extension | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1919 | using no_retraction_cball [OF True, of a] | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1920 | by (auto simp: retract_of_def retraction_def) | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1921 | next | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1922 | case False | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1923 | then show ?thesis | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1924 | unfolding contractible_def nullhomotopic_from_sphere_extension | 
| 71172 | 1925 | using less_eq_real_def by auto | 
| 64006 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1926 | qed | 
| 
0de4736dad8b
new theorems including the theory FurtherTopology
 paulson <lp15@cam.ac.uk> parents: 
63928diff
changeset | 1927 | |
| 68617 | 1928 | corollary connected_sphere_eq: | 
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1929 | fixes a :: "'a :: euclidean_space" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1930 |   shows "connected(sphere a r) \<longleftrightarrow> 2 \<le> DIM('a) \<or> r \<le> 0"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1931 | (is "?lhs = ?rhs") | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1932 | proof (cases r "0::real" rule: linorder_cases) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1933 | case less | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1934 | then show ?thesis by auto | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1935 | next | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1936 | case equal | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1937 | then show ?thesis by auto | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1938 | next | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1939 | case greater | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1940 | show ?thesis | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1941 | proof | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1942 | assume L: ?lhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1943 |     have "False" if 1: "DIM('a) = 1"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1944 | proof - | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1945 |       obtain x y where xy: "sphere a r = {x,y}" "x \<noteq> y"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1946 | using sphere_1D_doubleton [OF 1 greater] | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1947 | by (metis dist_self greater insertI1 less_add_same_cancel1 mem_sphere mult_2 not_le zero_le_dist) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1948 | then have "finite (sphere a r)" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1949 | by auto | 
| 68022 | 1950 | with L \<open>r > 0\<close> xy show "False" | 
| 1951 | using connected_finite_iff_sing by auto | |
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1952 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1953 | with greater show ?rhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1954 | by (metis DIM_ge_Suc0 One_nat_def Suc_1 le_antisym not_less_eq_eq) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1955 | next | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1956 | assume ?rhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1957 | then show ?lhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1958 | using connected_sphere greater by auto | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1959 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1960 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1961 | |
| 68617 | 1962 | corollary path_connected_sphere_eq: | 
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1963 | fixes a :: "'a :: euclidean_space" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1964 |   shows "path_connected(sphere a r) \<longleftrightarrow> 2 \<le> DIM('a) \<or> r \<le> 0"
 | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1965 | (is "?lhs = ?rhs") | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1966 | proof | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1967 | assume ?lhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1968 | then show ?rhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1969 | using connected_sphere_eq path_connected_imp_connected by blast | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1970 | next | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1971 | assume R: ?rhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1972 | then show ?lhs | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1973 | by (auto simp: contractible_imp_path_connected contractible_sphere path_connected_sphere) | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1974 | qed | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 1975 | |
| 64122 | 1976 | proposition frontier_subset_retraction: | 
| 1977 | fixes S :: "'a::euclidean_space set" | |
| 1978 | assumes "bounded S" and fros: "frontier S \<subseteq> T" | |
| 1979 | and contf: "continuous_on (closure S) f" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 1980 | and fim: "f \<in> S \<rightarrow> T" | 
| 64122 | 1981 | and fid: "\<And>x. x \<in> T \<Longrightarrow> f x = x" | 
| 1982 | shows "S \<subseteq> T" | |
| 1983 | proof (rule ccontr) | |
| 1984 | assume "\<not> S \<subseteq> T" | |
| 1985 | then obtain a where "a \<in> S" "a \<notin> T" by blast | |
| 1986 | define g where "g \<equiv> \<lambda>z. if z \<in> closure S then f z else z" | |
| 1987 | have "continuous_on (closure S \<union> closure(-S)) g" | |
| 80095 | 1988 | unfolding g_def using fros fid frontier_closures | 
| 1989 | by (intro continuous_on_cases) (auto simp: contf) | |
| 64122 | 1990 | moreover have "closure S \<union> closure(- S) = UNIV" | 
| 1991 | using closure_Un by fastforce | |
| 1992 | ultimately have contg: "continuous_on UNIV g" by metis | |
| 1993 | obtain B where "0 < B" and B: "closure S \<subseteq> ball a B" | |
| 1994 | using \<open>bounded S\<close> bounded_subset_ballD by blast | |
| 1995 | have notga: "g x \<noteq> a" for x | |
| 1996 | unfolding g_def using fros fim \<open>a \<notin> T\<close> | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 1997 | by (metis PiE Un_iff \<open>a \<in> S\<close> closure_Un_frontier fid subsetD) | 
| 64122 | 1998 | define h where "h \<equiv> (\<lambda>y. a + (B / norm(y - a)) *\<^sub>R (y - a)) \<circ> g" | 
| 1999 | have "\<not> (frontier (cball a B) retract_of (cball a B))" | |
| 2000 | by (metis no_retraction_cball \<open>0 < B\<close>) | |
| 2001 | then have "\<And>k. \<not> retraction (cball a B) (frontier (cball a B)) k" | |
| 2002 | by (simp add: retract_of_def) | |
| 2003 | moreover have "retraction (cball a B) (frontier (cball a B)) h" | |
| 2004 | unfolding retraction_def | |
| 2005 | proof (intro conjI ballI) | |
| 2006 | show "frontier (cball a B) \<subseteq> cball a B" | |
| 68022 | 2007 | by force | 
| 64122 | 2008 | show "continuous_on (cball a B) h" | 
| 2009 | unfolding h_def | |
| 68022 | 2010 | by (intro continuous_intros) (use contg continuous_on_subset notga in auto) | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2011 | show "h \<in> cball a B \<rightarrow> frontier (cball a B)" | 
| 64122 | 2012 | using \<open>0 < B\<close> by (auto simp: h_def notga dist_norm) | 
| 2013 | show "\<And>x. x \<in> frontier (cball a B) \<Longrightarrow> h x = x" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2014 | using notga \<open>0 < B\<close> | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2015 | apply (simp add: g_def h_def field_simps) | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2016 | by (metis B dist_commute dist_norm mem_ball order_less_irrefl subset_eq) | 
| 64122 | 2017 | qed | 
| 2018 | ultimately show False by simp | |
| 2019 | qed | |
| 2020 | ||
| 68617 | 2021 | subsubsection \<open>Punctured affine hulls, etc\<close> | 
| 2022 | ||
| 2023 | lemma rel_frontier_deformation_retract_of_punctured_convex: | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2024 | fixes S :: "'a::euclidean_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2025 | assumes "convex S" "convex T" "bounded S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2026 | and arelS: "a \<in> rel_interior S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2027 | and relS: "rel_frontier S \<subseteq> T" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2028 | and affS: "T \<subseteq> affine hull S" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 2029 |   obtains r where "homotopic_with_canon (\<lambda>x. True) (T - {a}) (T - {a}) id r"
 | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2030 |                   "retraction (T - {a}) (rel_frontier S) r"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2031 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2032 | have "\<exists>d. 0 < d \<and> (a + d *\<^sub>R l) \<in> rel_frontier S \<and> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2033 | (\<forall>e. 0 \<le> e \<and> e < d \<longrightarrow> (a + e *\<^sub>R l) \<in> rel_interior S)" | 
| 80095 | 2034 | if "(a + l) \<in> affine hull S" "l \<noteq> 0" for l | 
| 2035 | using ray_to_rel_frontier [OF \<open>bounded S\<close> arelS] that by metis | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2036 | then obtain dd | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2037 | where dd1: "\<And>l. \<lbrakk>(a + l) \<in> affine hull S; l \<noteq> 0\<rbrakk> \<Longrightarrow> 0 < dd l \<and> (a + dd l *\<^sub>R l) \<in> rel_frontier S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2038 | and dd2: "\<And>l e. \<lbrakk>(a + l) \<in> affine hull S; e < dd l; 0 \<le> e; l \<noteq> 0\<rbrakk> | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2039 | \<Longrightarrow> (a + e *\<^sub>R l) \<in> rel_interior S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2040 | by metis+ | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2041 | have aaffS: "a \<in> affine hull S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2042 | by (meson arelS subsetD hull_inc rel_interior_subset) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2043 |   have "((\<lambda>z. z - a) ` (affine hull S - {a})) = ((\<lambda>z. z - a) ` (affine hull S)) - {0}"
 | 
| 68022 | 2044 | by auto | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2045 |   moreover have "continuous_on (((\<lambda>z. z - a) ` (affine hull S)) - {0}) (\<lambda>x. dd x *\<^sub>R x)"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2046 | proof (rule continuous_on_compact_surface_projection) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2047 | show "compact (rel_frontier ((\<lambda>z. z - a) ` S))" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2048 | by (simp add: \<open>bounded S\<close> bounded_translation_minus compact_rel_frontier_bounded) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2049 | have releq: "rel_frontier ((\<lambda>z. z - a) ` S) = (\<lambda>z. z - a) ` rel_frontier S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2050 | using rel_frontier_translation [of "-a"] add.commute by simp | 
| 68022 | 2051 |     also have "\<dots> \<subseteq> (\<lambda>z. z - a) ` (affine hull S) - {0}"
 | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2052 | using rel_frontier_affine_hull arelS rel_frontier_def by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2053 |     finally show "rel_frontier ((\<lambda>z. z - a) ` S) \<subseteq> (\<lambda>z. z - a) ` (affine hull S) - {0}" .
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2054 | show "cone ((\<lambda>z. z - a) ` (affine hull S))" | 
| 68022 | 2055 | by (rule subspace_imp_cone) | 
| 2056 | (use aaffS in \<open>simp add: subspace_affine image_comp o_def affine_translation_aux [of a]\<close>) | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2057 | show "(0 < k \<and> k *\<^sub>R x \<in> rel_frontier ((\<lambda>z. z - a) ` S)) \<longleftrightarrow> (dd x = k)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2058 |          if x: "x \<in> (\<lambda>z. z - a) ` (affine hull S) - {0}" for k x
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2059 | proof | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2060 | show "dd x = k \<Longrightarrow> 0 < k \<and> k *\<^sub>R x \<in> rel_frontier ((\<lambda>z. z - a) ` S)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2061 | using dd1 [of x] that image_iff by (fastforce simp add: releq) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2062 | next | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2063 | assume k: "0 < k \<and> k *\<^sub>R x \<in> rel_frontier ((\<lambda>z. z - a) ` S)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2064 | have False if "dd x < k" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2065 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2066 | have "k \<noteq> 0" "a + k *\<^sub>R x \<in> closure S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2067 | using k closure_translation [of "-a"] | 
| 69661 | 2068 | by (auto simp: rel_frontier_def cong: image_cong_simp) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2069 | then have segsub: "open_segment a (a + k *\<^sub>R x) \<subseteq> rel_interior S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2070 | by (metis rel_interior_closure_convex_segment [OF \<open>convex S\<close> arelS]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2071 | have "x \<noteq> 0" and xaffS: "a + x \<in> affine hull S" | 
| 68022 | 2072 | using x by auto | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2073 | then have "0 < dd x" and inS: "a + dd x *\<^sub>R x \<in> rel_frontier S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2074 | using dd1 by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2075 | moreover have "a + dd x *\<^sub>R x \<in> open_segment a (a + k *\<^sub>R x)" | 
| 80095 | 2076 | unfolding in_segment | 
| 2077 | proof (intro conjI exI) | |
| 2078 | show "a + dd x *\<^sub>R x = (1 - dd x / k) *\<^sub>R a + (dd x / k) *\<^sub>R (a + k *\<^sub>R x)" | |
| 2079 | using k by (simp add: that algebra_simps) | |
| 2080 | qed (use \<open>x \<noteq> 0\<close> \<open>0 < dd x\<close> that in auto) | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2081 | ultimately show ?thesis | 
| 68022 | 2082 | using segsub by (auto simp: rel_frontier_def) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2083 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2084 | moreover have False if "k < dd x" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2085 | using x k that rel_frontier_def | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2086 | by (fastforce simp: algebra_simps releq dest!: dd2) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2087 | ultimately show "dd x = k" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2088 | by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2089 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2090 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2091 |   ultimately have *: "continuous_on ((\<lambda>z. z - a) ` (affine hull S - {a})) (\<lambda>x. dd x *\<^sub>R x)"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2092 | by auto | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2093 |   have "continuous_on (affine hull S - {a}) ((\<lambda>x. a + dd x *\<^sub>R x) \<circ> (\<lambda>z. z - a))"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2094 | by (intro * continuous_intros continuous_on_compose) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2095 |   with affS have contdd: "continuous_on (T - {a}) ((\<lambda>x. a + dd x *\<^sub>R x) \<circ> (\<lambda>z. z - a))"
 | 
| 68022 | 2096 | by (blast intro: continuous_on_subset) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2097 | show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2098 | proof | 
| 80095 | 2099 |     show "homotopic_with_canon (\<lambda>x. True) (T - {a}) (T - {a}) id (\<lambda>x. a + dd (x-a) *\<^sub>R (x-a))"
 | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2100 | proof (rule homotopic_with_linear) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2101 |       show "continuous_on (T - {a}) id"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2102 | by (intro continuous_intros continuous_on_compose) | 
| 80095 | 2103 |       show "continuous_on (T - {a}) (\<lambda>x. a + dd (x-a) *\<^sub>R (x-a))"
 | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2104 | using contdd by (simp add: o_def) | 
| 80095 | 2105 |       show "closed_segment (id x) (a + dd (x-a) *\<^sub>R (x-a)) \<subseteq> T - {a}"
 | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2106 |            if "x \<in> T - {a}" for x
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2107 | proof (clarsimp simp: in_segment, intro conjI) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2108 | fix u::real assume u: "0 \<le> u" "u \<le> 1" | 
| 80095 | 2109 | have "a + dd (x-a) *\<^sub>R (x-a) \<in> T" | 
| 68022 | 2110 | by (metis DiffD1 DiffD2 add.commute add.right_neutral affS dd1 diff_add_cancel relS singletonI subsetCE that) | 
| 80095 | 2111 | then show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x-a) *\<^sub>R (x-a)) \<in> T" | 
| 68022 | 2112 | using convexD [OF \<open>convex T\<close>] that u by simp | 
| 80095 | 2113 | have iff: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + d *\<^sub>R (x-a)) = a \<longleftrightarrow> | 
| 2114 | (1 - u + u * d) *\<^sub>R (x-a) = 0" for d | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2115 | by (auto simp: algebra_simps) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2116 | have "x \<in> T" "x \<noteq> a" using that by auto | 
| 80095 | 2117 | then have axa: "a + (x-a) \<in> affine hull S" | 
| 69712 | 2118 | by (metis (no_types) add.commute affS diff_add_cancel rev_subsetD) | 
| 80095 | 2119 | then have "\<not> dd (x-a) \<le> 0 \<and> a + dd (x-a) *\<^sub>R (x-a) \<in> rel_frontier S" | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2120 | using \<open>x \<noteq> a\<close> dd1 by fastforce | 
| 80095 | 2121 | with \<open>x \<noteq> a\<close> show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x-a) *\<^sub>R (x-a)) \<noteq> a" | 
| 76786 | 2122 | using less_eq_real_def mult_le_0_iff not_less u by (fastforce simp: iff) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2123 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2124 | qed | 
| 80095 | 2125 |     show "retraction (T - {a}) (rel_frontier S) (\<lambda>x. a + dd (x-a) *\<^sub>R (x-a))"
 | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2126 | proof (simp add: retraction_def, intro conjI ballI) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2127 |       show "rel_frontier S \<subseteq> T - {a}"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2128 | using arelS relS rel_frontier_def by fastforce | 
| 80095 | 2129 |       show "continuous_on (T - {a}) (\<lambda>x. a + dd (x-a) *\<^sub>R (x-a))"
 | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2130 | using contdd by (simp add: o_def) | 
| 80095 | 2131 |       show "(\<lambda>x. a + dd (x-a) *\<^sub>R (x-a)) \<in> (T - {a}) \<rightarrow> rel_frontier S"
 | 
| 2132 | unfolding Pi_iff using affS dd1 subset_eq by force | |
| 2133 | show "a + dd (x-a) *\<^sub>R (x-a) = x" if x: "x \<in> rel_frontier S" for x | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2134 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2135 | have "x \<noteq> a" | 
| 68022 | 2136 | using that arelS by (auto simp: rel_frontier_def) | 
| 80095 | 2137 | have False if "dd (x-a) < 1" | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2138 | proof - | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2139 | have "x \<in> closure S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2140 | using x by (auto simp: rel_frontier_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2141 | then have segsub: "open_segment a x \<subseteq> rel_interior S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2142 | by (metis rel_interior_closure_convex_segment [OF \<open>convex S\<close> arelS]) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2143 | have xaffS: "x \<in> affine hull S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2144 | using affS relS x by auto | 
| 80095 | 2145 | then have "0 < dd (x-a)" and inS: "a + dd (x-a) *\<^sub>R (x-a) \<in> rel_frontier S" | 
| 68022 | 2146 | using dd1 by (auto simp: \<open>x \<noteq> a\<close>) | 
| 80095 | 2147 | moreover have "a + dd (x-a) *\<^sub>R (x-a) \<in> open_segment a x" | 
| 2148 | unfolding in_segment | |
| 2149 | proof (intro exI conjI) | |
| 2150 | show "a + dd (x-a) *\<^sub>R (x-a) = (1 - dd (x-a)) *\<^sub>R a + (dd (x-a)) *\<^sub>R x" | |
| 2151 | by (simp add: algebra_simps) | |
| 2152 | qed (use \<open>x \<noteq> a\<close> \<open>0 < dd (x-a)\<close> that in auto) | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2153 | ultimately show ?thesis | 
| 68022 | 2154 | using segsub by (auto simp: rel_frontier_def) | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2155 | qed | 
| 80095 | 2156 | moreover have False if "1 < dd (x-a)" | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2157 | using x that dd2 [of "x - a" 1] \<open>x \<noteq> a\<close> closure_affine_hull | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2158 | by (auto simp: rel_frontier_def) | 
| 80095 | 2159 | ultimately have "dd (x-a) = 1" \<comment> \<open>similar to another proof above\<close> | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2160 | by fastforce | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2161 | with that show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2162 | by (simp add: rel_frontier_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2163 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2164 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2165 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2166 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2167 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2168 | corollary rel_frontier_retract_of_punctured_affine_hull: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2169 | fixes S :: "'a::euclidean_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2170 | assumes "bounded S" "convex S" "a \<in> rel_interior S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2171 |     shows "rel_frontier S retract_of (affine hull S - {a})"
 | 
| 76786 | 2172 | by (meson assms convex_affine_hull dual_order.refl rel_frontier_affine_hull | 
| 2173 | rel_frontier_deformation_retract_of_punctured_convex retract_of_def) | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2174 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2175 | corollary rel_boundary_retract_of_punctured_affine_hull: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2176 | fixes S :: "'a::euclidean_space set" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2177 | assumes "compact S" "convex S" "a \<in> rel_interior S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2178 |     shows "(S - rel_interior S) retract_of (affine hull S - {a})"
 | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2179 | by (metis assms closure_closed compact_eq_bounded_closed rel_frontier_def | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2180 | rel_frontier_retract_of_punctured_affine_hull) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2181 | |
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2182 | lemma homotopy_eqv_rel_frontier_punctured_convex: | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2183 | fixes S :: "'a::euclidean_space set" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2184 | assumes "convex S" "bounded S" "a \<in> rel_interior S" "convex T" "rel_frontier S \<subseteq> T" "T \<subseteq> affine hull S" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2185 |   shows "(rel_frontier S) homotopy_eqv (T - {a})"
 | 
| 76786 | 2186 | by (meson assms deformation_retract_imp_homotopy_eqv homotopy_equivalent_space_sym | 
| 2187 | rel_frontier_deformation_retract_of_punctured_convex[of S T]) | |
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2188 | |
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2189 | lemma homotopy_eqv_rel_frontier_punctured_affine_hull: | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2190 | fixes S :: "'a::euclidean_space set" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2191 | assumes "convex S" "bounded S" "a \<in> rel_interior S" | 
| 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2192 |     shows "(rel_frontier S) homotopy_eqv (affine hull S - {a})"
 | 
| 76786 | 2193 | by (simp add: assms homotopy_eqv_rel_frontier_punctured_convex rel_frontier_affine_hull) | 
| 64789 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 2194 | |
| 64394 | 2195 | lemma path_connected_sphere_gen: | 
| 2196 | assumes "convex S" "bounded S" "aff_dim S \<noteq> 1" | |
| 2197 | shows "path_connected(rel_frontier S)" | |
| 76786 | 2198 | proof - | 
| 2199 | have "convex (closure S)" | |
| 2200 | using assms by auto | |
| 64394 | 2201 | then show ?thesis | 
| 76786 | 2202 | by (metis Diff_empty aff_dim_affine_hull assms convex_affine_hull convex_imp_path_connected equals0I | 
| 2203 | path_connected_punctured_convex rel_frontier_def rel_frontier_retract_of_punctured_affine_hull retract_of_path_connected) | |
| 64394 | 2204 | qed | 
| 2205 | ||
| 2206 | lemma connected_sphere_gen: | |
| 2207 | assumes "convex S" "bounded S" "aff_dim S \<noteq> 1" | |
| 2208 | shows "connected(rel_frontier S)" | |
| 2209 | by (simp add: assms path_connected_imp_connected path_connected_sphere_gen) | |
| 2210 | ||
| 68617 | 2211 | subsubsection\<open>Borsuk-style characterization of separation\<close> | 
| 63301 | 2212 | |
| 2213 | lemma continuous_on_Borsuk_map: | |
| 80095 | 2214 | "a \<notin> S \<Longrightarrow> continuous_on S (\<lambda>x. inverse(norm (x-a)) *\<^sub>R (x-a))" | 
| 63301 | 2215 | by (rule continuous_intros | force)+ | 
| 2216 | ||
| 2217 | lemma Borsuk_map_into_sphere: | |
| 80095 | 2218 | "(\<lambda>x. inverse(norm (x-a)) *\<^sub>R (x-a)) \<in> S \<rightarrow> sphere 0 1 \<longleftrightarrow> (a \<notin> S)" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2219 | proof - | 
| 80095 | 2220 | have "\<And>x. \<lbrakk>a \<notin> S; x \<in> S\<rbrakk> \<Longrightarrow> inverse (norm (x-a)) * norm (x-a) = 1" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2221 | by (metis left_inverse norm_eq_zero right_minus_eq) | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2222 | then show ?thesis | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2223 | by force | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2224 | qed | 
| 63301 | 2225 | |
| 2226 | lemma Borsuk_maps_homotopic_in_path_component: | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2227 | assumes "path_component (- S) a b" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2228 | shows "homotopic_with_canon (\<lambda>x. True) S (sphere 0 1) | 
| 80095 | 2229 | (\<lambda>x. inverse(norm(x-a)) *\<^sub>R (x-a)) | 
| 63301 | 2230 | (\<lambda>x. inverse(norm(x - b)) *\<^sub>R (x - b))" | 
| 2231 | proof - | |
| 80095 | 2232 | obtain g where g: "path g" "path_image g \<subseteq> -S" "pathstart g = a" "pathfinish g = b" | 
| 63301 | 2233 | using assms by (auto simp: path_component_def) | 
| 80095 | 2234 | define h where "h \<equiv> \<lambda>z. (snd z - (g \<circ> fst) z) /\<^sub>R norm (snd z - (g \<circ> fst) z)" | 
| 2235 |   have "continuous_on ({0..1} \<times> S) h"
 | |
| 2236 | unfolding h_def using g by (intro continuous_intros) (auto simp: path_defs) | |
| 2237 | moreover | |
| 2238 |   have "h ` ({0..1} \<times> S) \<subseteq> sphere 0 1"
 | |
| 2239 | unfolding h_def using g by (auto simp: divide_simps path_defs) | |
| 2240 | ultimately show ?thesis | |
| 2241 | using g by (auto simp: h_def path_defs homotopic_with_def) | |
| 63301 | 2242 | qed | 
| 2243 | ||
| 2244 | lemma non_extensible_Borsuk_map: | |
| 2245 | fixes a :: "'a :: euclidean_space" | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2246 | assumes "compact S" and cin: "C \<in> components(- S)" and boc: "bounded C" and "a \<in> C" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2247 | shows "\<not> (\<exists>g. continuous_on (S \<union> C) g \<and> | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2248 | g \<in> (S \<union> C) \<rightarrow> sphere 0 1 \<and> | 
| 80095 | 2249 | (\<forall>x \<in> S. g x = inverse(norm(x-a)) *\<^sub>R (x-a)))" | 
| 63301 | 2250 | proof - | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2251 | have "closed S" using assms by (simp add: compact_imp_closed) | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2252 | have "C \<subseteq> -S" | 
| 63301 | 2253 | using assms by (simp add: in_components_subset) | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2254 | with \<open>a \<in> C\<close> have "a \<notin> S" by blast | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2255 | then have ceq: "C = connected_component_set (- S) a" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2256 | by (metis \<open>a \<in> C\<close> cin components_iff connected_component_eq) | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2257 | then have "bounded (S \<union> connected_component_set (- S) a)" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2258 | using \<open>compact S\<close> boc compact_imp_bounded by auto | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2259 | with bounded_subset_ballD obtain r where "0 < r" and r: "(S \<union> connected_component_set (- S) a) \<subseteq> ball a r" | 
| 63301 | 2260 | by blast | 
| 2261 |   { fix g
 | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2262 | assume "continuous_on (S \<union> C) g" | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2263 | "g \<in> (S \<union> C) \<rightarrow> sphere 0 1" | 
| 80095 | 2264 | and [simp]: "\<And>x. x \<in> S \<Longrightarrow> g x = (x-a) /\<^sub>R norm (x-a)" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2265 | then have norm_g1[simp]: "\<And>x. x \<in> S \<union> C \<Longrightarrow> norm (g x) = 1" | 
| 63301 | 2266 | by force | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2267 | have cb_eq: "cball a r = (S \<union> connected_component_set (- S) a) \<union> | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2268 | (cball a r - connected_component_set (- S) a)" | 
| 63301 | 2269 | using ball_subset_cball [of a r] r by auto | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2270 | have cont1: "continuous_on (S \<union> connected_component_set (- S) a) | 
| 63301 | 2271 | (\<lambda>x. a + r *\<^sub>R g x)" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2272 | using \<open>continuous_on (S \<union> C) g\<close> ceq | 
| 76786 | 2273 | by (intro continuous_intros) blast | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2274 | have cont2: "continuous_on (cball a r - connected_component_set (- S) a) | 
| 80095 | 2275 | (\<lambda>x. a + r *\<^sub>R ((x-a) /\<^sub>R norm (x-a)))" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2276 | by (rule continuous_intros | force simp: \<open>a \<notin> S\<close>)+ | 
| 63301 | 2277 | have 1: "continuous_on (cball a r) | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2278 | (\<lambda>x. if connected_component (- S) a x | 
| 63301 | 2279 | then a + r *\<^sub>R g x | 
| 80095 | 2280 | else a + r *\<^sub>R ((x-a) /\<^sub>R norm (x-a)))" | 
| 63301 | 2281 | apply (subst cb_eq) | 
| 2282 | apply (rule continuous_on_cases [OF _ _ cont1 cont2]) | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2283 | using \<open>closed S\<close> ceq cin | 
| 76786 | 2284 | by (force simp: closed_Diff open_Compl closed_Un_complement_component open_connected_component)+ | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2285 | have 2: "(\<lambda>x. a + r *\<^sub>R g x) ` (cball a r \<inter> connected_component_set (- S) a) | 
| 63301 | 2286 | \<subseteq> sphere a r " | 
| 2287 | using \<open>0 < r\<close> by (force simp: dist_norm ceq) | |
| 2288 | have "retraction (cball a r) (sphere a r) | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2289 | (\<lambda>x. if x \<in> connected_component_set (- S) a | 
| 63301 | 2290 | then a + r *\<^sub>R g x | 
| 80095 | 2291 | else a + r *\<^sub>R ((x-a) /\<^sub>R norm (x-a)))" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2292 | using \<open>0 < r\<close> \<open>a \<notin> S\<close> \<open>a \<in> C\<close> r | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2293 | by (auto simp: norm_minus_commute retraction_def Pi_iff ceq dist_norm abs_if | 
| 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2294 | mult_less_0_iff divide_simps 1 2) | 
| 63301 | 2295 | then have False | 
| 2296 | using no_retraction_cball | |
| 2297 | [OF \<open>0 < r\<close>, of a, unfolded retract_of_def, simplified, rule_format, | |
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
76786diff
changeset | 2298 | of "\<lambda>x. if x \<in> connected_component_set (- S) a | 
| 63301 | 2299 | then a + r *\<^sub>R g x | 
| 80095 | 2300 | else a + r *\<^sub>R inverse(norm(x-a)) *\<^sub>R (x-a)"] | 
| 63301 | 2301 | by blast | 
| 2302 | } | |
| 2303 | then show ?thesis | |
| 2304 | by blast | |
| 2305 | qed | |
| 2306 | ||
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2307 | subsubsection \<open>Proving surjectivity via Brouwer fixpoint theorem\<close> | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2308 | |
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2309 | lemma brouwer_surjective: | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2310 | fixes f :: "'n::euclidean_space \<Rightarrow> 'n" | 
| 80095 | 2311 |   assumes T: "compact T" "convex T" "T \<noteq> {}"
 | 
| 2312 | and f: "continuous_on T f" | |
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2313 | and "\<And>x y. \<lbrakk>x\<in>S; y\<in>T\<rbrakk> \<Longrightarrow> x + (y - f y) \<in> T" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2314 | and "x \<in> S" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2315 | shows "\<exists>y\<in>T. f y = x" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2316 | proof - | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2317 | have *: "\<And>x y. f y = x \<longleftrightarrow> x + (y - f y) = y" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2318 | by (auto simp add: algebra_simps) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2319 | show ?thesis | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2320 | unfolding * | 
| 80095 | 2321 | proof (rule brouwer[OF T]) | 
| 2322 | show "continuous_on T (\<lambda>y. x + (y - f y))" | |
| 2323 | by (intro continuous_intros f) | |
| 2324 | qed (use assms in auto) | |
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2325 | qed | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2326 | |
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2327 | lemma brouwer_surjective_cball: | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2328 | fixes f :: "'n::euclidean_space \<Rightarrow> 'n" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2329 | assumes "continuous_on (cball a e) f" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2330 | and "e > 0" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2331 | and "x \<in> S" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2332 | and "\<And>x y. \<lbrakk>x\<in>S; y\<in>cball a e\<rbrakk> \<Longrightarrow> x + (y - f y) \<in> cball a e" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2333 | shows "\<exists>y\<in>cball a e. f y = x" | 
| 76786 | 2334 | by (smt (verit, best) assms brouwer_surjective cball_eq_empty compact_cball convex_cball) | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2335 | |
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2336 | subsubsection \<open>Inverse function theorem\<close> | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2337 | |
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2338 | text \<open>See Sussmann: "Multidifferential calculus", Theorem 2.1.1\<close> | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2339 | |
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2340 | lemma sussmann_open_mapping: | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2341 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2342 | assumes "open S" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2343 | and contf: "continuous_on S f" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2344 | and "x \<in> S" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2345 | and derf: "(f has_derivative f') (at x)" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2346 | and "bounded_linear g'" "f' \<circ> g' = id" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2347 | and "T \<subseteq> S" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2348 | and x: "x \<in> interior T" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2349 | shows "f x \<in> interior (f ` T)" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2350 | proof - | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2351 | interpret f': bounded_linear f' | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2352 | using assms unfolding has_derivative_def by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2353 | interpret g': bounded_linear g' | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2354 | using assms by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2355 | obtain B where B: "0 < B" "\<forall>x. norm (g' x) \<le> norm x * B" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2356 | using bounded_linear.pos_bounded[OF assms(5)] by blast | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2357 | hence *: "1 / (2 * B) > 0" by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2358 | obtain e0 where e0: | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2359 | "0 < e0" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2360 | "\<forall>y. norm (y - x) < e0 \<longrightarrow> norm (f y - f x - f' (y - x)) \<le> 1 / (2 * B) * norm (y - x)" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2361 | using derf unfolding has_derivative_at_alt | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2362 | using * by blast | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2363 | obtain e1 where e1: "0 < e1" "cball x e1 \<subseteq> T" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2364 | using mem_interior_cball x by blast | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2365 | have *: "0 < e0 / B" "0 < e1 / B" using e0 e1 B by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2366 | obtain e where e: "0 < e" "e < e0 / B" "e < e1 / B" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2367 | using field_lbound_gt_zero[OF *] by blast | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2368 | have lem: "\<exists>y\<in>cball (f x) e. f (x + g' (y - f x)) = z" if "z\<in>cball (f x) (e / 2)" for z | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2369 | proof (rule brouwer_surjective_cball) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2370 | have z: "z \<in> S" if as: "y \<in>cball (f x) e" "z = x + (g' y - g' (f x))" for y z | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2371 | proof- | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2372 | have "dist x z = norm (g' (f x) - g' y)" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2373 | unfolding as(2) and dist_norm by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2374 | also have "\<dots> \<le> norm (f x - y) * B" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2375 | by (metis B(2) g'.diff) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2376 | also have "\<dots> \<le> e * B" | 
| 79566 | 2377 | by (metis B(1) dist_norm mem_cball mult_le_cancel_right_pos that(1)) | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2378 | also have "\<dots> \<le> e1" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2379 | using B(1) e(3) pos_less_divide_eq by fastforce | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2380 | finally have "z \<in> cball x e1" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2381 | by force | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2382 | then show "z \<in> S" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2383 | using e1 assms(7) by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2384 | qed | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2385 | show "continuous_on (cball (f x) e) (\<lambda>y. f (x + g' (y - f x)))" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2386 | unfolding g'.diff | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2387 | proof (rule continuous_on_compose2 [OF _ _ order_refl, of _ _ f]) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2388 | show "continuous_on ((\<lambda>y. x + (g' y - g' (f x))) ` cball (f x) e) f" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2389 | by (rule continuous_on_subset[OF contf]) (use z in blast) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2390 | show "continuous_on (cball (f x) e) (\<lambda>y. x + (g' y - g' (f x)))" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2391 | by (intro continuous_intros linear_continuous_on[OF \<open>bounded_linear g'\<close>]) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2392 | qed | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2393 | next | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2394 | fix y z | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2395 | assume y: "y \<in> cball (f x) (e / 2)" and z: "z \<in> cball (f x) e" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2396 | have "norm (g' (z - f x)) \<le> norm (z - f x) * B" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2397 | using B by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2398 | also have "\<dots> \<le> e * B" | 
| 79566 | 2399 | by (metis B(1) z dist_norm mem_cball norm_minus_commute mult_le_cancel_right_pos) | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2400 | also have "\<dots> < e0" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2401 | using B(1) e(2) pos_less_divide_eq by blast | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2402 | finally have *: "norm (x + g' (z - f x) - x) < e0" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2403 | by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2404 | have **: "f x + f' (x + g' (z - f x) - x) = z" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2405 | using assms(6)[unfolded o_def id_def,THEN cong] | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2406 | by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2407 | have "norm (f x - (y + (z - f (x + g' (z - f x))))) \<le> | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2408 | norm (f (x + g' (z - f x)) - z) + norm (f x - y)" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2409 | using norm_triangle_ineq[of "f (x + g'(z - f x)) - z" "f x - y"] | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2410 | by (auto simp add: algebra_simps) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2411 | also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + norm (f x - y)" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2412 | using e0(2)[rule_format, OF *] | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2413 | by (simp only: algebra_simps **) auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2414 | also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + e/2" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2415 | using y by (auto simp: dist_norm) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2416 | also have "\<dots> \<le> 1 / (B * 2) * B * norm (z - f x) + e/2" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2417 | using * B by (auto simp add: field_simps) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2418 | also have "\<dots> \<le> 1 / 2 * norm (z - f x) + e/2" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2419 | by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2420 | also have "\<dots> \<le> e/2 + e/2" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2421 | using B(1) \<open>norm (z - f x) * B \<le> e * B\<close> by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2422 | finally show "y + (z - f (x + g' (z - f x))) \<in> cball (f x) e" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2423 | by (auto simp: dist_norm) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2424 | qed (use e that in auto) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2425 | show ?thesis | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2426 | unfolding mem_interior | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2427 | proof (intro exI conjI subsetI) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2428 | fix y | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2429 | assume "y \<in> ball (f x) (e / 2)" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2430 | then have *: "y \<in> cball (f x) (e / 2)" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2431 | by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2432 | obtain z where z: "z \<in> cball (f x) e" "f (x + g' (z - f x)) = y" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2433 | using lem * by blast | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2434 | then have "norm (g' (z - f x)) \<le> norm (z - f x) * B" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2435 | using B | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2436 | by (auto simp add: field_simps) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2437 | also have "\<dots> \<le> e * B" | 
| 79566 | 2438 | by (metis B(1) dist_norm mem_cball norm_minus_commute mult_le_cancel_right_pos z(1)) | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2439 | also have "\<dots> \<le> e1" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2440 | using e B unfolding less_divide_eq by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2441 | finally have "x + g'(z - f x) \<in> T" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2442 | by (metis add_diff_cancel diff_diff_add dist_norm e1(2) mem_cball norm_minus_commute subset_eq) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2443 | then show "y \<in> f ` T" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2444 | using z by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2445 | qed (use e in auto) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2446 | qed | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2447 | |
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2448 | text \<open>Hence the following eccentric variant of the inverse function theorem. | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2449 | This has no continuity assumptions, but we do need the inverse function. | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2450 | We could put \<open>f' \<circ> g = I\<close> but this happens to fit with the minimal linear | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2451 | algebra theory I've set up so far.\<close> | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2452 | |
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2453 | lemma has_derivative_inverse_strong: | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2454 | fixes f :: "'n::euclidean_space \<Rightarrow> 'n" | 
| 76786 | 2455 | assumes S: "open S" "x \<in> S" | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2456 | and contf: "continuous_on S f" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2457 | and gf: "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2458 | and derf: "(f has_derivative f') (at x)" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2459 | and id: "f' \<circ> g' = id" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2460 | shows "(g has_derivative g') (at (f x))" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2461 | proof - | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2462 | have linf: "bounded_linear f'" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2463 | using derf unfolding has_derivative_def by auto | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2464 | then have ling: "bounded_linear g'" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2465 | unfolding linear_conv_bounded_linear[symmetric] | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2466 | using id right_inverse_linear by blast | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2467 | moreover have "g' \<circ> f' = id" | 
| 76786 | 2468 | using id linear_inverse_left linear_linear linf ling by blast | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2469 | moreover have *: "\<And>T. \<lbrakk>T \<subseteq> S; x \<in> interior T\<rbrakk> \<Longrightarrow> f x \<in> interior (f ` T)" | 
| 76786 | 2470 | using S derf contf id ling sussmann_open_mapping by blast | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2471 | have "continuous (at (f x)) g" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2472 | unfolding continuous_at Lim_at | 
| 76786 | 2473 | proof (intro strip) | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2474 | fix e :: real | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2475 | assume "e > 0" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2476 | then have "f x \<in> interior (f ` (ball x e \<inter> S))" | 
| 76786 | 2477 | by (simp add: "*" S interior_open) | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2478 | then obtain d where d: "0 < d" "ball (f x) d \<subseteq> f ` (ball x e \<inter> S)" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2479 | unfolding mem_interior by blast | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2480 | show "\<exists>d>0. \<forall>y. 0 < dist y (f x) \<and> dist y (f x) < d \<longrightarrow> dist (g y) (g (f x)) < e" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2481 | proof (intro exI allI impI conjI) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2482 | fix y | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2483 | assume "0 < dist y (f x) \<and> dist y (f x) < d" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2484 | then have "g y \<in> g ` f ` (ball x e \<inter> S)" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2485 | by (metis d(2) dist_commute mem_ball rev_image_eqI subset_iff) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2486 | then show "dist (g y) (g (f x)) < e" | 
| 76786 | 2487 | using \<open>x \<in> S\<close> by (simp add: gf dist_commute image_iff) | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2488 | qed (use d in auto) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2489 | qed | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2490 | moreover have "f x \<in> interior (f ` S)" | 
| 76786 | 2491 | using "*" S interior_eq by blast | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2492 | moreover have "f (g y) = y" if "y \<in> interior (f ` S)" for y | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2493 | by (metis gf imageE interiorE subsetD that) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2494 | ultimately show ?thesis using assms | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2495 | by (metis has_derivative_inverse_basic_x open_interior) | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2496 | qed | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2497 | |
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2498 | text \<open>A rewrite based on the other domain.\<close> | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2499 | |
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2500 | lemma has_derivative_inverse_strong_x: | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2501 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2502 | assumes "open S" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2503 | and "g y \<in> S" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2504 | and "continuous_on S f" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2505 | and "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2506 | and "(f has_derivative f') (at (g y))" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2507 | and "f' \<circ> g' = id" | 
| 76786 | 2508 | and f: "f (g y) = y" | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2509 | shows "(g has_derivative g') (at y)" | 
| 76786 | 2510 | using has_derivative_inverse_strong[OF assms(1-6)] by (simp add: f) | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2511 | |
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2512 | text \<open>On a region.\<close> | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2513 | |
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2514 | theorem has_derivative_inverse_on: | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2515 | fixes f :: "'n::euclidean_space \<Rightarrow> 'n" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2516 | assumes "open S" | 
| 76786 | 2517 | and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f'(x)) (at x)" | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2518 | and "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2519 | and "f' x \<circ> g' x = id" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2520 | and "x \<in> S" | 
| 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2521 | shows "(g has_derivative g'(x)) (at (f x))" | 
| 76786 | 2522 | by (meson assms continuous_on_eq_continuous_at has_derivative_continuous has_derivative_inverse_strong) | 
| 70620 
f95193669ad7
removed Brouwer_Fixpoint from imports of Derivative
 immler parents: 
70136diff
changeset | 2523 | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 2524 | end |