| author | wenzelm |
| Fri, 11 Apr 2014 23:26:31 +0200 | |
| changeset 56547 | e9bb73d7b6cf |
| parent 56371 | fb9ae0727548 |
| child 56571 | f4635657d66f |
| permissions | -rw-r--r-- |
| 53572 | 1 |
(* Author: John Harrison |
2 |
Author: Robert Himmelmann, TU Muenchen (translation from HOL light) |
|
3 |
*) |
|
| 36432 | 4 |
|
| 53572 | 5 |
header {* Fashoda meet theorem *}
|
| 36432 | 6 |
|
7 |
theory Fashoda |
|
|
37674
f86de9c00c47
convert theorem path_connected_sphere to euclidean_space class
huffman
parents:
37489
diff
changeset
|
8 |
imports Brouwer_Fixpoint Path_Connected Cartesian_Euclidean_Space |
| 36432 | 9 |
begin |
10 |
||
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44647
diff
changeset
|
11 |
(* move *) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44647
diff
changeset
|
12 |
|
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44647
diff
changeset
|
13 |
lemma cart_eq_inner_axis: "a $ i = a \<bullet> axis i 1" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44647
diff
changeset
|
14 |
by (simp add: inner_axis) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44647
diff
changeset
|
15 |
|
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44647
diff
changeset
|
16 |
lemma axis_in_Basis: "a \<in> Basis \<Longrightarrow> axis i a \<in> Basis" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44647
diff
changeset
|
17 |
by (auto simp add: Basis_vec_def axis_eq_axis) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44647
diff
changeset
|
18 |
|
|
56273
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
19 |
lemma divide_nonneg_nonneg: |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
20 |
fixes a b :: "'a :: {linordered_field, field_inverse_zero}"
|
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
21 |
shows "a \<ge> 0 \<Longrightarrow> b \<ge> 0 \<Longrightarrow> 0 \<le> a / b" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
22 |
by (cases "b = 0") (auto intro!: divide_nonneg_pos) |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
23 |
|
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
24 |
subsection {*Bijections between intervals. *}
|
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
25 |
|
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
26 |
definition interval_bij :: "'a \<times> 'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<Rightarrow> 'a::euclidean_space" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
27 |
where "interval_bij = |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
28 |
(\<lambda>(a, b) (u, v) x. (\<Sum>i\<in>Basis. (u\<bullet>i + (x\<bullet>i - a\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (v\<bullet>i - u\<bullet>i)) *\<^sub>R i))" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
29 |
|
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
30 |
lemma interval_bij_affine: |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
31 |
"interval_bij (a,b) (u,v) = (\<lambda>x. (\<Sum>i\<in>Basis. ((v\<bullet>i - u\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (x\<bullet>i)) *\<^sub>R i) + |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
32 |
(\<Sum>i\<in>Basis. (u\<bullet>i - (v\<bullet>i - u\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (a\<bullet>i)) *\<^sub>R i))" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
33 |
by (auto simp: setsum_addf[symmetric] scaleR_add_left[symmetric] interval_bij_def fun_eq_iff |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
34 |
field_simps inner_simps add_divide_distrib[symmetric] intro!: setsum_cong) |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
35 |
|
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
36 |
lemma continuous_interval_bij: |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
37 |
fixes a b :: "'a::euclidean_space" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
38 |
shows "continuous (at x) (interval_bij (a, b) (u, v))" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
39 |
by (auto simp add: divide_inverse interval_bij_def intro!: continuous_setsum continuous_intros) |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
40 |
|
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
41 |
lemma continuous_on_interval_bij: "continuous_on s (interval_bij (a, b) (u, v))" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
42 |
apply(rule continuous_at_imp_continuous_on) |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
43 |
apply (rule, rule continuous_interval_bij) |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
44 |
done |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
45 |
|
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
46 |
lemma in_interval_interval_bij: |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
47 |
fixes a b u v x :: "'a::euclidean_space" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
48 |
assumes "x \<in> cbox a b" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
49 |
and "cbox u v \<noteq> {}"
|
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
50 |
shows "interval_bij (a, b) (u, v) x \<in> cbox u v" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
51 |
apply (simp only: interval_bij_def split_conv mem_box inner_setsum_left_Basis cong: ball_cong) |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
52 |
apply safe |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
53 |
proof - |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
54 |
fix i :: 'a |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
55 |
assume i: "i \<in> Basis" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
56 |
have "cbox a b \<noteq> {}"
|
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
57 |
using assms by auto |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
58 |
with i have *: "a\<bullet>i \<le> b\<bullet>i" "u\<bullet>i \<le> v\<bullet>i" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
59 |
using assms(2) by (auto simp add: box_eq_empty) |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
60 |
have x: "a\<bullet>i\<le>x\<bullet>i" "x\<bullet>i\<le>b\<bullet>i" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
61 |
using assms(1)[unfolded mem_box] using i by auto |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
62 |
have "0 \<le> (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i)" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
63 |
using * x by (auto intro!: mult_nonneg_nonneg divide_nonneg_nonneg) |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
64 |
then show "u \<bullet> i \<le> u \<bullet> i + (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i)" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
65 |
using * by auto |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
66 |
have "((x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (v \<bullet> i - u \<bullet> i) \<le> 1 * (v \<bullet> i - u \<bullet> i)" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
67 |
apply (rule mult_right_mono) |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
68 |
unfolding divide_le_eq_1 |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
69 |
using * x |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
70 |
apply auto |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
71 |
done |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
72 |
then show "u \<bullet> i + (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i) \<le> v \<bullet> i" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
73 |
using * by auto |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
74 |
qed |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
75 |
|
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
76 |
lemma interval_bij_bij: |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
77 |
"\<forall>(i::'a::euclidean_space)\<in>Basis. a\<bullet>i < b\<bullet>i \<and> u\<bullet>i < v\<bullet>i \<Longrightarrow> |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
78 |
interval_bij (a, b) (u, v) (interval_bij (u, v) (a, b) x) = x" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
79 |
by (auto simp: interval_bij_def euclidean_eq_iff[where 'a='a]) |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
80 |
|
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
81 |
lemma interval_bij_bij_cart: fixes x::"real^'n" assumes "\<forall>i. a$i < b$i \<and> u$i < v$i" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
82 |
shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x" |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
83 |
using assms by (intro interval_bij_bij) (auto simp: Basis_vec_def inner_axis) |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
84 |
|
| 53572 | 85 |
|
86 |
subsection {* Fashoda meet theorem *}
|
|
| 36432 | 87 |
|
| 53572 | 88 |
lemma infnorm_2: |
89 |
fixes x :: "real^2" |
|
90 |
shows "infnorm x = max (abs (x$1)) (abs (x$2))" |
|
91 |
unfolding infnorm_cart UNIV_2 by (rule cSup_eq) auto |
|
| 36432 | 92 |
|
| 53572 | 93 |
lemma infnorm_eq_1_2: |
94 |
fixes x :: "real^2" |
|
95 |
shows "infnorm x = 1 \<longleftrightarrow> |
|
96 |
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)" |
|
| 36432 | 97 |
unfolding infnorm_2 by auto |
98 |
||
| 53572 | 99 |
lemma infnorm_eq_1_imp: |
100 |
fixes x :: "real^2" |
|
101 |
assumes "infnorm x = 1" |
|
102 |
shows "abs (x$1) \<le> 1" and "abs (x$2) \<le> 1" |
|
| 36432 | 103 |
using assms unfolding infnorm_eq_1_2 by auto |
104 |
||
| 53572 | 105 |
lemma fashoda_unit: |
106 |
fixes f g :: "real \<Rightarrow> real^2" |
|
| 56188 | 107 |
assumes "f ` {-1 .. 1} \<subseteq> cbox (-1) 1"
|
108 |
and "g ` {-1 .. 1} \<subseteq> cbox (-1) 1"
|
|
109 |
and "continuous_on {-1 .. 1} f"
|
|
110 |
and "continuous_on {-1 .. 1} g"
|
|
| 53572 | 111 |
and "f (- 1)$1 = - 1" |
112 |
and "f 1$1 = 1" "g (- 1) $2 = -1" |
|
113 |
and "g 1 $2 = 1" |
|
| 56188 | 114 |
shows "\<exists>s\<in>{-1 .. 1}. \<exists>t\<in>{-1 .. 1}. f s = g t"
|
| 53572 | 115 |
proof (rule ccontr) |
116 |
assume "\<not> ?thesis" |
|
117 |
note as = this[unfolded bex_simps,rule_format] |
|
| 36432 | 118 |
def sqprojection \<equiv> "\<lambda>z::real^2. (inverse (infnorm z)) *\<^sub>R z" |
| 53572 | 119 |
def negatex \<equiv> "\<lambda>x::real^2. (vector [-(x$1), x$2])::real^2" |
120 |
have lem1: "\<forall>z::real^2. infnorm (negatex z) = infnorm z" |
|
| 36432 | 121 |
unfolding negatex_def infnorm_2 vector_2 by auto |
| 53572 | 122 |
have lem2: "\<forall>z. z \<noteq> 0 \<longrightarrow> infnorm (sqprojection z) = 1" |
123 |
unfolding sqprojection_def |
|
124 |
unfolding infnorm_mul[unfolded scalar_mult_eq_scaleR] |
|
125 |
unfolding abs_inverse real_abs_infnorm |
|
| 53628 | 126 |
apply (subst infnorm_eq_0[symmetric]) |
| 53572 | 127 |
apply auto |
128 |
done |
|
129 |
let ?F = "\<lambda>w::real^2. (f \<circ> (\<lambda>x. x$1)) w - (g \<circ> (\<lambda>x. x$2)) w" |
|
| 56188 | 130 |
have *: "\<And>i. (\<lambda>x::real^2. x $ i) ` cbox (- 1) 1 = {-1 .. 1}"
|
| 53572 | 131 |
apply (rule set_eqI) |
| 56188 | 132 |
unfolding image_iff Bex_def mem_interval_cart interval_cbox_cart |
| 53572 | 133 |
apply rule |
134 |
defer |
|
135 |
apply (rule_tac x="vec x" in exI) |
|
136 |
apply auto |
|
137 |
done |
|
138 |
{
|
|
139 |
fix x |
|
| 56188 | 140 |
assume "x \<in> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w) ` (cbox (- 1) (1::real^2))" |
| 55675 | 141 |
then obtain w :: "real^2" where w: |
| 56188 | 142 |
"w \<in> cbox (- 1) 1" |
| 55675 | 143 |
"x = (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w" |
144 |
unfolding image_iff .. |
|
| 53572 | 145 |
then have "x \<noteq> 0" |
146 |
using as[of "w$1" "w$2"] |
|
| 56188 | 147 |
unfolding mem_interval_cart atLeastAtMost_iff |
| 53572 | 148 |
by auto |
149 |
} note x0 = this |
|
150 |
have 21: "\<And>i::2. i \<noteq> 1 \<Longrightarrow> i = 2" |
|
151 |
using UNIV_2 by auto |
|
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
53628
diff
changeset
|
152 |
have 1: "box (- 1) (1::real^2) \<noteq> {}"
|
| 53572 | 153 |
unfolding interval_eq_empty_cart by auto |
| 56188 | 154 |
have 2: "continuous_on (cbox -1 1) (negatex \<circ> sqprojection \<circ> ?F)" |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56273
diff
changeset
|
155 |
apply (intro continuous_intros continuous_on_component) |
| 53572 | 156 |
unfolding * |
157 |
apply (rule assms)+ |
|
158 |
apply (subst sqprojection_def) |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56273
diff
changeset
|
159 |
apply (intro continuous_intros) |
| 53572 | 160 |
apply (simp add: infnorm_eq_0 x0) |
161 |
apply (rule linear_continuous_on) |
|
162 |
proof - |
|
163 |
show "bounded_linear negatex" |
|
164 |
apply (rule bounded_linearI') |
|
165 |
unfolding vec_eq_iff |
|
166 |
proof (rule_tac[!] allI) |
|
167 |
fix i :: 2 |
|
168 |
fix x y :: "real^2" |
|
169 |
fix c :: real |
|
170 |
show "negatex (x + y) $ i = |
|
171 |
(negatex x + negatex y) $ i" "negatex (c *\<^sub>R x) $ i = (c *\<^sub>R negatex x) $ i" |
|
172 |
apply - |
|
173 |
apply (case_tac[!] "i\<noteq>1") |
|
174 |
prefer 3 |
|
175 |
apply (drule_tac[1-2] 21) |
|
176 |
unfolding negatex_def |
|
177 |
apply (auto simp add:vector_2) |
|
178 |
done |
|
179 |
qed |
|
|
44647
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44531
diff
changeset
|
180 |
qed |
| 56188 | 181 |
have 3: "(negatex \<circ> sqprojection \<circ> ?F) ` cbox (-1) 1 \<subseteq> cbox (-1) 1" |
| 53572 | 182 |
unfolding subset_eq |
183 |
apply rule |
|
184 |
proof - |
|
185 |
case goal1 |
|
| 55675 | 186 |
then obtain y :: "real^2" where y: |
| 56188 | 187 |
"y \<in> cbox -1 1" |
| 55675 | 188 |
"x = (negatex \<circ> sqprojection \<circ> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w)) y" |
189 |
unfolding image_iff .. |
|
| 53572 | 190 |
have "?F y \<noteq> 0" |
191 |
apply (rule x0) |
|
192 |
using y(1) |
|
193 |
apply auto |
|
194 |
done |
|
195 |
then have *: "infnorm (sqprojection (?F y)) = 1" |
|
| 53628 | 196 |
unfolding y o_def |
197 |
by - (rule lem2[rule_format]) |
|
| 53572 | 198 |
have "infnorm x = 1" |
| 53628 | 199 |
unfolding *[symmetric] y o_def |
200 |
by (rule lem1[rule_format]) |
|
| 56188 | 201 |
then show "x \<in> cbox (-1) 1" |
202 |
unfolding mem_interval_cart interval_cbox_cart infnorm_2 |
|
| 53572 | 203 |
apply - |
204 |
apply rule |
|
205 |
proof - |
|
206 |
case goal1 |
|
207 |
then show ?case |
|
208 |
apply (cases "i = 1") |
|
209 |
defer |
|
210 |
apply (drule 21) |
|
211 |
apply auto |
|
212 |
done |
|
213 |
qed |
|
214 |
qed |
|
| 55675 | 215 |
obtain x :: "real^2" where x: |
| 56188 | 216 |
"x \<in> cbox -1 1" |
| 55675 | 217 |
"(negatex \<circ> sqprojection \<circ> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w)) x = x" |
| 56188 | 218 |
apply (rule brouwer_weak[of "cbox -1 (1::real^2)" "negatex \<circ> sqprojection \<circ> ?F"]) |
219 |
apply (rule compact_cbox convex_box)+ |
|
220 |
unfolding interior_cbox |
|
| 53572 | 221 |
apply (rule 1 2 3)+ |
| 55675 | 222 |
apply blast |
| 53572 | 223 |
done |
224 |
have "?F x \<noteq> 0" |
|
225 |
apply (rule x0) |
|
226 |
using x(1) |
|
227 |
apply auto |
|
228 |
done |
|
229 |
then have *: "infnorm (sqprojection (?F x)) = 1" |
|
| 53628 | 230 |
unfolding o_def |
231 |
by (rule lem2[rule_format]) |
|
| 53572 | 232 |
have nx: "infnorm x = 1" |
| 53628 | 233 |
apply (subst x(2)[symmetric]) |
234 |
unfolding *[symmetric] o_def |
|
| 53572 | 235 |
apply (rule lem1[rule_format]) |
236 |
done |
|
237 |
have "\<forall>x i. x \<noteq> 0 \<longrightarrow> (0 < (sqprojection x)$i \<longleftrightarrow> 0 < x$i)" |
|
238 |
and "\<forall>x i. x \<noteq> 0 \<longrightarrow> ((sqprojection x)$i < 0 \<longleftrightarrow> x$i < 0)" |
|
239 |
apply - |
|
240 |
apply (rule_tac[!] allI impI)+ |
|
241 |
proof - |
|
242 |
fix x :: "real^2" |
|
243 |
fix i :: 2 |
|
244 |
assume x: "x \<noteq> 0" |
|
245 |
have "inverse (infnorm x) > 0" |
|
| 53628 | 246 |
using x[unfolded infnorm_pos_lt[symmetric]] by auto |
| 53572 | 247 |
then show "(0 < sqprojection x $ i) = (0 < x $ i)" |
248 |
and "(sqprojection x $ i < 0) = (x $ i < 0)" |
|
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44136
diff
changeset
|
249 |
unfolding sqprojection_def vector_component_simps vector_scaleR_component real_scaleR_def |
| 53572 | 250 |
unfolding zero_less_mult_iff mult_less_0_iff |
251 |
by (auto simp add: field_simps) |
|
252 |
qed |
|
| 36432 | 253 |
note lem3 = this[rule_format] |
| 53572 | 254 |
have x1: "x $ 1 \<in> {- 1..1::real}" "x $ 2 \<in> {- 1..1::real}"
|
255 |
using x(1) unfolding mem_interval_cart by auto |
|
256 |
then have nz: "f (x $ 1) - g (x $ 2) \<noteq> 0" |
|
257 |
unfolding right_minus_eq |
|
258 |
apply - |
|
259 |
apply (rule as) |
|
260 |
apply auto |
|
261 |
done |
|
262 |
have "x $ 1 = -1 \<or> x $ 1 = 1 \<or> x $ 2 = -1 \<or> x $ 2 = 1" |
|
263 |
using nx unfolding infnorm_eq_1_2 by auto |
|
264 |
then show False |
|
265 |
proof - |
|
266 |
fix P Q R S |
|
267 |
presume "P \<or> Q \<or> R \<or> S" |
|
268 |
and "P \<Longrightarrow> False" |
|
269 |
and "Q \<Longrightarrow> False" |
|
270 |
and "R \<Longrightarrow> False" |
|
271 |
and "S \<Longrightarrow> False" |
|
272 |
then show False by auto |
|
273 |
next |
|
274 |
assume as: "x$1 = 1" |
|
275 |
then have *: "f (x $ 1) $ 1 = 1" |
|
276 |
using assms(6) by auto |
|
| 36432 | 277 |
have "sqprojection (f (x$1) - g (x$2)) $ 1 < 0" |
|
44136
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents:
41958
diff
changeset
|
278 |
using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]] |
| 53572 | 279 |
unfolding as negatex_def vector_2 |
280 |
by auto |
|
281 |
moreover |
|
| 56188 | 282 |
from x1 have "g (x $ 2) \<in> cbox (-1) 1" |
| 53572 | 283 |
apply - |
284 |
apply (rule assms(2)[unfolded subset_eq,rule_format]) |
|
285 |
apply auto |
|
286 |
done |
|
287 |
ultimately show False |
|
288 |
unfolding lem3[OF nz] vector_component_simps * mem_interval_cart |
|
289 |
apply (erule_tac x=1 in allE) |
|
290 |
apply auto |
|
291 |
done |
|
292 |
next |
|
293 |
assume as: "x$1 = -1" |
|
294 |
then have *: "f (x $ 1) $ 1 = - 1" |
|
295 |
using assms(5) by auto |
|
| 36432 | 296 |
have "sqprojection (f (x$1) - g (x$2)) $ 1 > 0" |
|
44136
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents:
41958
diff
changeset
|
297 |
using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]] |
| 53572 | 298 |
unfolding as negatex_def vector_2 |
299 |
by auto |
|
300 |
moreover |
|
| 56188 | 301 |
from x1 have "g (x $ 2) \<in> cbox (-1) 1" |
| 53572 | 302 |
apply - |
303 |
apply (rule assms(2)[unfolded subset_eq,rule_format]) |
|
304 |
apply auto |
|
305 |
done |
|
306 |
ultimately show False |
|
307 |
unfolding lem3[OF nz] vector_component_simps * mem_interval_cart |
|
308 |
apply (erule_tac x=1 in allE) |
|
309 |
apply auto |
|
310 |
done |
|
311 |
next |
|
312 |
assume as: "x$2 = 1" |
|
313 |
then have *: "g (x $ 2) $ 2 = 1" |
|
314 |
using assms(8) by auto |
|
| 36432 | 315 |
have "sqprojection (f (x$1) - g (x$2)) $ 2 > 0" |
|
44136
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents:
41958
diff
changeset
|
316 |
using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]] |
| 53572 | 317 |
unfolding as negatex_def vector_2 |
318 |
by auto |
|
319 |
moreover |
|
| 56188 | 320 |
from x1 have "f (x $ 1) \<in> cbox (-1) 1" |
| 53572 | 321 |
apply - |
322 |
apply (rule assms(1)[unfolded subset_eq,rule_format]) |
|
323 |
apply auto |
|
324 |
done |
|
325 |
ultimately show False |
|
326 |
unfolding lem3[OF nz] vector_component_simps * mem_interval_cart |
|
327 |
apply (erule_tac x=2 in allE) |
|
328 |
apply auto |
|
329 |
done |
|
330 |
next |
|
331 |
assume as: "x$2 = -1" |
|
332 |
then have *: "g (x $ 2) $ 2 = - 1" |
|
333 |
using assms(7) by auto |
|
| 36432 | 334 |
have "sqprojection (f (x$1) - g (x$2)) $ 2 < 0" |
|
44136
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
huffman
parents:
41958
diff
changeset
|
335 |
using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]] |
| 53572 | 336 |
unfolding as negatex_def vector_2 |
337 |
by auto |
|
338 |
moreover |
|
| 56188 | 339 |
from x1 have "f (x $ 1) \<in> cbox (-1) 1" |
| 53572 | 340 |
apply - |
341 |
apply (rule assms(1)[unfolded subset_eq,rule_format]) |
|
342 |
apply auto |
|
343 |
done |
|
344 |
ultimately show False |
|
345 |
unfolding lem3[OF nz] vector_component_simps * mem_interval_cart |
|
346 |
apply (erule_tac x=2 in allE) |
|
347 |
apply auto |
|
348 |
done |
|
349 |
qed auto |
|
350 |
qed |
|
| 36432 | 351 |
|
| 53572 | 352 |
lemma fashoda_unit_path: |
353 |
fixes f g :: "real \<Rightarrow> real^2" |
|
354 |
assumes "path f" |
|
355 |
and "path g" |
|
| 56188 | 356 |
and "path_image f \<subseteq> cbox (-1) 1" |
357 |
and "path_image g \<subseteq> cbox (-1) 1" |
|
| 53572 | 358 |
and "(pathstart f)$1 = -1" |
359 |
and "(pathfinish f)$1 = 1" |
|
360 |
and "(pathstart g)$2 = -1" |
|
361 |
and "(pathfinish g)$2 = 1" |
|
362 |
obtains z where "z \<in> path_image f" and "z \<in> path_image g" |
|
363 |
proof - |
|
| 36432 | 364 |
note assms=assms[unfolded path_def pathstart_def pathfinish_def path_image_def] |
365 |
def iscale \<equiv> "\<lambda>z::real. inverse 2 *\<^sub>R (z + 1)" |
|
| 53572 | 366 |
have isc: "iscale ` {- 1..1} \<subseteq> {0..1}"
|
367 |
unfolding iscale_def by auto |
|
368 |
have "\<exists>s\<in>{- 1..1}. \<exists>t\<in>{- 1..1}. (f \<circ> iscale) s = (g \<circ> iscale) t"
|
|
369 |
proof (rule fashoda_unit) |
|
| 56188 | 370 |
show "(f \<circ> iscale) ` {- 1..1} \<subseteq> cbox -1 1" "(g \<circ> iscale) ` {- 1..1} \<subseteq> cbox -1 1"
|
|
56154
f0a927235162
more complete set of lemmas wrt. image and composition
haftmann
parents:
55675
diff
changeset
|
371 |
using isc and assms(3-4) by (auto simp add: image_comp [symmetric]) |
| 53572 | 372 |
have *: "continuous_on {- 1..1} iscale"
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56273
diff
changeset
|
373 |
unfolding iscale_def by (rule continuous_intros)+ |
| 36432 | 374 |
show "continuous_on {- 1..1} (f \<circ> iscale)" "continuous_on {- 1..1} (g \<circ> iscale)"
|
| 53572 | 375 |
apply - |
376 |
apply (rule_tac[!] continuous_on_compose[OF *]) |
|
377 |
apply (rule_tac[!] continuous_on_subset[OF _ isc]) |
|
378 |
apply (rule assms)+ |
|
379 |
done |
|
380 |
have *: "(1 / 2) *\<^sub>R (1 + (1::real^1)) = 1" |
|
381 |
unfolding vec_eq_iff by auto |
|
382 |
show "(f \<circ> iscale) (- 1) $ 1 = - 1" |
|
383 |
and "(f \<circ> iscale) 1 $ 1 = 1" |
|
384 |
and "(g \<circ> iscale) (- 1) $ 2 = -1" |
|
385 |
and "(g \<circ> iscale) 1 $ 2 = 1" |
|
386 |
unfolding o_def iscale_def |
|
387 |
using assms |
|
388 |
by (auto simp add: *) |
|
389 |
qed |
|
| 55675 | 390 |
then obtain s t where st: |
391 |
"s \<in> {- 1..1}"
|
|
392 |
"t \<in> {- 1..1}"
|
|
393 |
"(f \<circ> iscale) s = (g \<circ> iscale) t" |
|
| 56188 | 394 |
by auto |
| 53572 | 395 |
show thesis |
| 53628 | 396 |
apply (rule_tac z = "f (iscale s)" in that) |
| 55675 | 397 |
using st |
| 53572 | 398 |
unfolding o_def path_image_def image_iff |
399 |
apply - |
|
400 |
apply (rule_tac x="iscale s" in bexI) |
|
401 |
prefer 3 |
|
402 |
apply (rule_tac x="iscale t" in bexI) |
|
403 |
using isc[unfolded subset_eq, rule_format] |
|
404 |
apply auto |
|
405 |
done |
|
406 |
qed |
|
| 36432 | 407 |
|
| 53627 | 408 |
lemma fashoda: |
409 |
fixes b :: "real^2" |
|
410 |
assumes "path f" |
|
411 |
and "path g" |
|
| 56188 | 412 |
and "path_image f \<subseteq> cbox a b" |
413 |
and "path_image g \<subseteq> cbox a b" |
|
| 53627 | 414 |
and "(pathstart f)$1 = a$1" |
415 |
and "(pathfinish f)$1 = b$1" |
|
416 |
and "(pathstart g)$2 = a$2" |
|
417 |
and "(pathfinish g)$2 = b$2" |
|
418 |
obtains z where "z \<in> path_image f" and "z \<in> path_image g" |
|
419 |
proof - |
|
420 |
fix P Q S |
|
421 |
presume "P \<or> Q \<or> S" "P \<Longrightarrow> thesis" and "Q \<Longrightarrow> thesis" and "S \<Longrightarrow> thesis" |
|
422 |
then show thesis |
|
423 |
by auto |
|
424 |
next |
|
| 56188 | 425 |
have "cbox a b \<noteq> {}"
|
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
53628
diff
changeset
|
426 |
using assms(3) using path_image_nonempty[of f] by auto |
| 53627 | 427 |
then have "a \<le> b" |
428 |
unfolding interval_eq_empty_cart less_eq_vec_def by (auto simp add: not_less) |
|
429 |
then show "a$1 = b$1 \<or> a$2 = b$2 \<or> (a$1 < b$1 \<and> a$2 < b$2)" |
|
430 |
unfolding less_eq_vec_def forall_2 by auto |
|
431 |
next |
|
432 |
assume as: "a$1 = b$1" |
|
433 |
have "\<exists>z\<in>path_image g. z$2 = (pathstart f)$2" |
|
434 |
apply (rule connected_ivt_component_cart) |
|
435 |
apply (rule connected_path_image assms)+ |
|
436 |
apply (rule pathstart_in_path_image) |
|
437 |
apply (rule pathfinish_in_path_image) |
|
| 36432 | 438 |
unfolding assms using assms(3)[unfolded path_image_def subset_eq,rule_format,of "f 0"] |
| 53627 | 439 |
unfolding pathstart_def |
| 56188 | 440 |
apply (auto simp add: less_eq_vec_def mem_interval_cart) |
| 53627 | 441 |
done |
| 55675 | 442 |
then obtain z :: "real^2" where z: "z \<in> path_image g" "z $ 2 = pathstart f $ 2" .. |
| 56188 | 443 |
have "z \<in> cbox a b" |
| 53627 | 444 |
using z(1) assms(4) |
445 |
unfolding path_image_def |
|
| 56188 | 446 |
by blast |
| 53627 | 447 |
then have "z = f 0" |
448 |
unfolding vec_eq_iff forall_2 |
|
449 |
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
|
450 |
using assms(3)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "f 0" 1] |
| 53627 | 451 |
unfolding mem_interval_cart |
452 |
apply (erule_tac x=1 in allE) |
|
453 |
using as |
|
454 |
apply auto |
|
455 |
done |
|
456 |
then show thesis |
|
457 |
apply - |
|
458 |
apply (rule that[OF _ z(1)]) |
|
459 |
unfolding path_image_def |
|
460 |
apply auto |
|
461 |
done |
|
462 |
next |
|
463 |
assume as: "a$2 = b$2" |
|
464 |
have "\<exists>z\<in>path_image f. z$1 = (pathstart g)$1" |
|
465 |
apply (rule connected_ivt_component_cart) |
|
466 |
apply (rule connected_path_image assms)+ |
|
467 |
apply (rule pathstart_in_path_image) |
|
468 |
apply (rule pathfinish_in_path_image) |
|
469 |
unfolding assms |
|
470 |
using assms(4)[unfolded path_image_def subset_eq,rule_format,of "g 0"] |
|
471 |
unfolding pathstart_def |
|
| 56188 | 472 |
apply (auto simp add: less_eq_vec_def mem_interval_cart) |
| 53627 | 473 |
done |
| 55675 | 474 |
then obtain z where z: "z \<in> path_image f" "z $ 1 = pathstart g $ 1" .. |
| 56188 | 475 |
have "z \<in> cbox a b" |
| 53627 | 476 |
using z(1) assms(3) |
477 |
unfolding path_image_def |
|
| 56188 | 478 |
by blast |
| 53627 | 479 |
then have "z = g 0" |
480 |
unfolding vec_eq_iff forall_2 |
|
481 |
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
|
482 |
using assms(4)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "g 0" 2] |
| 53627 | 483 |
unfolding mem_interval_cart |
484 |
apply (erule_tac x=2 in allE) |
|
485 |
using as |
|
486 |
apply auto |
|
487 |
done |
|
488 |
then show thesis |
|
489 |
apply - |
|
490 |
apply (rule that[OF z(1)]) |
|
491 |
unfolding path_image_def |
|
492 |
apply auto |
|
493 |
done |
|
494 |
next |
|
495 |
assume as: "a $ 1 < b $ 1 \<and> a $ 2 < b $ 2" |
|
| 56188 | 496 |
have int_nem: "cbox (-1) (1::real^2) \<noteq> {}"
|
| 53627 | 497 |
unfolding interval_eq_empty_cart by auto |
| 55675 | 498 |
obtain z :: "real^2" where z: |
499 |
"z \<in> (interval_bij (a, b) (- 1, 1) \<circ> f) ` {0..1}"
|
|
500 |
"z \<in> (interval_bij (a, b) (- 1, 1) \<circ> g) ` {0..1}"
|
|
| 53627 | 501 |
apply (rule fashoda_unit_path[of "interval_bij (a,b) (- 1,1) \<circ> f" "interval_bij (a,b) (- 1,1) \<circ> g"]) |
| 36432 | 502 |
unfolding path_def path_image_def pathstart_def pathfinish_def |
| 53627 | 503 |
apply (rule_tac[1-2] continuous_on_compose) |
504 |
apply (rule assms[unfolded path_def] continuous_on_interval_bij)+ |
|
505 |
unfolding subset_eq |
|
506 |
apply(rule_tac[1-2] ballI) |
|
507 |
proof - |
|
508 |
fix x |
|
509 |
assume "x \<in> (interval_bij (a, b) (- 1, 1) \<circ> f) ` {0..1}"
|
|
| 55675 | 510 |
then obtain y where y: |
511 |
"y \<in> {0..1}"
|
|
512 |
"x = (interval_bij (a, b) (- 1, 1) \<circ> f) y" |
|
513 |
unfolding image_iff .. |
|
| 56188 | 514 |
show "x \<in> cbox -1 1" |
| 53627 | 515 |
unfolding y o_def |
516 |
apply (rule in_interval_interval_bij) |
|
517 |
using y(1) |
|
518 |
using assms(3)[unfolded path_image_def subset_eq] int_nem |
|
519 |
apply auto |
|
520 |
done |
|
521 |
next |
|
522 |
fix x |
|
523 |
assume "x \<in> (interval_bij (a, b) (- 1, 1) \<circ> g) ` {0..1}"
|
|
| 55675 | 524 |
then obtain y where y: |
525 |
"y \<in> {0..1}"
|
|
526 |
"x = (interval_bij (a, b) (- 1, 1) \<circ> g) y" |
|
527 |
unfolding image_iff .. |
|
| 56188 | 528 |
show "x \<in> cbox -1 1" |
| 53627 | 529 |
unfolding y o_def |
530 |
apply (rule in_interval_interval_bij) |
|
531 |
using y(1) |
|
532 |
using assms(4)[unfolded path_image_def subset_eq] int_nem |
|
533 |
apply auto |
|
534 |
done |
|
535 |
next |
|
536 |
show "(interval_bij (a, b) (- 1, 1) \<circ> f) 0 $ 1 = -1" |
|
537 |
and "(interval_bij (a, b) (- 1, 1) \<circ> f) 1 $ 1 = 1" |
|
538 |
and "(interval_bij (a, b) (- 1, 1) \<circ> g) 0 $ 2 = -1" |
|
539 |
and "(interval_bij (a, b) (- 1, 1) \<circ> g) 1 $ 2 = 1" |
|
| 56188 | 540 |
using assms as |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44647
diff
changeset
|
541 |
by (simp_all add: axis_in_Basis cart_eq_inner_axis pathstart_def pathfinish_def interval_bij_def) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44647
diff
changeset
|
542 |
(simp_all add: inner_axis) |
| 53627 | 543 |
qed |
| 55675 | 544 |
from z(1) obtain zf where zf: |
545 |
"zf \<in> {0..1}"
|
|
546 |
"z = (interval_bij (a, b) (- 1, 1) \<circ> f) zf" |
|
547 |
unfolding image_iff .. |
|
548 |
from z(2) obtain zg where zg: |
|
549 |
"zg \<in> {0..1}"
|
|
550 |
"z = (interval_bij (a, b) (- 1, 1) \<circ> g) zg" |
|
551 |
unfolding image_iff .. |
|
| 53627 | 552 |
have *: "\<forall>i. (- 1) $ i < (1::real^2) $ i \<and> a $ i < b $ i" |
553 |
unfolding forall_2 |
|
554 |
using as |
|
555 |
by auto |
|
556 |
show thesis |
|
557 |
apply (rule_tac z="interval_bij (- 1,1) (a,b) z" in that) |
|
558 |
apply (subst zf) |
|
559 |
defer |
|
560 |
apply (subst zg) |
|
561 |
unfolding o_def interval_bij_bij_cart[OF *] path_image_def |
|
562 |
using zf(1) zg(1) |
|
563 |
apply auto |
|
564 |
done |
|
565 |
qed |
|
| 36432 | 566 |
|
| 53627 | 567 |
|
568 |
subsection {* Some slightly ad hoc lemmas I use below *}
|
|
| 36432 | 569 |
|
| 53627 | 570 |
lemma segment_vertical: |
571 |
fixes a :: "real^2" |
|
572 |
assumes "a$1 = b$1" |
|
573 |
shows "x \<in> closed_segment a b \<longleftrightarrow> |
|
574 |
x$1 = a$1 \<and> x$1 = b$1 \<and> (a$2 \<le> x$2 \<and> x$2 \<le> b$2 \<or> b$2 \<le> x$2 \<and> x$2 \<le> a$2)" |
|
575 |
(is "_ = ?R") |
|
576 |
proof - |
|
| 36432 | 577 |
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" |
| 53627 | 578 |
{
|
579 |
presume "?L \<Longrightarrow> ?R" and "?R \<Longrightarrow> ?L" |
|
580 |
then show ?thesis |
|
581 |
unfolding closed_segment_def mem_Collect_eq |
|
| 53628 | 582 |
unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[symmetric] vector_component_simps |
| 53627 | 583 |
by blast |
584 |
} |
|
585 |
{
|
|
586 |
assume ?L |
|
| 55675 | 587 |
then obtain u where u: |
588 |
"x $ 1 = (1 - u) * a $ 1 + u * b $ 1" |
|
589 |
"x $ 2 = (1 - u) * a $ 2 + u * b $ 2" |
|
590 |
"0 \<le> u" |
|
591 |
"u \<le> 1" |
|
592 |
by blast |
|
| 53627 | 593 |
{ fix b a
|
594 |
assume "b + u * a > a + u * b" |
|
595 |
then have "(1 - u) * b > (1 - u) * a" |
|
596 |
by (auto simp add:field_simps) |
|
597 |
then have "b \<ge> a" |
|
598 |
apply (drule_tac mult_less_imp_less_left) |
|
599 |
using u |
|
600 |
apply auto |
|
601 |
done |
|
602 |
then have "u * a \<le> u * b" |
|
603 |
apply - |
|
604 |
apply (rule mult_left_mono[OF _ u(3)]) |
|
605 |
using u(3-4) |
|
606 |
apply (auto simp add: field_simps) |
|
607 |
done |
|
608 |
} note * = this |
|
609 |
{
|
|
610 |
fix a b |
|
611 |
assume "u * b > u * a" |
|
612 |
then have "(1 - u) * a \<le> (1 - u) * b" |
|
613 |
apply - |
|
614 |
apply (rule mult_left_mono) |
|
615 |
apply (drule mult_less_imp_less_left) |
|
616 |
using u |
|
617 |
apply auto |
|
618 |
done |
|
619 |
then have "a + u * b \<le> b + u * a" |
|
620 |
by (auto simp add: field_simps) |
|
621 |
} note ** = this |
|
622 |
then show ?R |
|
623 |
unfolding u assms |
|
624 |
using u |
|
625 |
by (auto simp add:field_simps not_le intro: * **) |
|
626 |
} |
|
627 |
{
|
|
628 |
assume ?R |
|
629 |
then show ?L |
|
630 |
proof (cases "x$2 = b$2") |
|
631 |
case True |
|
632 |
then show ?L |
|
633 |
apply (rule_tac x="(x$2 - a$2) / (b$2 - a$2)" in exI) |
|
634 |
unfolding assms True |
|
635 |
using `?R` |
|
636 |
apply (auto simp add: field_simps) |
|
637 |
done |
|
638 |
next |
|
639 |
case False |
|
640 |
then show ?L |
|
641 |
apply (rule_tac x="1 - (x$2 - b$2) / (a$2 - b$2)" in exI) |
|
642 |
unfolding assms |
|
643 |
using `?R` |
|
644 |
apply (auto simp add: field_simps) |
|
645 |
done |
|
646 |
qed |
|
647 |
} |
|
648 |
qed |
|
| 36432 | 649 |
|
| 53627 | 650 |
lemma segment_horizontal: |
651 |
fixes a :: "real^2" |
|
652 |
assumes "a$2 = b$2" |
|
653 |
shows "x \<in> closed_segment a b \<longleftrightarrow> |
|
654 |
x$2 = a$2 \<and> x$2 = b$2 \<and> (a$1 \<le> x$1 \<and> x$1 \<le> b$1 \<or> b$1 \<le> x$1 \<and> x$1 \<le> a$1)" |
|
655 |
(is "_ = ?R") |
|
656 |
proof - |
|
| 36432 | 657 |
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" |
| 53627 | 658 |
{
|
659 |
presume "?L \<Longrightarrow> ?R" and "?R \<Longrightarrow> ?L" |
|
660 |
then show ?thesis |
|
661 |
unfolding closed_segment_def mem_Collect_eq |
|
| 53628 | 662 |
unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[symmetric] vector_component_simps |
| 53627 | 663 |
by blast |
664 |
} |
|
665 |
{
|
|
666 |
assume ?L |
|
| 55675 | 667 |
then obtain u where u: |
668 |
"x $ 1 = (1 - u) * a $ 1 + u * b $ 1" |
|
669 |
"x $ 2 = (1 - u) * a $ 2 + u * b $ 2" |
|
670 |
"0 \<le> u" |
|
671 |
"u \<le> 1" |
|
672 |
by blast |
|
| 53627 | 673 |
{
|
674 |
fix b a |
|
675 |
assume "b + u * a > a + u * b" |
|
676 |
then have "(1 - u) * b > (1 - u) * a" |
|
| 53628 | 677 |
by (auto simp add: field_simps) |
| 53627 | 678 |
then have "b \<ge> a" |
679 |
apply (drule_tac mult_less_imp_less_left) |
|
680 |
using u |
|
681 |
apply auto |
|
682 |
done |
|
683 |
then have "u * a \<le> u * b" |
|
684 |
apply - |
|
685 |
apply (rule mult_left_mono[OF _ u(3)]) |
|
686 |
using u(3-4) |
|
687 |
apply (auto simp add: field_simps) |
|
688 |
done |
|
689 |
} note * = this |
|
690 |
{
|
|
691 |
fix a b |
|
692 |
assume "u * b > u * a" |
|
693 |
then have "(1 - u) * a \<le> (1 - u) * b" |
|
694 |
apply - |
|
695 |
apply (rule mult_left_mono) |
|
696 |
apply (drule mult_less_imp_less_left) |
|
697 |
using u |
|
698 |
apply auto |
|
699 |
done |
|
700 |
then have "a + u * b \<le> b + u * a" |
|
701 |
by (auto simp add: field_simps) |
|
702 |
} note ** = this |
|
703 |
then show ?R |
|
704 |
unfolding u assms |
|
705 |
using u |
|
706 |
by (auto simp add: field_simps not_le intro: * **) |
|
707 |
} |
|
708 |
{
|
|
709 |
assume ?R |
|
710 |
then show ?L |
|
711 |
proof (cases "x$1 = b$1") |
|
712 |
case True |
|
713 |
then show ?L |
|
714 |
apply (rule_tac x="(x$1 - a$1) / (b$1 - a$1)" in exI) |
|
715 |
unfolding assms True |
|
716 |
using `?R` |
|
717 |
apply (auto simp add: field_simps) |
|
718 |
done |
|
719 |
next |
|
720 |
case False |
|
721 |
then show ?L |
|
722 |
apply (rule_tac x="1 - (x$1 - b$1) / (a$1 - b$1)" in exI) |
|
723 |
unfolding assms |
|
724 |
using `?R` |
|
725 |
apply (auto simp add: field_simps) |
|
726 |
done |
|
727 |
qed |
|
728 |
} |
|
729 |
qed |
|
| 36432 | 730 |
|
| 53627 | 731 |
|
732 |
subsection {* Useful Fashoda corollary pointed out to me by Tom Hales *}
|
|
| 36432 | 733 |
|
| 53627 | 734 |
lemma fashoda_interlace: |
735 |
fixes a :: "real^2" |
|
736 |
assumes "path f" |
|
737 |
and "path g" |
|
| 56188 | 738 |
and "path_image f \<subseteq> cbox a b" |
739 |
and "path_image g \<subseteq> cbox a b" |
|
| 53627 | 740 |
and "(pathstart f)$2 = a$2" |
741 |
and "(pathfinish f)$2 = a$2" |
|
742 |
and "(pathstart g)$2 = a$2" |
|
743 |
and "(pathfinish g)$2 = a$2" |
|
744 |
and "(pathstart f)$1 < (pathstart g)$1" |
|
745 |
and "(pathstart g)$1 < (pathfinish f)$1" |
|
746 |
and "(pathfinish f)$1 < (pathfinish g)$1" |
|
747 |
obtains z where "z \<in> path_image f" and "z \<in> path_image g" |
|
748 |
proof - |
|
| 56188 | 749 |
have "cbox a b \<noteq> {}"
|
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
53628
diff
changeset
|
750 |
using path_image_nonempty[of f] 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
|
751 |
note ab=this[unfolded interval_eq_empty_cart not_ex forall_2 not_less] |
| 56188 | 752 |
have "pathstart f \<in> cbox a b" |
753 |
and "pathfinish f \<in> cbox a b" |
|
754 |
and "pathstart g \<in> cbox a b" |
|
755 |
and "pathfinish g \<in> cbox a b" |
|
| 53628 | 756 |
using pathstart_in_path_image pathfinish_in_path_image |
757 |
using assms(3-4) |
|
758 |
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
|
759 |
note startfin = this[unfolded mem_interval_cart forall_2] |
| 36432 | 760 |
let ?P1 = "linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2]) +++ |
761 |
linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f) +++ f +++ |
|
762 |
linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2]) +++ |
|
763 |
linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2])" |
|
764 |
let ?P2 = "linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g) +++ g +++ |
|
765 |
linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1]) +++ |
|
766 |
linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1]) +++ |
|
767 |
linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3])" |
|
768 |
let ?a = "vector[a$1 - 2, a$2 - 3]" |
|
769 |
let ?b = "vector[b$1 + 2, b$2 + 3]" |
|
| 53627 | 770 |
have P1P2: "path_image ?P1 = path_image (linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2])) \<union> |
| 36432 | 771 |
path_image (linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f)) \<union> path_image f \<union> |
772 |
path_image (linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2])) \<union> |
|
773 |
path_image (linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2]))" |
|
774 |
"path_image ?P2 = path_image(linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g)) \<union> path_image g \<union> |
|
775 |
path_image(linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1])) \<union> |
|
776 |
path_image(linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1])) \<union> |
|
777 |
path_image(linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3]))" using assms(1-2) |
|
778 |
by(auto simp add: path_image_join path_linepath) |
|
| 56188 | 779 |
have abab: "cbox a b \<subseteq> cbox ?a ?b" |
780 |
unfolding interval_cbox_cart[symmetric] |
|
| 53627 | 781 |
by (auto simp add:less_eq_vec_def forall_2 vector_2) |
| 55675 | 782 |
obtain z where |
783 |
"z \<in> path_image |
|
784 |
(linepath (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) +++ |
|
785 |
linepath (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f) +++ |
|
786 |
f +++ |
|
787 |
linepath (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) +++ |
|
788 |
linepath (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]))" |
|
789 |
"z \<in> path_image |
|
790 |
(linepath (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g) +++ |
|
791 |
g +++ |
|
792 |
linepath (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1]) +++ |
|
793 |
linepath (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1]) +++ |
|
794 |
linepath (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]))" |
|
| 53627 | 795 |
apply (rule fashoda[of ?P1 ?P2 ?a ?b]) |
796 |
unfolding pathstart_join pathfinish_join pathstart_linepath pathfinish_linepath vector_2 |
|
797 |
proof - |
|
| 53628 | 798 |
show "path ?P1" and "path ?P2" |
| 53627 | 799 |
using assms by auto |
| 56188 | 800 |
have "path_image ?P1 \<subseteq> cbox ?a ?b" |
| 53627 | 801 |
unfolding P1P2 path_image_linepath |
802 |
apply (rule Un_least)+ |
|
803 |
defer 3 |
|
| 56188 | 804 |
apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format]) |
| 53627 | 805 |
unfolding mem_interval_cart forall_2 vector_2 |
806 |
using ab startfin abab assms(3) |
|
807 |
using assms(9-) |
|
808 |
unfolding assms |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
809 |
apply (auto simp add: field_simps box_def) |
| 53627 | 810 |
done |
| 56188 | 811 |
then show "path_image ?P1 \<subseteq> cbox ?a ?b" . |
812 |
have "path_image ?P2 \<subseteq> cbox ?a ?b" |
|
| 53627 | 813 |
unfolding P1P2 path_image_linepath |
814 |
apply (rule Un_least)+ |
|
815 |
defer 2 |
|
| 56188 | 816 |
apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format]) |
| 53627 | 817 |
unfolding mem_interval_cart forall_2 vector_2 |
818 |
using ab startfin abab assms(4) |
|
819 |
using assms(9-) |
|
820 |
unfolding assms |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
821 |
apply (auto simp add: field_simps box_def) |
| 53627 | 822 |
done |
| 56188 | 823 |
then show "path_image ?P2 \<subseteq> cbox ?a ?b" . |
| 53627 | 824 |
show "a $ 1 - 2 = a $ 1 - 2" |
825 |
and "b $ 1 + 2 = b $ 1 + 2" |
|
826 |
and "pathstart g $ 2 - 3 = a $ 2 - 3" |
|
827 |
and "b $ 2 + 3 = b $ 2 + 3" |
|
828 |
by (auto simp add: assms) |
|
| 53628 | 829 |
qed |
830 |
note z=this[unfolded P1P2 path_image_linepath] |
|
| 53627 | 831 |
show thesis |
832 |
apply (rule that[of z]) |
|
833 |
proof - |
|
| 36432 | 834 |
have "(z \<in> closed_segment (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) \<or> |
| 53627 | 835 |
z \<in> closed_segment (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f)) \<or> |
836 |
z \<in> closed_segment (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) \<or> |
|
837 |
z \<in> closed_segment (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]) \<Longrightarrow> |
|
838 |
(((z \<in> closed_segment (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g)) \<or> |
|
839 |
z \<in> closed_segment (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1])) \<or> |
|
840 |
z \<in> closed_segment (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1])) \<or> |
|
841 |
z \<in> closed_segment (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]) \<Longrightarrow> False" |
|
842 |
apply (simp only: segment_vertical segment_horizontal vector_2) |
|
843 |
proof - |
|
844 |
case goal1 note as=this |
|
| 56188 | 845 |
have "pathfinish f \<in> cbox a b" |
| 53627 | 846 |
using assms(3) pathfinish_in_path_image[of f] by auto |
| 53628 | 847 |
then have "1 + b $ 1 \<le> pathfinish f $ 1 \<Longrightarrow> False" |
| 53627 | 848 |
unfolding mem_interval_cart forall_2 by auto |
849 |
then have "z$1 \<noteq> pathfinish f$1" |
|
| 53628 | 850 |
using as(2) |
851 |
using assms ab |
|
852 |
by (auto simp add: field_simps) |
|
| 56188 | 853 |
moreover have "pathstart f \<in> cbox a b" |
| 53628 | 854 |
using assms(3) pathstart_in_path_image[of f] |
855 |
by auto |
|
| 53627 | 856 |
then have "1 + b $ 1 \<le> pathstart f $ 1 \<Longrightarrow> False" |
| 53628 | 857 |
unfolding mem_interval_cart forall_2 |
858 |
by auto |
|
| 53627 | 859 |
then have "z$1 \<noteq> pathstart f$1" |
| 53628 | 860 |
using as(2) using assms ab |
861 |
by (auto simp add: field_simps) |
|
| 53627 | 862 |
ultimately have *: "z$2 = a$2 - 2" |
| 53628 | 863 |
using goal1(1) |
864 |
by auto |
|
| 53627 | 865 |
have "z$1 \<noteq> pathfinish g$1" |
| 53628 | 866 |
using as(2) |
867 |
using assms ab |
|
868 |
by (auto simp add: field_simps *) |
|
| 56188 | 869 |
moreover have "pathstart g \<in> cbox a b" |
| 53628 | 870 |
using assms(4) pathstart_in_path_image[of g] |
871 |
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
|
872 |
note this[unfolded mem_interval_cart forall_2] |
| 53627 | 873 |
then have "z$1 \<noteq> pathstart g$1" |
| 53628 | 874 |
using as(1) |
875 |
using assms ab |
|
876 |
by (auto simp add: field_simps *) |
|
| 36432 | 877 |
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" |
| 53628 | 878 |
using as(2) |
879 |
unfolding * assms |
|
880 |
by (auto simp add: field_simps) |
|
| 53627 | 881 |
then show False |
882 |
unfolding * using ab by auto |
|
883 |
qed |
|
884 |
then have "z \<in> path_image f \<or> z \<in> path_image g" |
|
885 |
using z unfolding Un_iff by blast |
|
| 56188 | 886 |
then have z': "z \<in> cbox a b" |
| 53628 | 887 |
using assms(3-4) |
888 |
by auto |
|
| 53627 | 889 |
have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart f $ 1 \<or> z $ 1 = pathfinish f $ 1) \<Longrightarrow> |
890 |
z = pathstart f \<or> z = pathfinish f" |
|
| 53628 | 891 |
unfolding vec_eq_iff forall_2 assms |
892 |
by auto |
|
| 53627 | 893 |
with z' show "z \<in> path_image f" |
894 |
using z(1) |
|
895 |
unfolding Un_iff mem_interval_cart forall_2 |
|
896 |
apply - |
|
897 |
apply (simp only: segment_vertical segment_horizontal vector_2) |
|
898 |
unfolding assms |
|
899 |
apply auto |
|
900 |
done |
|
901 |
have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart g $ 1 \<or> z $ 1 = pathfinish g $ 1) \<Longrightarrow> |
|
902 |
z = pathstart g \<or> z = pathfinish g" |
|
| 53628 | 903 |
unfolding vec_eq_iff forall_2 assms |
904 |
by auto |
|
| 53627 | 905 |
with z' show "z \<in> path_image g" |
906 |
using z(2) |
|
907 |
unfolding Un_iff mem_interval_cart forall_2 |
|
908 |
apply (simp only: segment_vertical segment_horizontal vector_2) |
|
909 |
unfolding assms |
|
910 |
apply auto |
|
911 |
done |
|
912 |
qed |
|
913 |
qed |
|
| 36432 | 914 |
|
915 |
(** The Following still needs to be translated. Maybe I will do that later. |
|
916 |
||
917 |
(* ------------------------------------------------------------------------- *) |
|
918 |
(* Complement in dimension N >= 2 of set homeomorphic to any interval in *) |
|
919 |
(* any dimension is (path-)connected. This naively generalizes the argument *) |
|
920 |
(* in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer *) |
|
921 |
(* fixed point theorem", American Mathematical Monthly 1984. *) |
|
922 |
(* ------------------------------------------------------------------------- *) |
|
923 |
||
924 |
let RETRACTION_INJECTIVE_IMAGE_INTERVAL = prove |
|
925 |
(`!p:real^M->real^N a b. |
|
926 |
~(interval[a,b] = {}) /\
|
|
927 |
p continuous_on interval[a,b] /\ |
|
928 |
(!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ p x = p y ==> x = y) |
|
929 |
==> ?f. f continuous_on (:real^N) /\ |
|
930 |
IMAGE f (:real^N) SUBSET (IMAGE p (interval[a,b])) /\ |
|
931 |
(!x. x IN (IMAGE p (interval[a,b])) ==> f x = x)`, |
|
932 |
REPEAT STRIP_TAC THEN |
|
933 |
FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [INJECTIVE_ON_LEFT_INVERSE]) THEN |
|
934 |
DISCH_THEN(X_CHOOSE_TAC `q:real^N->real^M`) THEN |
|
935 |
SUBGOAL_THEN `(q:real^N->real^M) continuous_on |
|
936 |
(IMAGE p (interval[a:real^M,b]))` |
|
937 |
ASSUME_TAC THENL |
|
938 |
[MATCH_MP_TAC CONTINUOUS_ON_INVERSE THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]; |
|
939 |
ALL_TAC] THEN |
|
940 |
MP_TAC(ISPECL [`q:real^N->real^M`; |
|
941 |
`IMAGE (p:real^M->real^N) |
|
942 |
(interval[a,b])`; |
|
943 |
`a:real^M`; `b:real^M`] |
|
944 |
TIETZE_CLOSED_INTERVAL) THEN |
|
945 |
ASM_SIMP_TAC[COMPACT_INTERVAL; COMPACT_CONTINUOUS_IMAGE; |
|
946 |
COMPACT_IMP_CLOSED] THEN |
|
947 |
ANTS_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN |
|
948 |
DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^M` STRIP_ASSUME_TAC) THEN |
|
949 |
EXISTS_TAC `(p:real^M->real^N) o (r:real^N->real^M)` THEN |
|
950 |
REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; o_THM; IN_UNIV] THEN |
|
951 |
CONJ_TAC THENL [ALL_TAC; ASM SET_TAC[]] THEN |
|
952 |
MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN ASM_REWRITE_TAC[] THEN |
|
953 |
FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP(REWRITE_RULE[IMP_CONJ] |
|
954 |
CONTINUOUS_ON_SUBSET)) THEN ASM SET_TAC[]);; |
|
955 |
||
956 |
let UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove |
|
957 |
(`!s:real^N->bool a b:real^M. |
|
958 |
s homeomorphic (interval[a,b]) |
|
959 |
==> !x. ~(x IN s) ==> ~bounded(path_component((:real^N) DIFF s) x)`, |
|
960 |
REPEAT GEN_TAC THEN REWRITE_TAC[homeomorphic; homeomorphism] THEN |
|
961 |
REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN |
|
962 |
MAP_EVERY X_GEN_TAC [`p':real^N->real^M`; `p:real^M->real^N`] THEN |
|
963 |
DISCH_TAC THEN |
|
964 |
SUBGOAL_THEN |
|
965 |
`!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ |
|
966 |
(p:real^M->real^N) x = p y ==> x = y` |
|
967 |
ASSUME_TAC THENL [ASM_MESON_TAC[]; ALL_TAC] THEN |
|
968 |
FIRST_X_ASSUM(MP_TAC o funpow 4 CONJUNCT2) THEN |
|
969 |
DISCH_THEN(CONJUNCTS_THEN2 (SUBST1_TAC o SYM) ASSUME_TAC) THEN |
|
970 |
ASM_CASES_TAC `interval[a:real^M,b] = {}` THEN
|
|
971 |
ASM_REWRITE_TAC[IMAGE_CLAUSES; DIFF_EMPTY; PATH_COMPONENT_UNIV; |
|
972 |
NOT_BOUNDED_UNIV] THEN |
|
973 |
ABBREV_TAC `s = (:real^N) DIFF (IMAGE p (interval[a:real^M,b]))` THEN |
|
974 |
X_GEN_TAC `c:real^N` THEN REPEAT STRIP_TAC THEN |
|
975 |
SUBGOAL_THEN `(c:real^N) IN s` ASSUME_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN |
|
976 |
SUBGOAL_THEN `bounded((path_component s c) UNION |
|
977 |
(IMAGE (p:real^M->real^N) (interval[a,b])))` |
|
978 |
MP_TAC THENL |
|
979 |
[ASM_SIMP_TAC[BOUNDED_UNION; COMPACT_IMP_BOUNDED; COMPACT_IMP_BOUNDED; |
|
980 |
COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; |
|
981 |
ALL_TAC] THEN |
|
982 |
DISCH_THEN(MP_TAC o SPEC `c:real^N` o MATCH_MP BOUNDED_SUBSET_BALL) THEN |
|
983 |
REWRITE_TAC[UNION_SUBSET] THEN |
|
984 |
DISCH_THEN(X_CHOOSE_THEN `B:real` STRIP_ASSUME_TAC) THEN |
|
985 |
MP_TAC(ISPECL [`p:real^M->real^N`; `a:real^M`; `b:real^M`] |
|
986 |
RETRACTION_INJECTIVE_IMAGE_INTERVAL) THEN |
|
987 |
ASM_REWRITE_TAC[SUBSET; IN_UNIV] THEN |
|
988 |
DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^N` MP_TAC) THEN |
|
989 |
DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC |
|
990 |
(CONJUNCTS_THEN2 MP_TAC ASSUME_TAC)) THEN |
|
991 |
REWRITE_TAC[FORALL_IN_IMAGE; IN_UNIV] THEN DISCH_TAC THEN |
|
992 |
ABBREV_TAC `q = \z:real^N. if z IN path_component s c then r(z) else z` THEN |
|
993 |
SUBGOAL_THEN |
|
994 |
`(q:real^N->real^N) continuous_on |
|
995 |
(closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)))` |
|
996 |
MP_TAC THENL |
|
997 |
[EXPAND_TAC "q" THEN MATCH_MP_TAC CONTINUOUS_ON_CASES THEN |
|
998 |
REWRITE_TAC[CLOSED_CLOSURE; CONTINUOUS_ON_ID; GSYM OPEN_CLOSED] THEN |
|
999 |
REPEAT CONJ_TAC THENL |
|
1000 |
[MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN |
|
1001 |
ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED; |
|
1002 |
COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; |
|
1003 |
ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; |
|
1004 |
ALL_TAC] THEN |
|
1005 |
X_GEN_TAC `z:real^N` THEN |
|
1006 |
REWRITE_TAC[SET_RULE `~(z IN (s DIFF t) /\ z IN t)`] THEN |
|
1007 |
STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN |
|
1008 |
MP_TAC(ISPECL |
|
1009 |
[`path_component s (z:real^N)`; `path_component s (c:real^N)`] |
|
1010 |
OPEN_INTER_CLOSURE_EQ_EMPTY) THEN |
|
1011 |
ASM_REWRITE_TAC[GSYM DISJOINT; PATH_COMPONENT_DISJOINT] THEN ANTS_TAC THENL |
|
1012 |
[MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN |
|
1013 |
ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED; |
|
1014 |
COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL]; |
|
1015 |
REWRITE_TAC[DISJOINT; EXTENSION; IN_INTER; NOT_IN_EMPTY] THEN |
|
1016 |
DISCH_THEN(MP_TAC o SPEC `z:real^N`) THEN ASM_REWRITE_TAC[] THEN |
|
1017 |
GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [IN] THEN |
|
1018 |
REWRITE_TAC[PATH_COMPONENT_REFL_EQ] THEN ASM SET_TAC[]]; |
|
1019 |
ALL_TAC] THEN |
|
1020 |
SUBGOAL_THEN |
|
1021 |
`closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)) = |
|
1022 |
(:real^N)` |
|
1023 |
SUBST1_TAC THENL |
|
1024 |
[MATCH_MP_TAC(SET_RULE `s SUBSET t ==> t UNION (UNIV DIFF s) = UNIV`) THEN |
|
1025 |
REWRITE_TAC[CLOSURE_SUBSET]; |
|
1026 |
DISCH_TAC] THEN |
|
1027 |
MP_TAC(ISPECL |
|
1028 |
[`(\x. &2 % c - x) o |
|
1029 |
(\x. c + B / norm(x - c) % (x - c)) o (q:real^N->real^N)`; |
|
1030 |
`cball(c:real^N,B)`] |
|
1031 |
BROUWER) THEN |
|
1032 |
REWRITE_TAC[NOT_IMP; GSYM CONJ_ASSOC; COMPACT_CBALL; CONVEX_CBALL] THEN |
|
1033 |
ASM_SIMP_TAC[CBALL_EQ_EMPTY; REAL_LT_IMP_LE; REAL_NOT_LT] THEN |
|
1034 |
SUBGOAL_THEN `!x. ~((q:real^N->real^N) x = c)` ASSUME_TAC THENL |
|
1035 |
[X_GEN_TAC `x:real^N` THEN EXPAND_TAC "q" THEN |
|
1036 |
REWRITE_TAC[NORM_EQ_0; VECTOR_SUB_EQ] THEN COND_CASES_TAC THEN |
|
1037 |
ASM SET_TAC[PATH_COMPONENT_REFL_EQ]; |
|
1038 |
ALL_TAC] THEN |
|
1039 |
REPEAT CONJ_TAC THENL |
|
1040 |
[MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN |
|
1041 |
SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN |
|
1042 |
MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN CONJ_TAC THENL |
|
1043 |
[ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; ALL_TAC] THEN |
|
1044 |
MATCH_MP_TAC CONTINUOUS_ON_ADD THEN REWRITE_TAC[CONTINUOUS_ON_CONST] THEN |
|
1045 |
MATCH_MP_TAC CONTINUOUS_ON_MUL THEN |
|
1046 |
SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN |
|
1047 |
REWRITE_TAC[o_DEF; real_div; LIFT_CMUL] THEN |
|
1048 |
MATCH_MP_TAC CONTINUOUS_ON_CMUL THEN |
|
1049 |
MATCH_MP_TAC(REWRITE_RULE[o_DEF] CONTINUOUS_ON_INV) THEN |
|
1050 |
ASM_REWRITE_TAC[FORALL_IN_IMAGE; NORM_EQ_0; VECTOR_SUB_EQ] THEN |
|
1051 |
SUBGOAL_THEN |
|
1052 |
`(\x:real^N. lift(norm(x - c))) = (lift o norm) o (\x. x - c)` |
|
1053 |
SUBST1_TAC THENL [REWRITE_TAC[o_DEF]; ALL_TAC] THEN |
|
1054 |
MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN |
|
1055 |
ASM_SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST; |
|
1056 |
CONTINUOUS_ON_LIFT_NORM]; |
|
1057 |
REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; IN_CBALL; o_THM; dist] THEN |
|
1058 |
X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN |
|
1059 |
REWRITE_TAC[VECTOR_ARITH `c - (&2 % c - (c + x)) = x`] THEN |
|
1060 |
REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN |
|
1061 |
ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN |
|
1062 |
ASM_REAL_ARITH_TAC; |
|
1063 |
REWRITE_TAC[NOT_EXISTS_THM; TAUT `~(c /\ b) <=> c ==> ~b`] THEN |
|
1064 |
REWRITE_TAC[IN_CBALL; o_THM; dist] THEN |
|
1065 |
X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN |
|
1066 |
REWRITE_TAC[VECTOR_ARITH `&2 % c - (c + x') = x <=> --x' = x - c`] THEN |
|
1067 |
ASM_CASES_TAC `(x:real^N) IN path_component s c` THENL |
|
1068 |
[MATCH_MP_TAC(NORM_ARITH `norm(y) < B /\ norm(x) = B ==> ~(--x = y)`) THEN |
|
1069 |
REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN |
|
1070 |
ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN |
|
1071 |
ASM_SIMP_TAC[REAL_ARITH `&0 < B ==> abs B = B`] THEN |
|
1072 |
UNDISCH_TAC `path_component s c SUBSET ball(c:real^N,B)` THEN |
|
1073 |
REWRITE_TAC[SUBSET; IN_BALL] THEN ASM_MESON_TAC[dist; NORM_SUB]; |
|
1074 |
EXPAND_TAC "q" THEN REWRITE_TAC[] THEN ASM_REWRITE_TAC[] THEN |
|
1075 |
REWRITE_TAC[VECTOR_ARITH `--(c % x) = x <=> (&1 + c) % x = vec 0`] THEN |
|
1076 |
ASM_REWRITE_TAC[DE_MORGAN_THM; VECTOR_SUB_EQ; VECTOR_MUL_EQ_0] THEN |
|
1077 |
SUBGOAL_THEN `~(x:real^N = c)` ASSUME_TAC THENL |
|
1078 |
[ASM_MESON_TAC[PATH_COMPONENT_REFL; IN]; ALL_TAC] THEN |
|
1079 |
ASM_REWRITE_TAC[] THEN |
|
1080 |
MATCH_MP_TAC(REAL_ARITH `&0 < x ==> ~(&1 + x = &0)`) THEN |
|
1081 |
ASM_SIMP_TAC[REAL_LT_DIV; NORM_POS_LT; VECTOR_SUB_EQ]]]);; |
|
1082 |
||
1083 |
let PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove |
|
1084 |
(`!s:real^N->bool a b:real^M. |
|
1085 |
2 <= dimindex(:N) /\ s homeomorphic interval[a,b] |
|
1086 |
==> path_connected((:real^N) DIFF s)`, |
|
1087 |
REPEAT STRIP_TAC THEN REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN |
|
1088 |
FIRST_ASSUM(MP_TAC o MATCH_MP |
|
1089 |
UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN |
|
1090 |
ASM_REWRITE_TAC[SET_RULE `~(x IN s) <=> x IN (UNIV DIFF s)`] THEN |
|
1091 |
ABBREV_TAC `t = (:real^N) DIFF s` THEN |
|
1092 |
DISCH_TAC THEN MAP_EVERY X_GEN_TAC [`x:real^N`; `y:real^N`] THEN |
|
1093 |
STRIP_TAC THEN FIRST_ASSUM(MP_TAC o MATCH_MP HOMEOMORPHIC_COMPACTNESS) THEN |
|
1094 |
REWRITE_TAC[COMPACT_INTERVAL] THEN |
|
1095 |
DISCH_THEN(MP_TAC o MATCH_MP COMPACT_IMP_BOUNDED) THEN |
|
1096 |
REWRITE_TAC[BOUNDED_POS; LEFT_IMP_EXISTS_THM] THEN |
|
1097 |
X_GEN_TAC `B:real` THEN STRIP_TAC THEN |
|
1098 |
SUBGOAL_THEN `(?u:real^N. u IN path_component t x /\ B < norm(u)) /\ |
|
1099 |
(?v:real^N. v IN path_component t y /\ B < norm(v))` |
|
1100 |
STRIP_ASSUME_TAC THENL |
|
1101 |
[ASM_MESON_TAC[BOUNDED_POS; REAL_NOT_LE]; ALL_TAC] THEN |
|
1102 |
MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `u:real^N` THEN |
|
1103 |
CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN |
|
1104 |
MATCH_MP_TAC PATH_COMPONENT_SYM THEN |
|
1105 |
MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `v:real^N` THEN |
|
1106 |
CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN |
|
1107 |
MATCH_MP_TAC PATH_COMPONENT_OF_SUBSET THEN |
|
1108 |
EXISTS_TAC `(:real^N) DIFF cball(vec 0,B)` THEN CONJ_TAC THENL |
|
1109 |
[EXPAND_TAC "t" THEN MATCH_MP_TAC(SET_RULE |
|
1110 |
`s SUBSET t ==> (u DIFF t) SUBSET (u DIFF s)`) THEN |
|
1111 |
ASM_REWRITE_TAC[SUBSET; IN_CBALL_0]; |
|
1112 |
MP_TAC(ISPEC `cball(vec 0:real^N,B)` |
|
1113 |
PATH_CONNECTED_COMPLEMENT_BOUNDED_CONVEX) THEN |
|
1114 |
ASM_REWRITE_TAC[BOUNDED_CBALL; CONVEX_CBALL] THEN |
|
1115 |
REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN |
|
1116 |
DISCH_THEN MATCH_MP_TAC THEN |
|
1117 |
ASM_REWRITE_TAC[IN_DIFF; IN_UNIV; IN_CBALL_0; REAL_NOT_LE]]);; |
|
1118 |
||
1119 |
(* ------------------------------------------------------------------------- *) |
|
1120 |
(* In particular, apply all these to the special case of an arc. *) |
|
1121 |
(* ------------------------------------------------------------------------- *) |
|
1122 |
||
1123 |
let RETRACTION_ARC = prove |
|
1124 |
(`!p. arc p |
|
1125 |
==> ?f. f continuous_on (:real^N) /\ |
|
1126 |
IMAGE f (:real^N) SUBSET path_image p /\ |
|
1127 |
(!x. x IN path_image p ==> f x = x)`, |
|
1128 |
REWRITE_TAC[arc; path; path_image] THEN REPEAT STRIP_TAC THEN |
|
1129 |
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
|
1130 |
ASM_REWRITE_TAC[INTERVAL_EQ_EMPTY_CART_1; DROP_VEC; REAL_NOT_LT; REAL_POS]);; |
| 36432 | 1131 |
|
1132 |
let PATH_CONNECTED_ARC_COMPLEMENT = prove |
|
1133 |
(`!p. 2 <= dimindex(:N) /\ arc p |
|
1134 |
==> path_connected((:real^N) DIFF path_image p)`, |
|
1135 |
REWRITE_TAC[arc; path] THEN REPEAT STRIP_TAC THEN SIMP_TAC[path_image] THEN |
|
1136 |
MP_TAC(ISPECL [`path_image p:real^N->bool`; `vec 0:real^1`; `vec 1:real^1`] |
|
1137 |
PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN |
|
1138 |
ASM_REWRITE_TAC[path_image] THEN DISCH_THEN MATCH_MP_TAC THEN |
|
1139 |
ONCE_REWRITE_TAC[HOMEOMORPHIC_SYM] THEN |
|
1140 |
MATCH_MP_TAC HOMEOMORPHIC_COMPACT THEN |
|
1141 |
EXISTS_TAC `p:real^1->real^N` THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]);; |
|
1142 |
||
1143 |
let CONNECTED_ARC_COMPLEMENT = prove |
|
1144 |
(`!p. 2 <= dimindex(:N) /\ arc p |
|
1145 |
==> connected((:real^N) DIFF path_image p)`, |
|
1146 |
SIMP_TAC[PATH_CONNECTED_ARC_COMPLEMENT; PATH_CONNECTED_IMP_CONNECTED]);; *) |
|
1147 |
||
1148 |
end |