author | nipkow |
Mon, 13 Sep 2010 11:13:15 +0200 | |
changeset 39302 | d7728f65b353 |
parent 37674 | f86de9c00c47 |
child 41958 | 5abc60a017e0 |
permissions | -rw-r--r-- |
36432 | 1 |
(* Author: John Harrison |
2 |
Translation from HOL light: Robert Himmelmann, TU Muenchen *) |
|
3 |
||
4 |
header {* Fashoda meet theorem. *} |
|
5 |
||
6 |
theory Fashoda |
|
37674
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
huffman
parents:
37489
diff
changeset
|
7 |
imports Brouwer_Fixpoint Path_Connected Cartesian_Euclidean_Space |
36432 | 8 |
begin |
9 |
||
10 |
subsection {*Fashoda meet theorem. *} |
|
11 |
||
12 |
lemma infnorm_2: "infnorm (x::real^2) = max (abs(x$1)) (abs(x$2))" |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
13 |
unfolding infnorm_cart UNIV_2 apply(rule Sup_eq) by auto |
36432 | 14 |
|
15 |
lemma infnorm_eq_1_2: "infnorm (x::real^2) = 1 \<longleftrightarrow> |
|
16 |
(abs(x$1) \<le> 1 \<and> abs(x$2) \<le> 1 \<and> (x$1 = -1 \<or> x$1 = 1 \<or> x$2 = -1 \<or> x$2 = 1))" |
|
17 |
unfolding infnorm_2 by auto |
|
18 |
||
19 |
lemma infnorm_eq_1_imp: assumes "infnorm (x::real^2) = 1" shows "abs(x$1) \<le> 1" "abs(x$2) \<le> 1" |
|
20 |
using assms unfolding infnorm_eq_1_2 by auto |
|
21 |
||
22 |
lemma fashoda_unit: fixes f g::"real \<Rightarrow> real^2" |
|
23 |
assumes "f ` {- 1..1} \<subseteq> {- 1..1}" "g ` {- 1..1} \<subseteq> {- 1..1}" |
|
24 |
"continuous_on {- 1..1} f" "continuous_on {- 1..1} g" |
|
25 |
"f (- 1)$1 = - 1" "f 1$1 = 1" "g (- 1) $2 = -1" "g 1 $2 = 1" |
|
26 |
shows "\<exists>s\<in>{- 1..1}. \<exists>t\<in>{- 1..1}. f s = g t" proof(rule ccontr) |
|
27 |
case goal1 note as = this[unfolded bex_simps,rule_format] |
|
28 |
def sqprojection \<equiv> "\<lambda>z::real^2. (inverse (infnorm z)) *\<^sub>R z" |
|
29 |
def negatex \<equiv> "\<lambda>x::real^2. (vector [-(x$1), x$2])::real^2" |
|
30 |
have lem1:"\<forall>z::real^2. infnorm(negatex z) = infnorm z" |
|
31 |
unfolding negatex_def infnorm_2 vector_2 by auto |
|
32 |
have lem2:"\<forall>z. z\<noteq>0 \<longrightarrow> infnorm(sqprojection z) = 1" unfolding sqprojection_def |
|
33 |
unfolding infnorm_mul[unfolded smult_conv_scaleR] unfolding abs_inverse real_abs_infnorm |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
34 |
apply(subst infnorm_eq_0[THEN sym]) by auto |
36432 | 35 |
let ?F = "(\<lambda>w::real^2. (f \<circ> (\<lambda>x. x$1)) w - (g \<circ> (\<lambda>x. x$2)) w)" |
36 |
have *:"\<And>i. (\<lambda>x::real^2. x $ i) ` {- 1..1} = {- 1..1::real}" |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
37674
diff
changeset
|
37 |
apply(rule set_eqI) unfolding image_iff Bex_def mem_interval_cart apply rule defer |
36432 | 38 |
apply(rule_tac x="vec x" in exI) by auto |
39 |
{ fix x assume "x \<in> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w) ` {- 1..1::real^2}" |
|
40 |
then guess w unfolding image_iff .. note w = this |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
41 |
hence "x \<noteq> 0" using as[of "w$1" "w$2"] unfolding mem_interval_cart by auto} note x0=this |
36432 | 42 |
have 21:"\<And>i::2. i\<noteq>1 \<Longrightarrow> i=2" using UNIV_2 by auto |
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
43 |
have 1:"{- 1<..<1::real^2} \<noteq> {}" unfolding interval_eq_empty_cart by auto |
36432 | 44 |
have 2:"continuous_on {- 1..1} (negatex \<circ> sqprojection \<circ> ?F)" apply(rule continuous_on_intros continuous_on_component continuous_on_vec1)+ |
45 |
prefer 2 apply(rule continuous_on_intros continuous_on_component continuous_on_vec1)+ unfolding * |
|
46 |
apply(rule assms)+ apply(rule continuous_on_compose,subst sqprojection_def) |
|
47 |
apply(rule continuous_on_mul ) apply(rule continuous_at_imp_continuous_on,rule) apply(rule continuous_at_inv[unfolded o_def]) |
|
48 |
apply(rule continuous_at_infnorm) unfolding infnorm_eq_0 defer apply(rule continuous_on_id) apply(rule linear_continuous_on) proof- |
|
49 |
show "bounded_linear negatex" apply(rule bounded_linearI') unfolding Cart_eq proof(rule_tac[!] allI) fix i::2 and x y::"real^2" and c::real |
|
36593
fb69c8cd27bd
define linear algebra concepts using scaleR instead of (op *s); generalized many lemmas, though a few theorems that used to work on type int^'n are a bit less general
huffman
parents:
36583
diff
changeset
|
50 |
show "negatex (x + y) $ i = (negatex x + negatex y) $ i" "negatex (c *\<^sub>R x) $ i = (c *\<^sub>R negatex x) $ i" |
36432 | 51 |
apply-apply(case_tac[!] "i\<noteq>1") prefer 3 apply(drule_tac[1-2] 21) |
52 |
unfolding negatex_def by(auto simp add:vector_2 ) qed qed(insert x0, auto) |
|
53 |
have 3:"(negatex \<circ> sqprojection \<circ> ?F) ` {- 1..1} \<subseteq> {- 1..1}" unfolding subset_eq apply rule proof- |
|
54 |
case goal1 then guess y unfolding image_iff .. note y=this have "?F y \<noteq> 0" apply(rule x0) using y(1) by auto |
|
55 |
hence *:"infnorm (sqprojection (?F y)) = 1" unfolding y o_def apply- by(rule lem2[rule_format]) |
|
56 |
have "infnorm x = 1" unfolding *[THEN sym] y o_def by(rule lem1[rule_format]) |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
57 |
thus "x\<in>{- 1..1}" unfolding mem_interval_cart infnorm_2 apply- apply rule |
36432 | 58 |
proof-case goal1 thus ?case apply(cases "i=1") defer apply(drule 21) by auto qed qed |
59 |
guess x apply(rule brouwer_weak[of "{- 1..1::real^2}" "negatex \<circ> sqprojection \<circ> ?F"]) |
|
60 |
apply(rule compact_interval convex_interval)+ unfolding interior_closed_interval |
|
61 |
apply(rule 1 2 3)+ . note x=this |
|
62 |
have "?F x \<noteq> 0" apply(rule x0) using x(1) by auto |
|
63 |
hence *:"infnorm (sqprojection (?F x)) = 1" unfolding o_def by(rule lem2[rule_format]) |
|
64 |
have nx:"infnorm x = 1" apply(subst x(2)[THEN sym]) unfolding *[THEN sym] o_def by(rule lem1[rule_format]) |
|
65 |
have "\<forall>x i. x \<noteq> 0 \<longrightarrow> (0 < (sqprojection x)$i \<longleftrightarrow> 0 < x$i)" "\<forall>x i. x \<noteq> 0 \<longrightarrow> ((sqprojection x)$i < 0 \<longleftrightarrow> x$i < 0)" |
|
66 |
apply- apply(rule_tac[!] allI impI)+ proof- fix x::"real^2" and i::2 assume x:"x\<noteq>0" |
|
67 |
have "inverse (infnorm x) > 0" using x[unfolded infnorm_pos_lt[THEN sym]] by auto |
|
68 |
thus "(0 < sqprojection x $ i) = (0 < x $ i)" "(sqprojection x $ i < 0) = (x $ i < 0)" |
|
69 |
unfolding sqprojection_def vector_component_simps Cart_nth.scaleR real_scaleR_def |
|
70 |
unfolding zero_less_mult_iff mult_less_0_iff by(auto simp add:field_simps) qed |
|
71 |
note lem3 = this[rule_format] |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
72 |
have x1:"x $ 1 \<in> {- 1..1::real}" "x $ 2 \<in> {- 1..1::real}" using x(1) unfolding mem_interval_cart by auto |
36432 | 73 |
hence nz:"f (x $ 1) - g (x $ 2) \<noteq> 0" unfolding right_minus_eq apply-apply(rule as) by auto |
74 |
have "x $ 1 = -1 \<or> x $ 1 = 1 \<or> x $ 2 = -1 \<or> x $ 2 = 1" using nx unfolding infnorm_eq_1_2 by auto |
|
75 |
thus False proof- fix P Q R S |
|
76 |
presume "P \<or> Q \<or> R \<or> S" "P\<Longrightarrow>False" "Q\<Longrightarrow>False" "R\<Longrightarrow>False" "S\<Longrightarrow>False" thus False by auto |
|
77 |
next assume as:"x$1 = 1" |
|
78 |
hence *:"f (x $ 1) $ 1 = 1" using assms(6) by auto |
|
79 |
have "sqprojection (f (x$1) - g (x$2)) $ 1 < 0" |
|
80 |
using x(2)[unfolded o_def Cart_eq,THEN spec[where x=1]] |
|
81 |
unfolding as negatex_def vector_2 by auto moreover |
|
82 |
from x1 have "g (x $ 2) \<in> {- 1..1}" apply-apply(rule assms(2)[unfolded subset_eq,rule_format]) by auto |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
83 |
ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval_cart |
36432 | 84 |
apply(erule_tac x=1 in allE) by auto |
85 |
next assume as:"x$1 = -1" |
|
86 |
hence *:"f (x $ 1) $ 1 = - 1" using assms(5) by auto |
|
87 |
have "sqprojection (f (x$1) - g (x$2)) $ 1 > 0" |
|
88 |
using x(2)[unfolded o_def Cart_eq,THEN spec[where x=1]] |
|
89 |
unfolding as negatex_def vector_2 by auto moreover |
|
90 |
from x1 have "g (x $ 2) \<in> {- 1..1}" apply-apply(rule assms(2)[unfolded subset_eq,rule_format]) by auto |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
91 |
ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval_cart |
36432 | 92 |
apply(erule_tac x=1 in allE) by auto |
93 |
next assume as:"x$2 = 1" |
|
94 |
hence *:"g (x $ 2) $ 2 = 1" using assms(8) by auto |
|
95 |
have "sqprojection (f (x$1) - g (x$2)) $ 2 > 0" |
|
96 |
using x(2)[unfolded o_def Cart_eq,THEN spec[where x=2]] |
|
97 |
unfolding as negatex_def vector_2 by auto moreover |
|
98 |
from x1 have "f (x $ 1) \<in> {- 1..1}" apply-apply(rule assms(1)[unfolded subset_eq,rule_format]) by auto |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
99 |
ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval_cart |
36432 | 100 |
apply(erule_tac x=2 in allE) by auto |
101 |
next assume as:"x$2 = -1" |
|
102 |
hence *:"g (x $ 2) $ 2 = - 1" using assms(7) by auto |
|
103 |
have "sqprojection (f (x$1) - g (x$2)) $ 2 < 0" |
|
104 |
using x(2)[unfolded o_def Cart_eq,THEN spec[where x=2]] |
|
105 |
unfolding as negatex_def vector_2 by auto moreover |
|
106 |
from x1 have "f (x $ 1) \<in> {- 1..1}" apply-apply(rule assms(1)[unfolded subset_eq,rule_format]) by auto |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
107 |
ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval_cart |
36432 | 108 |
apply(erule_tac x=2 in allE) by auto qed(auto) qed |
109 |
||
110 |
lemma fashoda_unit_path: fixes f ::"real \<Rightarrow> real^2" and g ::"real \<Rightarrow> real^2" |
|
111 |
assumes "path f" "path g" "path_image f \<subseteq> {- 1..1}" "path_image g \<subseteq> {- 1..1}" |
|
112 |
"(pathstart f)$1 = -1" "(pathfinish f)$1 = 1" "(pathstart g)$2 = -1" "(pathfinish g)$2 = 1" |
|
113 |
obtains z where "z \<in> path_image f" "z \<in> path_image g" proof- |
|
114 |
note assms=assms[unfolded path_def pathstart_def pathfinish_def path_image_def] |
|
115 |
def iscale \<equiv> "\<lambda>z::real. inverse 2 *\<^sub>R (z + 1)" |
|
116 |
have isc:"iscale ` {- 1..1} \<subseteq> {0..1}" unfolding iscale_def by(auto) |
|
117 |
have "\<exists>s\<in>{- 1..1}. \<exists>t\<in>{- 1..1}. (f \<circ> iscale) s = (g \<circ> iscale) t" proof(rule fashoda_unit) |
|
118 |
show "(f \<circ> iscale) ` {- 1..1} \<subseteq> {- 1..1}" "(g \<circ> iscale) ` {- 1..1} \<subseteq> {- 1..1}" |
|
119 |
using isc and assms(3-4) unfolding image_compose by auto |
|
120 |
have *:"continuous_on {- 1..1} iscale" unfolding iscale_def by(rule continuous_on_intros)+ |
|
121 |
show "continuous_on {- 1..1} (f \<circ> iscale)" "continuous_on {- 1..1} (g \<circ> iscale)" |
|
122 |
apply-apply(rule_tac[!] continuous_on_compose[OF *]) apply(rule_tac[!] continuous_on_subset[OF _ isc]) |
|
123 |
by(rule assms)+ have *:"(1 / 2) *\<^sub>R (1 + (1::real^1)) = 1" unfolding Cart_eq by auto |
|
124 |
show "(f \<circ> iscale) (- 1) $ 1 = - 1" "(f \<circ> iscale) 1 $ 1 = 1" "(g \<circ> iscale) (- 1) $ 2 = -1" "(g \<circ> iscale) 1 $ 2 = 1" |
|
125 |
unfolding o_def iscale_def using assms by(auto simp add:*) qed |
|
126 |
then guess s .. from this(2) guess t .. note st=this |
|
127 |
show thesis apply(rule_tac z="f (iscale s)" in that) |
|
128 |
using st `s\<in>{- 1..1}` unfolding o_def path_image_def image_iff apply- |
|
129 |
apply(rule_tac x="iscale s" in bexI) prefer 3 apply(rule_tac x="iscale t" in bexI) |
|
130 |
using isc[unfolded subset_eq, rule_format] by auto qed |
|
131 |
||
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
132 |
(* move *) |
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
133 |
lemma interval_bij_bij_cart: fixes x::"real^'n" assumes "\<forall>i. a$i < b$i \<and> u$i < v$i" |
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
134 |
shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x" |
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
135 |
unfolding interval_bij_cart split_conv Cart_eq Cart_lambda_beta |
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
136 |
apply(rule,insert assms,erule_tac x=i in allE) by auto |
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
137 |
|
36432 | 138 |
lemma fashoda: fixes b::"real^2" |
139 |
assumes "path f" "path g" "path_image f \<subseteq> {a..b}" "path_image g \<subseteq> {a..b}" |
|
140 |
"(pathstart f)$1 = a$1" "(pathfinish f)$1 = b$1" |
|
141 |
"(pathstart g)$2 = a$2" "(pathfinish g)$2 = b$2" |
|
142 |
obtains z where "z \<in> path_image f" "z \<in> path_image g" proof- |
|
143 |
fix P Q S presume "P \<or> Q \<or> S" "P \<Longrightarrow> thesis" "Q \<Longrightarrow> thesis" "S \<Longrightarrow> thesis" thus thesis by auto |
|
144 |
next have "{a..b} \<noteq> {}" using assms(3) using path_image_nonempty by auto |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
145 |
hence "a \<le> b" unfolding interval_eq_empty_cart vector_le_def by(auto simp add: not_less) |
36432 | 146 |
thus "a$1 = b$1 \<or> a$2 = b$2 \<or> (a$1 < b$1 \<and> a$2 < b$2)" unfolding vector_le_def forall_2 by auto |
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
147 |
next assume as:"a$1 = b$1" have "\<exists>z\<in>path_image g. z$2 = (pathstart f)$2" apply(rule connected_ivt_component_cart) |
36432 | 148 |
apply(rule connected_path_image assms)+apply(rule pathstart_in_path_image,rule pathfinish_in_path_image) |
149 |
unfolding assms using assms(3)[unfolded path_image_def subset_eq,rule_format,of "f 0"] |
|
150 |
unfolding pathstart_def by(auto simp add: vector_le_def) then guess z .. note z=this |
|
151 |
have "z \<in> {a..b}" using z(1) assms(4) unfolding path_image_def by blast |
|
152 |
hence "z = f 0" unfolding Cart_eq forall_2 unfolding z(2) pathstart_def |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
153 |
using assms(3)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "f 0" 1] |
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
154 |
unfolding mem_interval_cart apply(erule_tac x=1 in allE) using as by auto |
36432 | 155 |
thus thesis apply-apply(rule that[OF _ z(1)]) unfolding path_image_def by auto |
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
156 |
next assume as:"a$2 = b$2" have "\<exists>z\<in>path_image f. z$1 = (pathstart g)$1" apply(rule connected_ivt_component_cart) |
36432 | 157 |
apply(rule connected_path_image assms)+apply(rule pathstart_in_path_image,rule pathfinish_in_path_image) |
158 |
unfolding assms using assms(4)[unfolded path_image_def subset_eq,rule_format,of "g 0"] |
|
159 |
unfolding pathstart_def by(auto simp add: vector_le_def) then guess z .. note z=this |
|
160 |
have "z \<in> {a..b}" using z(1) assms(3) unfolding path_image_def by blast |
|
161 |
hence "z = g 0" unfolding Cart_eq forall_2 unfolding z(2) pathstart_def |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
162 |
using assms(4)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "g 0" 2] |
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
163 |
unfolding mem_interval_cart apply(erule_tac x=2 in allE) using as by auto |
36432 | 164 |
thus thesis apply-apply(rule that[OF z(1)]) unfolding path_image_def by auto |
165 |
next assume as:"a $ 1 < b $ 1 \<and> a $ 2 < b $ 2" |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
166 |
have int_nem:"{- 1..1::real^2} \<noteq> {}" unfolding interval_eq_empty_cart by auto |
36432 | 167 |
guess z apply(rule fashoda_unit_path[of "interval_bij (a,b) (- 1,1) \<circ> f" "interval_bij (a,b) (- 1,1) \<circ> g"]) |
168 |
unfolding path_def path_image_def pathstart_def pathfinish_def |
|
169 |
apply(rule_tac[1-2] continuous_on_compose) apply(rule assms[unfolded path_def] continuous_on_interval_bij)+ |
|
170 |
unfolding subset_eq apply(rule_tac[1-2] ballI) |
|
171 |
proof- fix x assume "x \<in> (interval_bij (a, b) (- 1, 1) \<circ> f) ` {0..1}" |
|
172 |
then guess y unfolding image_iff .. note y=this |
|
173 |
show "x \<in> {- 1..1}" unfolding y o_def apply(rule in_interval_interval_bij) |
|
174 |
using y(1) using assms(3)[unfolded path_image_def subset_eq] int_nem by auto |
|
175 |
next fix x assume "x \<in> (interval_bij (a, b) (- 1, 1) \<circ> g) ` {0..1}" |
|
176 |
then guess y unfolding image_iff .. note y=this |
|
177 |
show "x \<in> {- 1..1}" unfolding y o_def apply(rule in_interval_interval_bij) |
|
178 |
using y(1) using assms(4)[unfolded path_image_def subset_eq] int_nem by auto |
|
179 |
next show "(interval_bij (a, b) (- 1, 1) \<circ> f) 0 $ 1 = -1" |
|
180 |
"(interval_bij (a, b) (- 1, 1) \<circ> f) 1 $ 1 = 1" |
|
181 |
"(interval_bij (a, b) (- 1, 1) \<circ> g) 0 $ 2 = -1" |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
182 |
"(interval_bij (a, b) (- 1, 1) \<circ> g) 1 $ 2 = 1" |
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
183 |
unfolding interval_bij_cart Cart_lambda_beta vector_component_simps o_def split_conv |
36432 | 184 |
unfolding assms[unfolded pathstart_def pathfinish_def] using as by auto qed note z=this |
185 |
from z(1) guess zf unfolding image_iff .. note zf=this |
|
186 |
from z(2) guess zg unfolding image_iff .. note zg=this |
|
187 |
have *:"\<forall>i. (- 1) $ i < (1::real^2) $ i \<and> a $ i < b $ i" unfolding forall_2 using as by auto |
|
188 |
show thesis apply(rule_tac z="interval_bij (- 1,1) (a,b) z" in that) |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
189 |
apply(subst zf) defer apply(subst zg) unfolding o_def interval_bij_bij_cart[OF *] path_image_def |
36432 | 190 |
using zf(1) zg(1) by auto qed |
191 |
||
192 |
subsection {*Some slightly ad hoc lemmas I use below*} |
|
193 |
||
194 |
lemma segment_vertical: fixes a::"real^2" assumes "a$1 = b$1" |
|
195 |
shows "x \<in> closed_segment a b \<longleftrightarrow> (x$1 = a$1 \<and> x$1 = b$1 \<and> |
|
196 |
(a$2 \<le> x$2 \<and> x$2 \<le> b$2 \<or> b$2 \<le> x$2 \<and> x$2 \<le> a$2))" (is "_ = ?R") |
|
197 |
proof- |
|
198 |
let ?L = "\<exists>u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \<and> x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \<and> 0 \<le> u \<and> u \<le> 1" |
|
199 |
{ presume "?L \<Longrightarrow> ?R" "?R \<Longrightarrow> ?L" thus ?thesis unfolding closed_segment_def mem_Collect_eq |
|
200 |
unfolding Cart_eq forall_2 smult_conv_scaleR[THEN sym] vector_component_simps by blast } |
|
201 |
{ assume ?L then guess u apply-apply(erule exE)apply(erule conjE)+ . note u=this |
|
202 |
{ fix b a assume "b + u * a > a + u * b" |
|
203 |
hence "(1 - u) * b > (1 - u) * a" by(auto simp add:field_simps) |
|
204 |
hence "b \<ge> a" apply(drule_tac mult_less_imp_less_left) using u by auto |
|
205 |
hence "u * a \<le> u * b" apply-apply(rule mult_left_mono[OF _ u(3)]) |
|
206 |
using u(3-4) by(auto simp add:field_simps) } note * = this |
|
207 |
{ fix a b assume "u * b > u * a" hence "(1 - u) * a \<le> (1 - u) * b" apply-apply(rule mult_left_mono) |
|
208 |
apply(drule mult_less_imp_less_left) using u by auto |
|
209 |
hence "a + u * b \<le> b + u * a" by(auto simp add:field_simps) } note ** = this |
|
210 |
thus ?R unfolding u assms using u by(auto simp add:field_simps not_le intro:* **) } |
|
211 |
{ assume ?R thus ?L proof(cases "x$2 = b$2") |
|
212 |
case True thus ?L apply(rule_tac x="(x$2 - a$2) / (b$2 - a$2)" in exI) unfolding assms True |
|
213 |
using `?R` by(auto simp add:field_simps) |
|
214 |
next case False thus ?L apply(rule_tac x="1 - (x$2 - b$2) / (a$2 - b$2)" in exI) unfolding assms using `?R` |
|
215 |
by(auto simp add:field_simps) |
|
216 |
qed } qed |
|
217 |
||
218 |
lemma segment_horizontal: fixes a::"real^2" assumes "a$2 = b$2" |
|
219 |
shows "x \<in> closed_segment a b \<longleftrightarrow> (x$2 = a$2 \<and> x$2 = b$2 \<and> |
|
220 |
(a$1 \<le> x$1 \<and> x$1 \<le> b$1 \<or> b$1 \<le> x$1 \<and> x$1 \<le> a$1))" (is "_ = ?R") |
|
221 |
proof- |
|
222 |
let ?L = "\<exists>u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \<and> x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \<and> 0 \<le> u \<and> u \<le> 1" |
|
223 |
{ presume "?L \<Longrightarrow> ?R" "?R \<Longrightarrow> ?L" thus ?thesis unfolding closed_segment_def mem_Collect_eq |
|
224 |
unfolding Cart_eq forall_2 smult_conv_scaleR[THEN sym] vector_component_simps by blast } |
|
225 |
{ assume ?L then guess u apply-apply(erule exE)apply(erule conjE)+ . note u=this |
|
226 |
{ fix b a assume "b + u * a > a + u * b" |
|
227 |
hence "(1 - u) * b > (1 - u) * a" by(auto simp add:field_simps) |
|
228 |
hence "b \<ge> a" apply(drule_tac mult_less_imp_less_left) using u by auto |
|
229 |
hence "u * a \<le> u * b" apply-apply(rule mult_left_mono[OF _ u(3)]) |
|
230 |
using u(3-4) by(auto simp add:field_simps) } note * = this |
|
231 |
{ fix a b assume "u * b > u * a" hence "(1 - u) * a \<le> (1 - u) * b" apply-apply(rule mult_left_mono) |
|
232 |
apply(drule mult_less_imp_less_left) using u by auto |
|
233 |
hence "a + u * b \<le> b + u * a" by(auto simp add:field_simps) } note ** = this |
|
234 |
thus ?R unfolding u assms using u by(auto simp add:field_simps not_le intro:* **) } |
|
235 |
{ assume ?R thus ?L proof(cases "x$1 = b$1") |
|
236 |
case True thus ?L apply(rule_tac x="(x$1 - a$1) / (b$1 - a$1)" in exI) unfolding assms True |
|
237 |
using `?R` by(auto simp add:field_simps) |
|
238 |
next case False thus ?L apply(rule_tac x="1 - (x$1 - b$1) / (a$1 - b$1)" in exI) unfolding assms using `?R` |
|
239 |
by(auto simp add:field_simps) |
|
240 |
qed } qed |
|
241 |
||
242 |
subsection {*useful Fashoda corollary pointed out to me by Tom Hales. *} |
|
243 |
||
244 |
lemma fashoda_interlace: fixes a::"real^2" |
|
245 |
assumes "path f" "path g" |
|
246 |
"path_image f \<subseteq> {a..b}" "path_image g \<subseteq> {a..b}" |
|
247 |
"(pathstart f)$2 = a$2" "(pathfinish f)$2 = a$2" |
|
248 |
"(pathstart g)$2 = a$2" "(pathfinish g)$2 = a$2" |
|
249 |
"(pathstart f)$1 < (pathstart g)$1" "(pathstart g)$1 < (pathfinish f)$1" |
|
250 |
"(pathfinish f)$1 < (pathfinish g)$1" |
|
251 |
obtains z where "z \<in> path_image f" "z \<in> path_image g" |
|
252 |
proof- |
|
253 |
have "{a..b} \<noteq> {}" using path_image_nonempty using assms(3) by auto |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
254 |
note ab=this[unfolded interval_eq_empty_cart not_ex forall_2 not_less] |
36432 | 255 |
have "pathstart f \<in> {a..b}" "pathfinish f \<in> {a..b}" "pathstart g \<in> {a..b}" "pathfinish g \<in> {a..b}" |
256 |
using pathstart_in_path_image pathfinish_in_path_image using assms(3-4) by auto |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
257 |
note startfin = this[unfolded mem_interval_cart forall_2] |
36432 | 258 |
let ?P1 = "linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2]) +++ |
259 |
linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f) +++ f +++ |
|
260 |
linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2]) +++ |
|
261 |
linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2])" |
|
262 |
let ?P2 = "linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g) +++ g +++ |
|
263 |
linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1]) +++ |
|
264 |
linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1]) +++ |
|
265 |
linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3])" |
|
266 |
let ?a = "vector[a$1 - 2, a$2 - 3]" |
|
267 |
let ?b = "vector[b$1 + 2, b$2 + 3]" |
|
268 |
have P1P2:"path_image ?P1 = path_image (linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2])) \<union> |
|
269 |
path_image (linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f)) \<union> path_image f \<union> |
|
270 |
path_image (linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2])) \<union> |
|
271 |
path_image (linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2]))" |
|
272 |
"path_image ?P2 = path_image(linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g)) \<union> path_image g \<union> |
|
273 |
path_image(linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1])) \<union> |
|
274 |
path_image(linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1])) \<union> |
|
275 |
path_image(linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3]))" using assms(1-2) |
|
276 |
by(auto simp add: path_image_join path_linepath) |
|
277 |
have abab: "{a..b} \<subseteq> {?a..?b}" by(auto simp add:vector_le_def forall_2 vector_2) |
|
278 |
guess z apply(rule fashoda[of ?P1 ?P2 ?a ?b]) |
|
279 |
unfolding pathstart_join pathfinish_join pathstart_linepath pathfinish_linepath vector_2 proof- |
|
280 |
show "path ?P1" "path ?P2" using assms by auto |
|
281 |
have "path_image ?P1 \<subseteq> {?a .. ?b}" unfolding P1P2 path_image_linepath apply(rule Un_least)+ defer 3 |
|
282 |
apply(rule_tac[1-4] convex_interval(1)[unfolded convex_contains_segment,rule_format]) |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
283 |
unfolding mem_interval_cart forall_2 vector_2 using ab startfin abab assms(3) |
36432 | 284 |
using assms(9-) unfolding assms by(auto simp add:field_simps) |
285 |
thus "path_image ?P1 \<subseteq> {?a .. ?b}" . |
|
286 |
have "path_image ?P2 \<subseteq> {?a .. ?b}" unfolding P1P2 path_image_linepath apply(rule Un_least)+ defer 2 |
|
287 |
apply(rule_tac[1-4] convex_interval(1)[unfolded convex_contains_segment,rule_format]) |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
288 |
unfolding mem_interval_cart forall_2 vector_2 using ab startfin abab assms(4) |
36432 | 289 |
using assms(9-) unfolding assms by(auto simp add:field_simps) |
290 |
thus "path_image ?P2 \<subseteq> {?a .. ?b}" . |
|
291 |
show "a $ 1 - 2 = a $ 1 - 2" "b $ 1 + 2 = b $ 1 + 2" "pathstart g $ 2 - 3 = a $ 2 - 3" "b $ 2 + 3 = b $ 2 + 3" |
|
292 |
by(auto simp add: assms) |
|
293 |
qed note z=this[unfolded P1P2 path_image_linepath] |
|
294 |
show thesis apply(rule that[of z]) proof- |
|
295 |
have "(z \<in> closed_segment (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) \<or> |
|
296 |
z \<in> closed_segment (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f)) \<or> |
|
297 |
z \<in> closed_segment (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) \<or> |
|
298 |
z \<in> closed_segment (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]) \<Longrightarrow> |
|
299 |
(((z \<in> closed_segment (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g)) \<or> |
|
300 |
z \<in> closed_segment (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1])) \<or> |
|
301 |
z \<in> closed_segment (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1])) \<or> |
|
302 |
z \<in> closed_segment (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]) \<Longrightarrow> False" |
|
303 |
apply(simp only: segment_vertical segment_horizontal vector_2) proof- case goal1 note as=this |
|
304 |
have "pathfinish f \<in> {a..b}" using assms(3) pathfinish_in_path_image[of f] by auto |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
305 |
hence "1 + b $ 1 \<le> pathfinish f $ 1 \<Longrightarrow> False" unfolding mem_interval_cart forall_2 by auto |
36432 | 306 |
hence "z$1 \<noteq> pathfinish f$1" using as(2) using assms ab by(auto simp add:field_simps) |
307 |
moreover have "pathstart f \<in> {a..b}" using assms(3) pathstart_in_path_image[of f] by auto |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
308 |
hence "1 + b $ 1 \<le> pathstart f $ 1 \<Longrightarrow> False" unfolding mem_interval_cart forall_2 by auto |
36432 | 309 |
hence "z$1 \<noteq> pathstart f$1" using as(2) using assms ab by(auto simp add:field_simps) |
310 |
ultimately have *:"z$2 = a$2 - 2" using goal1(1) by auto |
|
311 |
have "z$1 \<noteq> pathfinish g$1" using as(2) using assms ab by(auto simp add:field_simps *) |
|
312 |
moreover have "pathstart g \<in> {a..b}" using assms(4) pathstart_in_path_image[of g] by auto |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
313 |
note this[unfolded mem_interval_cart forall_2] |
36432 | 314 |
hence "z$1 \<noteq> pathstart g$1" using as(1) using assms ab by(auto simp add:field_simps *) |
315 |
ultimately have "a $ 2 - 1 \<le> z $ 2 \<and> z $ 2 \<le> b $ 2 + 3 \<or> b $ 2 + 3 \<le> z $ 2 \<and> z $ 2 \<le> a $ 2 - 1" |
|
316 |
using as(2) unfolding * assms by(auto simp add:field_simps) |
|
317 |
thus False unfolding * using ab by auto |
|
318 |
qed hence "z \<in> path_image f \<or> z \<in> path_image g" using z unfolding Un_iff by blast |
|
319 |
hence z':"z\<in>{a..b}" using assms(3-4) by auto |
|
320 |
have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart f $ 1 \<or> z $ 1 = pathfinish f $ 1) \<Longrightarrow> (z = pathstart f \<or> z = pathfinish f)" |
|
321 |
unfolding Cart_eq forall_2 assms by auto |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
322 |
with z' show "z\<in>path_image f" using z(1) unfolding Un_iff mem_interval_cart forall_2 apply- |
36432 | 323 |
apply(simp only: segment_vertical segment_horizontal vector_2) unfolding assms by auto |
324 |
have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart g $ 1 \<or> z $ 1 = pathfinish g $ 1) \<Longrightarrow> (z = pathstart g \<or> z = pathfinish g)" |
|
325 |
unfolding Cart_eq forall_2 assms by auto |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
326 |
with z' show "z\<in>path_image g" using z(2) unfolding Un_iff mem_interval_cart forall_2 apply- |
36432 | 327 |
apply(simp only: segment_vertical segment_horizontal vector_2) unfolding assms by auto |
328 |
qed qed |
|
329 |
||
330 |
(** The Following still needs to be translated. Maybe I will do that later. |
|
331 |
||
332 |
(* ------------------------------------------------------------------------- *) |
|
333 |
(* Complement in dimension N >= 2 of set homeomorphic to any interval in *) |
|
334 |
(* any dimension is (path-)connected. This naively generalizes the argument *) |
|
335 |
(* in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer *) |
|
336 |
(* fixed point theorem", American Mathematical Monthly 1984. *) |
|
337 |
(* ------------------------------------------------------------------------- *) |
|
338 |
||
339 |
let RETRACTION_INJECTIVE_IMAGE_INTERVAL = prove |
|
340 |
(`!p:real^M->real^N a b. |
|
341 |
~(interval[a,b] = {}) /\ |
|
342 |
p continuous_on interval[a,b] /\ |
|
343 |
(!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ p x = p y ==> x = y) |
|
344 |
==> ?f. f continuous_on (:real^N) /\ |
|
345 |
IMAGE f (:real^N) SUBSET (IMAGE p (interval[a,b])) /\ |
|
346 |
(!x. x IN (IMAGE p (interval[a,b])) ==> f x = x)`, |
|
347 |
REPEAT STRIP_TAC THEN |
|
348 |
FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [INJECTIVE_ON_LEFT_INVERSE]) THEN |
|
349 |
DISCH_THEN(X_CHOOSE_TAC `q:real^N->real^M`) THEN |
|
350 |
SUBGOAL_THEN `(q:real^N->real^M) continuous_on |
|
351 |
(IMAGE p (interval[a:real^M,b]))` |
|
352 |
ASSUME_TAC THENL |
|
353 |
[MATCH_MP_TAC CONTINUOUS_ON_INVERSE THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]; |
|
354 |
ALL_TAC] THEN |
|
355 |
MP_TAC(ISPECL [`q:real^N->real^M`; |
|
356 |
`IMAGE (p:real^M->real^N) |
|
357 |
(interval[a,b])`; |
|
358 |
`a:real^M`; `b:real^M`] |
|
359 |
TIETZE_CLOSED_INTERVAL) THEN |
|
360 |
ASM_SIMP_TAC[COMPACT_INTERVAL; COMPACT_CONTINUOUS_IMAGE; |
|
361 |
COMPACT_IMP_CLOSED] THEN |
|
362 |
ANTS_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN |
|
363 |
DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^M` STRIP_ASSUME_TAC) THEN |
|
364 |
EXISTS_TAC `(p:real^M->real^N) o (r:real^N->real^M)` THEN |
|
365 |
REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; o_THM; IN_UNIV] THEN |
|
366 |
CONJ_TAC THENL [ALL_TAC; ASM SET_TAC[]] THEN |
|
367 |
MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN ASM_REWRITE_TAC[] THEN |
|
368 |
FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP(REWRITE_RULE[IMP_CONJ] |
|
369 |
CONTINUOUS_ON_SUBSET)) THEN ASM SET_TAC[]);; |
|
370 |
||
371 |
let UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove |
|
372 |
(`!s:real^N->bool a b:real^M. |
|
373 |
s homeomorphic (interval[a,b]) |
|
374 |
==> !x. ~(x IN s) ==> ~bounded(path_component((:real^N) DIFF s) x)`, |
|
375 |
REPEAT GEN_TAC THEN REWRITE_TAC[homeomorphic; homeomorphism] THEN |
|
376 |
REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN |
|
377 |
MAP_EVERY X_GEN_TAC [`p':real^N->real^M`; `p:real^M->real^N`] THEN |
|
378 |
DISCH_TAC THEN |
|
379 |
SUBGOAL_THEN |
|
380 |
`!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ |
|
381 |
(p:real^M->real^N) x = p y ==> x = y` |
|
382 |
ASSUME_TAC THENL [ASM_MESON_TAC[]; ALL_TAC] THEN |
|
383 |
FIRST_X_ASSUM(MP_TAC o funpow 4 CONJUNCT2) THEN |
|
384 |
DISCH_THEN(CONJUNCTS_THEN2 (SUBST1_TAC o SYM) ASSUME_TAC) THEN |
|
385 |
ASM_CASES_TAC `interval[a:real^M,b] = {}` THEN |
|
386 |
ASM_REWRITE_TAC[IMAGE_CLAUSES; DIFF_EMPTY; PATH_COMPONENT_UNIV; |
|
387 |
NOT_BOUNDED_UNIV] THEN |
|
388 |
ABBREV_TAC `s = (:real^N) DIFF (IMAGE p (interval[a:real^M,b]))` THEN |
|
389 |
X_GEN_TAC `c:real^N` THEN REPEAT STRIP_TAC THEN |
|
390 |
SUBGOAL_THEN `(c:real^N) IN s` ASSUME_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN |
|
391 |
SUBGOAL_THEN `bounded((path_component s c) UNION |
|
392 |
(IMAGE (p:real^M->real^N) (interval[a,b])))` |
|
393 |
MP_TAC THENL |
|
394 |
[ASM_SIMP_TAC[BOUNDED_UNION; COMPACT_IMP_BOUNDED; COMPACT_IMP_BOUNDED; |
|
395 |
COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; |
|
396 |
ALL_TAC] THEN |
|
397 |
DISCH_THEN(MP_TAC o SPEC `c:real^N` o MATCH_MP BOUNDED_SUBSET_BALL) THEN |
|
398 |
REWRITE_TAC[UNION_SUBSET] THEN |
|
399 |
DISCH_THEN(X_CHOOSE_THEN `B:real` STRIP_ASSUME_TAC) THEN |
|
400 |
MP_TAC(ISPECL [`p:real^M->real^N`; `a:real^M`; `b:real^M`] |
|
401 |
RETRACTION_INJECTIVE_IMAGE_INTERVAL) THEN |
|
402 |
ASM_REWRITE_TAC[SUBSET; IN_UNIV] THEN |
|
403 |
DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^N` MP_TAC) THEN |
|
404 |
DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC |
|
405 |
(CONJUNCTS_THEN2 MP_TAC ASSUME_TAC)) THEN |
|
406 |
REWRITE_TAC[FORALL_IN_IMAGE; IN_UNIV] THEN DISCH_TAC THEN |
|
407 |
ABBREV_TAC `q = \z:real^N. if z IN path_component s c then r(z) else z` THEN |
|
408 |
SUBGOAL_THEN |
|
409 |
`(q:real^N->real^N) continuous_on |
|
410 |
(closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)))` |
|
411 |
MP_TAC THENL |
|
412 |
[EXPAND_TAC "q" THEN MATCH_MP_TAC CONTINUOUS_ON_CASES THEN |
|
413 |
REWRITE_TAC[CLOSED_CLOSURE; CONTINUOUS_ON_ID; GSYM OPEN_CLOSED] THEN |
|
414 |
REPEAT CONJ_TAC THENL |
|
415 |
[MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN |
|
416 |
ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED; |
|
417 |
COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; |
|
418 |
ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; |
|
419 |
ALL_TAC] THEN |
|
420 |
X_GEN_TAC `z:real^N` THEN |
|
421 |
REWRITE_TAC[SET_RULE `~(z IN (s DIFF t) /\ z IN t)`] THEN |
|
422 |
STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN |
|
423 |
MP_TAC(ISPECL |
|
424 |
[`path_component s (z:real^N)`; `path_component s (c:real^N)`] |
|
425 |
OPEN_INTER_CLOSURE_EQ_EMPTY) THEN |
|
426 |
ASM_REWRITE_TAC[GSYM DISJOINT; PATH_COMPONENT_DISJOINT] THEN ANTS_TAC THENL |
|
427 |
[MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN |
|
428 |
ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED; |
|
429 |
COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; |
|
430 |
REWRITE_TAC[DISJOINT; EXTENSION; IN_INTER; NOT_IN_EMPTY] THEN |
|
431 |
DISCH_THEN(MP_TAC o SPEC `z:real^N`) THEN ASM_REWRITE_TAC[] THEN |
|
432 |
GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [IN] THEN |
|
433 |
REWRITE_TAC[PATH_COMPONENT_REFL_EQ] THEN ASM SET_TAC[]]; |
|
434 |
ALL_TAC] THEN |
|
435 |
SUBGOAL_THEN |
|
436 |
`closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)) = |
|
437 |
(:real^N)` |
|
438 |
SUBST1_TAC THENL |
|
439 |
[MATCH_MP_TAC(SET_RULE `s SUBSET t ==> t UNION (UNIV DIFF s) = UNIV`) THEN |
|
440 |
REWRITE_TAC[CLOSURE_SUBSET]; |
|
441 |
DISCH_TAC] THEN |
|
442 |
MP_TAC(ISPECL |
|
443 |
[`(\x. &2 % c - x) o |
|
444 |
(\x. c + B / norm(x - c) % (x - c)) o (q:real^N->real^N)`; |
|
445 |
`cball(c:real^N,B)`] |
|
446 |
BROUWER) THEN |
|
447 |
REWRITE_TAC[NOT_IMP; GSYM CONJ_ASSOC; COMPACT_CBALL; CONVEX_CBALL] THEN |
|
448 |
ASM_SIMP_TAC[CBALL_EQ_EMPTY; REAL_LT_IMP_LE; REAL_NOT_LT] THEN |
|
449 |
SUBGOAL_THEN `!x. ~((q:real^N->real^N) x = c)` ASSUME_TAC THENL |
|
450 |
[X_GEN_TAC `x:real^N` THEN EXPAND_TAC "q" THEN |
|
451 |
REWRITE_TAC[NORM_EQ_0; VECTOR_SUB_EQ] THEN COND_CASES_TAC THEN |
|
452 |
ASM SET_TAC[PATH_COMPONENT_REFL_EQ]; |
|
453 |
ALL_TAC] THEN |
|
454 |
REPEAT CONJ_TAC THENL |
|
455 |
[MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN |
|
456 |
SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN |
|
457 |
MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN CONJ_TAC THENL |
|
458 |
[ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; ALL_TAC] THEN |
|
459 |
MATCH_MP_TAC CONTINUOUS_ON_ADD THEN REWRITE_TAC[CONTINUOUS_ON_CONST] THEN |
|
460 |
MATCH_MP_TAC CONTINUOUS_ON_MUL THEN |
|
461 |
SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN |
|
462 |
REWRITE_TAC[o_DEF; real_div; LIFT_CMUL] THEN |
|
463 |
MATCH_MP_TAC CONTINUOUS_ON_CMUL THEN |
|
464 |
MATCH_MP_TAC(REWRITE_RULE[o_DEF] CONTINUOUS_ON_INV) THEN |
|
465 |
ASM_REWRITE_TAC[FORALL_IN_IMAGE; NORM_EQ_0; VECTOR_SUB_EQ] THEN |
|
466 |
SUBGOAL_THEN |
|
467 |
`(\x:real^N. lift(norm(x - c))) = (lift o norm) o (\x. x - c)` |
|
468 |
SUBST1_TAC THENL [REWRITE_TAC[o_DEF]; ALL_TAC] THEN |
|
469 |
MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN |
|
470 |
ASM_SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST; |
|
471 |
CONTINUOUS_ON_LIFT_NORM]; |
|
472 |
REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; IN_CBALL; o_THM; dist] THEN |
|
473 |
X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN |
|
474 |
REWRITE_TAC[VECTOR_ARITH `c - (&2 % c - (c + x)) = x`] THEN |
|
475 |
REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN |
|
476 |
ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN |
|
477 |
ASM_REAL_ARITH_TAC; |
|
478 |
REWRITE_TAC[NOT_EXISTS_THM; TAUT `~(c /\ b) <=> c ==> ~b`] THEN |
|
479 |
REWRITE_TAC[IN_CBALL; o_THM; dist] THEN |
|
480 |
X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN |
|
481 |
REWRITE_TAC[VECTOR_ARITH `&2 % c - (c + x') = x <=> --x' = x - c`] THEN |
|
482 |
ASM_CASES_TAC `(x:real^N) IN path_component s c` THENL |
|
483 |
[MATCH_MP_TAC(NORM_ARITH `norm(y) < B /\ norm(x) = B ==> ~(--x = y)`) THEN |
|
484 |
REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN |
|
485 |
ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN |
|
486 |
ASM_SIMP_TAC[REAL_ARITH `&0 < B ==> abs B = B`] THEN |
|
487 |
UNDISCH_TAC `path_component s c SUBSET ball(c:real^N,B)` THEN |
|
488 |
REWRITE_TAC[SUBSET; IN_BALL] THEN ASM_MESON_TAC[dist; NORM_SUB]; |
|
489 |
EXPAND_TAC "q" THEN REWRITE_TAC[] THEN ASM_REWRITE_TAC[] THEN |
|
490 |
REWRITE_TAC[VECTOR_ARITH `--(c % x) = x <=> (&1 + c) % x = vec 0`] THEN |
|
491 |
ASM_REWRITE_TAC[DE_MORGAN_THM; VECTOR_SUB_EQ; VECTOR_MUL_EQ_0] THEN |
|
492 |
SUBGOAL_THEN `~(x:real^N = c)` ASSUME_TAC THENL |
|
493 |
[ASM_MESON_TAC[PATH_COMPONENT_REFL; IN]; ALL_TAC] THEN |
|
494 |
ASM_REWRITE_TAC[] THEN |
|
495 |
MATCH_MP_TAC(REAL_ARITH `&0 < x ==> ~(&1 + x = &0)`) THEN |
|
496 |
ASM_SIMP_TAC[REAL_LT_DIV; NORM_POS_LT; VECTOR_SUB_EQ]]]);; |
|
497 |
||
498 |
let PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove |
|
499 |
(`!s:real^N->bool a b:real^M. |
|
500 |
2 <= dimindex(:N) /\ s homeomorphic interval[a,b] |
|
501 |
==> path_connected((:real^N) DIFF s)`, |
|
502 |
REPEAT STRIP_TAC THEN REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN |
|
503 |
FIRST_ASSUM(MP_TAC o MATCH_MP |
|
504 |
UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN |
|
505 |
ASM_REWRITE_TAC[SET_RULE `~(x IN s) <=> x IN (UNIV DIFF s)`] THEN |
|
506 |
ABBREV_TAC `t = (:real^N) DIFF s` THEN |
|
507 |
DISCH_TAC THEN MAP_EVERY X_GEN_TAC [`x:real^N`; `y:real^N`] THEN |
|
508 |
STRIP_TAC THEN FIRST_ASSUM(MP_TAC o MATCH_MP HOMEOMORPHIC_COMPACTNESS) THEN |
|
509 |
REWRITE_TAC[COMPACT_INTERVAL] THEN |
|
510 |
DISCH_THEN(MP_TAC o MATCH_MP COMPACT_IMP_BOUNDED) THEN |
|
511 |
REWRITE_TAC[BOUNDED_POS; LEFT_IMP_EXISTS_THM] THEN |
|
512 |
X_GEN_TAC `B:real` THEN STRIP_TAC THEN |
|
513 |
SUBGOAL_THEN `(?u:real^N. u IN path_component t x /\ B < norm(u)) /\ |
|
514 |
(?v:real^N. v IN path_component t y /\ B < norm(v))` |
|
515 |
STRIP_ASSUME_TAC THENL |
|
516 |
[ASM_MESON_TAC[BOUNDED_POS; REAL_NOT_LE]; ALL_TAC] THEN |
|
517 |
MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `u:real^N` THEN |
|
518 |
CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN |
|
519 |
MATCH_MP_TAC PATH_COMPONENT_SYM THEN |
|
520 |
MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `v:real^N` THEN |
|
521 |
CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN |
|
522 |
MATCH_MP_TAC PATH_COMPONENT_OF_SUBSET THEN |
|
523 |
EXISTS_TAC `(:real^N) DIFF cball(vec 0,B)` THEN CONJ_TAC THENL |
|
524 |
[EXPAND_TAC "t" THEN MATCH_MP_TAC(SET_RULE |
|
525 |
`s SUBSET t ==> (u DIFF t) SUBSET (u DIFF s)`) THEN |
|
526 |
ASM_REWRITE_TAC[SUBSET; IN_CBALL_0]; |
|
527 |
MP_TAC(ISPEC `cball(vec 0:real^N,B)` |
|
528 |
PATH_CONNECTED_COMPLEMENT_BOUNDED_CONVEX) THEN |
|
529 |
ASM_REWRITE_TAC[BOUNDED_CBALL; CONVEX_CBALL] THEN |
|
530 |
REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN |
|
531 |
DISCH_THEN MATCH_MP_TAC THEN |
|
532 |
ASM_REWRITE_TAC[IN_DIFF; IN_UNIV; IN_CBALL_0; REAL_NOT_LE]]);; |
|
533 |
||
534 |
(* ------------------------------------------------------------------------- *) |
|
535 |
(* In particular, apply all these to the special case of an arc. *) |
|
536 |
(* ------------------------------------------------------------------------- *) |
|
537 |
||
538 |
let RETRACTION_ARC = prove |
|
539 |
(`!p. arc p |
|
540 |
==> ?f. f continuous_on (:real^N) /\ |
|
541 |
IMAGE f (:real^N) SUBSET path_image p /\ |
|
542 |
(!x. x IN path_image p ==> f x = x)`, |
|
543 |
REWRITE_TAC[arc; path; path_image] THEN REPEAT STRIP_TAC THEN |
|
544 |
MATCH_MP_TAC RETRACTION_INJECTIVE_IMAGE_INTERVAL THEN |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36593
diff
changeset
|
545 |
ASM_REWRITE_TAC[INTERVAL_EQ_EMPTY_CART_1; DROP_VEC; REAL_NOT_LT; REAL_POS]);; |
36432 | 546 |
|
547 |
let PATH_CONNECTED_ARC_COMPLEMENT = prove |
|
548 |
(`!p. 2 <= dimindex(:N) /\ arc p |
|
549 |
==> path_connected((:real^N) DIFF path_image p)`, |
|
550 |
REWRITE_TAC[arc; path] THEN REPEAT STRIP_TAC THEN SIMP_TAC[path_image] THEN |
|
551 |
MP_TAC(ISPECL [`path_image p:real^N->bool`; `vec 0:real^1`; `vec 1:real^1`] |
|
552 |
PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN |
|
553 |
ASM_REWRITE_TAC[path_image] THEN DISCH_THEN MATCH_MP_TAC THEN |
|
554 |
ONCE_REWRITE_TAC[HOMEOMORPHIC_SYM] THEN |
|
555 |
MATCH_MP_TAC HOMEOMORPHIC_COMPACT THEN |
|
556 |
EXISTS_TAC `p:real^1->real^N` THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]);; |
|
557 |
||
558 |
let CONNECTED_ARC_COMPLEMENT = prove |
|
559 |
(`!p. 2 <= dimindex(:N) /\ arc p |
|
560 |
==> connected((:real^N) DIFF path_image p)`, |
|
561 |
SIMP_TAC[PATH_CONNECTED_ARC_COMPLEMENT; PATH_CONNECTED_IMP_CONNECTED]);; *) |
|
562 |
||
563 |
end |