author | haftmann |
Thu, 19 Jun 2025 17:15:40 +0200 | |
changeset 82734 | 89347c0cc6a3 |
parent 78456 | 57f5127d2ff2 |
permissions | -rw-r--r-- |
53572 | 1 |
(* Author: John Harrison |
2 |
Author: Robert Himmelmann, TU Muenchen (translation from HOL light) |
|
3 |
*) |
|
36432 | 4 |
|
69722
b5163b2132c5
minor tagging updates in 13 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
5 |
section \<open>Fashoda Meet Theorem\<close> |
36432 | 6 |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63040
diff
changeset
|
7 |
theory Fashoda_Theorem |
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 |
||
69683 | 11 |
subsection \<open>Bijections between intervals\<close> |
56273
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
12 |
|
70136 | 13 |
definition\<^marker>\<open>tag important\<close> interval_bij :: "'a \<times> 'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<Rightarrow> 'a::euclidean_space" |
56273
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
14 |
where "interval_bij = |
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
15 |
(\<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
|
16 |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
17 |
lemma interval_bij_affine: |
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
18 |
"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
|
19 |
(\<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))" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
20 |
by (simp add: interval_bij_def algebra_simps add_divide_distrib diff_divide_distrib flip: sum.distrib scaleR_add_left) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
21 |
|
56273
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
22 |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
23 |
lemma continuous_interval_bij: |
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
24 |
fixes a b :: "'a::euclidean_space" |
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
25 |
shows "continuous (at x) (interval_bij (a, b) (u, v))" |
64267 | 26 |
by (auto simp add: divide_inverse interval_bij_def intro!: continuous_sum continuous_intros) |
56273
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
27 |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
28 |
lemma continuous_on_interval_bij: "continuous_on s (interval_bij (a, b) (u, v))" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
29 |
by (metis continuous_at_imp_continuous_on continuous_interval_bij) |
56273
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
30 |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
31 |
lemma in_interval_interval_bij: |
56273
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
32 |
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
|
33 |
assumes "x \<in> cbox a b" |
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
34 |
and "cbox u v \<noteq> {}" |
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
35 |
shows "interval_bij (a, b) (u, v) x \<in> cbox u v" |
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
36 |
proof - |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
37 |
have "\<And>i. i \<in> Basis \<Longrightarrow> 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)" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
38 |
by (smt (verit) assms box_ne_empty(1) divide_nonneg_nonneg mem_box(2) mult_nonneg_nonneg) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
39 |
moreover |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
40 |
have "\<And>i. i \<in> Basis \<Longrightarrow> 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" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
41 |
apply (simp add: divide_simps algebra_simps) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
42 |
by (smt (verit, best) assms box_ne_empty(1) left_diff_distrib mem_box(2) mult.commute mult_left_mono) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
43 |
ultimately show ?thesis |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
44 |
by (force simp only: interval_bij_def split_conv mem_box inner_sum_left_Basis) |
56273
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
45 |
qed |
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
46 |
|
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
47 |
lemma interval_bij_bij: |
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
48 |
"\<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
|
49 |
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
|
50 |
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
|
51 |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63040
diff
changeset
|
52 |
lemma interval_bij_bij_cart: fixes x::"real^'n" assumes "\<forall>i. a$i < b$i \<and> u$i < v$i" |
56273
def3bbe6f2a5
cleanup auxiliary proofs for Brouwer fixpoint theorem (removes ~2400 lines)
hoelzl
parents:
56189
diff
changeset
|
53 |
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
|
54 |
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
|
55 |
|
53572 | 56 |
|
69683 | 57 |
subsection \<open>Fashoda meet theorem\<close> |
36432 | 58 |
|
53572 | 59 |
lemma infnorm_2: |
60 |
fixes x :: "real^2" |
|
61945 | 61 |
shows "infnorm x = max \<bar>x$1\<bar> \<bar>x$2\<bar>" |
53572 | 62 |
unfolding infnorm_cart UNIV_2 by (rule cSup_eq) auto |
36432 | 63 |
|
53572 | 64 |
lemma infnorm_eq_1_2: |
65 |
fixes x :: "real^2" |
|
66 |
shows "infnorm x = 1 \<longleftrightarrow> |
|
61945 | 67 |
\<bar>x$1\<bar> \<le> 1 \<and> \<bar>x$2\<bar> \<le> 1 \<and> (x$1 = -1 \<or> x$1 = 1 \<or> x$2 = -1 \<or> x$2 = 1)" |
36432 | 68 |
unfolding infnorm_2 by auto |
69 |
||
53572 | 70 |
lemma infnorm_eq_1_imp: |
71 |
fixes x :: "real^2" |
|
72 |
assumes "infnorm x = 1" |
|
61945 | 73 |
shows "\<bar>x$1\<bar> \<le> 1" and "\<bar>x$2\<bar> \<le> 1" |
36432 | 74 |
using assms unfolding infnorm_eq_1_2 by auto |
75 |
||
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
76 |
proposition fashoda_unit: |
53572 | 77 |
fixes f g :: "real \<Rightarrow> real^2" |
56188 | 78 |
assumes "f ` {-1 .. 1} \<subseteq> cbox (-1) 1" |
79 |
and "g ` {-1 .. 1} \<subseteq> cbox (-1) 1" |
|
80 |
and "continuous_on {-1 .. 1} f" |
|
81 |
and "continuous_on {-1 .. 1} g" |
|
53572 | 82 |
and "f (- 1)$1 = - 1" |
83 |
and "f 1$1 = 1" "g (- 1) $2 = -1" |
|
84 |
and "g 1 $2 = 1" |
|
56188 | 85 |
shows "\<exists>s\<in>{-1 .. 1}. \<exists>t\<in>{-1 .. 1}. f s = g t" |
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
86 |
proof (rule ccontr) |
53572 | 87 |
assume "\<not> ?thesis" |
88 |
note as = this[unfolded bex_simps,rule_format] |
|
63040 | 89 |
define sqprojection |
90 |
where [abs_def]: "sqprojection z = (inverse (infnorm z)) *\<^sub>R z" for z :: "real^2" |
|
91 |
define negatex :: "real^2 \<Rightarrow> real^2" |
|
92 |
where "negatex x = (vector [-(x$1), x$2])" for x |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
93 |
have inf_nega: "\<And>z::real^2. infnorm (negatex z) = infnorm z" |
36432 | 94 |
unfolding negatex_def infnorm_2 vector_2 by auto |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
95 |
have inf_eq1: "\<And>z. z \<noteq> 0 \<Longrightarrow> infnorm (sqprojection z) = 1" |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
96 |
unfolding sqprojection_def infnorm_mul[unfolded scalar_mult_eq_scaleR] |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
97 |
by (simp add: real_abs_infnorm infnorm_eq_0) |
53572 | 98 |
let ?F = "\<lambda>w::real^2. (f \<circ> (\<lambda>x. x$1)) w - (g \<circ> (\<lambda>x. x$2)) w" |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
99 |
have *: "\<And>i. (\<lambda>x::real^2. x $ i) ` cbox (- 1) 1 = {-1..1}" |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
100 |
proof |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
101 |
show "(\<lambda>x::real^2. x $ i) ` cbox (- 1) 1 \<subseteq> {-1..1}" for i |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
102 |
by (auto simp: mem_box_cart) |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
103 |
show "{-1..1} \<subseteq> (\<lambda>x::real^2. x $ i) ` cbox (- 1) 1" for i |
73932
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
71633
diff
changeset
|
104 |
by (clarsimp simp: image_iff mem_box_cart Bex_def) (metis (no_types, opaque_lifting) vec_component) |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
105 |
qed |
53572 | 106 |
{ |
107 |
fix x |
|
56188 | 108 |
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 | 109 |
then obtain w :: "real^2" where w: |
56188 | 110 |
"w \<in> cbox (- 1) 1" |
55675 | 111 |
"x = (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w" |
112 |
unfolding image_iff .. |
|
53572 | 113 |
then have "x \<noteq> 0" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
114 |
using as[of "w$1" "w$2"] by (auto simp: mem_box_cart atLeastAtMost_iff) |
53572 | 115 |
} note x0 = this |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
116 |
let ?CB11 = "cbox (- 1) (1::real^2)" |
55675 | 117 |
obtain x :: "real^2" where x: |
58410
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents:
57418
diff
changeset
|
118 |
"x \<in> cbox (- 1) 1" |
55675 | 119 |
"(negatex \<circ> sqprojection \<circ> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w)) x = x" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
120 |
proof (rule brouwer_weak[of ?CB11 "negatex \<circ> sqprojection \<circ> ?F"]) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
121 |
show "compact ?CB11" "convex ?CB11" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
122 |
by (rule compact_cbox convex_box)+ |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
123 |
have "box (- 1) (1::real^2) \<noteq> {}" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
124 |
unfolding interval_eq_empty_cart by auto |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
125 |
then show "interior ?CB11 \<noteq> {}" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
126 |
by simp |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
127 |
have "negatex (x + y) $ i = (negatex x + negatex y) $ i \<and> negatex (c *\<^sub>R x) $ i = (c *\<^sub>R negatex x) $ i" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
128 |
for i x y c |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
129 |
using exhaust_2 [of i] by (auto simp: negatex_def) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
130 |
then have "bounded_linear negatex" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
131 |
by (simp add: bounded_linearI' vec_eq_iff) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
132 |
then show "continuous_on ?CB11 (negatex \<circ> sqprojection \<circ> ?F)" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
133 |
unfolding sqprojection_def |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
134 |
apply (intro continuous_intros continuous_on_component | use * assms in presburger)+ |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
135 |
apply (simp_all add: infnorm_eq_0 x0 linear_continuous_on) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
136 |
done |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
137 |
have "(negatex \<circ> sqprojection \<circ> ?F) ` ?CB11 \<subseteq> ?CB11" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
138 |
proof clarsimp |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
139 |
fix y :: "real^2" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
140 |
assume y: "y \<in> ?CB11" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
141 |
have "?F y \<noteq> 0" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
142 |
by (rule x0) (use y in auto) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
143 |
then have *: "infnorm (sqprojection (?F y)) = 1" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
144 |
using inf_eq1 by blast |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
145 |
show "negatex (sqprojection (f (y $ 1) - g (y $ 2))) \<in> cbox (-1) 1" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
146 |
unfolding mem_box_cart interval_cbox_cart infnorm_2 |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
147 |
by (smt (verit, del_insts) "*" component_le_infnorm_cart inf_nega neg_one_index o_apply one_index) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
148 |
qed |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
149 |
then show "negatex \<circ> sqprojection \<circ> ?F \<in> ?CB11 \<rightarrow> ?CB11" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
150 |
by blast |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
151 |
qed |
53572 | 152 |
have "?F x \<noteq> 0" |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
153 |
by (rule x0) (use x in auto) |
53572 | 154 |
then have *: "infnorm (sqprojection (?F x)) = 1" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
155 |
using inf_eq1 by blast |
53572 | 156 |
have nx: "infnorm x = 1" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
157 |
by (metis (no_types, lifting) "*" inf_nega o_apply x(2)) |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
158 |
have iff: "0 < sqprojection x$i \<longleftrightarrow> 0 < x$i" "sqprojection x$i < 0 \<longleftrightarrow> x$i < 0" if "x \<noteq> 0" for x i |
53572 | 159 |
proof - |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
160 |
have *: "inverse (infnorm x) > 0" |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
161 |
by (simp add: infnorm_pos_lt that) |
53572 | 162 |
then show "(0 < sqprojection x $ i) = (0 < x $ i)" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
163 |
by (simp add: sqprojection_def zero_less_mult_iff) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
164 |
show "(sqprojection x $ i < 0) = (x $ i < 0)" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
165 |
unfolding sqprojection_def |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
166 |
by (metis * pos_less_divideR_eq scaleR_zero_right vector_scaleR_component) |
53572 | 167 |
qed |
168 |
have x1: "x $ 1 \<in> {- 1..1::real}" "x $ 2 \<in> {- 1..1::real}" |
|
67673
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
169 |
using x(1) unfolding mem_box_cart by auto |
53572 | 170 |
then have nz: "f (x $ 1) - g (x $ 2) \<noteq> 0" |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
171 |
using as by auto |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
172 |
consider "x $ 1 = -1" | "x $ 1 = 1" | "x $ 2 = -1" | "x $ 2 = 1" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63040
diff
changeset
|
173 |
using nx unfolding infnorm_eq_1_2 by auto |
53572 | 174 |
then show False |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
175 |
proof cases |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
176 |
case 1 |
53572 | 177 |
then have *: "f (x $ 1) $ 1 = - 1" |
178 |
using assms(5) by auto |
|
36432 | 179 |
have "sqprojection (f (x$1) - g (x$2)) $ 1 > 0" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
180 |
by (smt (verit) "1" negatex_def o_apply vector_2(1) x(2)) |
53572 | 181 |
moreover |
56188 | 182 |
from x1 have "g (x $ 2) \<in> cbox (-1) 1" |
67982
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents:
67968
diff
changeset
|
183 |
using assms(2) by blast |
53572 | 184 |
ultimately show False |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
185 |
unfolding iff[OF nz] vector_component_simps * mem_box_cart |
67982
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents:
67968
diff
changeset
|
186 |
using not_le by auto |
53572 | 187 |
next |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
188 |
case 2 |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
189 |
then have *: "f (x $ 1) $ 1 = 1" |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
190 |
using assms(6) by auto |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
191 |
have "sqprojection (f (x$1) - g (x$2)) $ 1 < 0" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
192 |
by (smt (verit) "2" negatex_def o_apply vector_2(1) x(2) zero_less_one) |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
193 |
moreover have "g (x $ 2) \<in> cbox (-1) 1" |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
194 |
using assms(2) x1 by blast |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
195 |
ultimately show False |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
196 |
unfolding iff[OF nz] vector_component_simps * mem_box_cart |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
197 |
using not_le by auto |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
198 |
next |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
199 |
case 3 |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
200 |
then have *: "g (x $ 2) $ 2 = - 1" |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
201 |
using assms(7) by auto |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
202 |
moreover have "sqprojection (f (x$1) - g (x$2)) $ 2 < 0" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
203 |
by (smt (verit, ccfv_SIG) "3" negatex_def o_apply vector_2(2) x(2)) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
204 |
moreover from x1 have "f (x $ 1) \<in> cbox (-1) 1" |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
205 |
using assms(1) by blast |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
206 |
ultimately show False |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
207 |
by (smt (verit, del_insts) iff(2) mem_box_cart(2) neg_one_index nz vector_minus_component) |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
208 |
next |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
209 |
case 4 |
53572 | 210 |
then have *: "g (x $ 2) $ 2 = 1" |
211 |
using assms(8) by auto |
|
36432 | 212 |
have "sqprojection (f (x$1) - g (x$2)) $ 2 > 0" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
213 |
by (smt (verit, best) "4" negatex_def o_apply vector_2(2) x(2)) |
53572 | 214 |
moreover |
56188 | 215 |
from x1 have "f (x $ 1) \<in> cbox (-1) 1" |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
216 |
using assms(1) by blast |
53572 | 217 |
ultimately show False |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
218 |
by (smt (verit) "*" iff(1) mem_box_cart(2) nz one_index vector_minus_component) |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
219 |
qed |
53572 | 220 |
qed |
36432 | 221 |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
222 |
proposition fashoda_unit_path: |
53572 | 223 |
fixes f g :: "real \<Rightarrow> real^2" |
224 |
assumes "path f" |
|
225 |
and "path g" |
|
56188 | 226 |
and "path_image f \<subseteq> cbox (-1) 1" |
227 |
and "path_image g \<subseteq> cbox (-1) 1" |
|
53572 | 228 |
and "(pathstart f)$1 = -1" |
229 |
and "(pathfinish f)$1 = 1" |
|
230 |
and "(pathstart g)$2 = -1" |
|
231 |
and "(pathfinish g)$2 = 1" |
|
232 |
obtains z where "z \<in> path_image f" and "z \<in> path_image g" |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
233 |
proof - |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
234 |
note assms = assms[unfolded path_def pathstart_def pathfinish_def path_image_def] |
63040 | 235 |
define iscale where [abs_def]: "iscale z = inverse 2 *\<^sub>R (z + 1)" for z :: real |
53572 | 236 |
have isc: "iscale ` {- 1..1} \<subseteq> {0..1}" |
237 |
unfolding iscale_def by auto |
|
238 |
have "\<exists>s\<in>{- 1..1}. \<exists>t\<in>{- 1..1}. (f \<circ> iscale) s = (g \<circ> iscale) t" |
|
239 |
proof (rule fashoda_unit) |
|
58410
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents:
57418
diff
changeset
|
240 |
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
|
241 |
using isc and assms(3-4) by (auto simp add: image_comp [symmetric]) |
53572 | 242 |
have *: "continuous_on {- 1..1} iscale" |
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56273
diff
changeset
|
243 |
unfolding iscale_def by (rule continuous_intros)+ |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
244 |
show "continuous_on {- 1..1} (f \<circ> iscale)" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
245 |
using "*" assms(1) continuous_on_compose continuous_on_subset isc by blast |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
246 |
show "continuous_on {- 1..1} (g \<circ> iscale)" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
247 |
by (meson "*" assms(2) continuous_on_compose continuous_on_subset isc) |
53572 | 248 |
have *: "(1 / 2) *\<^sub>R (1 + (1::real^1)) = 1" |
249 |
unfolding vec_eq_iff by auto |
|
250 |
show "(f \<circ> iscale) (- 1) $ 1 = - 1" |
|
251 |
and "(f \<circ> iscale) 1 $ 1 = 1" |
|
252 |
and "(g \<circ> iscale) (- 1) $ 2 = -1" |
|
253 |
and "(g \<circ> iscale) 1 $ 2 = 1" |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
254 |
unfolding o_def iscale_def using assms by (auto simp add: *) |
53572 | 255 |
qed |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
256 |
then obtain s t where st: "s \<in> {- 1..1}" "t \<in> {- 1..1}" "(f \<circ> iscale) s = (g \<circ> iscale) t" |
56188 | 257 |
by auto |
53572 | 258 |
show thesis |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
259 |
proof |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
260 |
show "f (iscale s) \<in> path_image f" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
261 |
by (metis image_eqI image_subset_iff isc path_image_def st(1)) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
262 |
show "f (iscale s) \<in> path_image g" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
263 |
by (metis comp_def image_eqI image_subset_iff isc path_image_def st(2) st(3)) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
264 |
qed |
53572 | 265 |
qed |
36432 | 266 |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
267 |
theorem fashoda: |
53627 | 268 |
fixes b :: "real^2" |
269 |
assumes "path f" |
|
270 |
and "path g" |
|
56188 | 271 |
and "path_image f \<subseteq> cbox a b" |
272 |
and "path_image g \<subseteq> cbox a b" |
|
53627 | 273 |
and "(pathstart f)$1 = a$1" |
274 |
and "(pathfinish f)$1 = b$1" |
|
275 |
and "(pathstart g)$2 = a$2" |
|
276 |
and "(pathfinish g)$2 = b$2" |
|
277 |
obtains z where "z \<in> path_image f" and "z \<in> path_image g" |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
278 |
proof - |
53627 | 279 |
fix P Q S |
280 |
presume "P \<or> Q \<or> S" "P \<Longrightarrow> thesis" and "Q \<Longrightarrow> thesis" and "S \<Longrightarrow> thesis" |
|
281 |
then show thesis |
|
282 |
by auto |
|
283 |
next |
|
56188 | 284 |
have "cbox a b \<noteq> {}" |
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
53628
diff
changeset
|
285 |
using assms(3) using path_image_nonempty[of f] by auto |
53627 | 286 |
then have "a \<le> b" |
287 |
unfolding interval_eq_empty_cart less_eq_vec_def by (auto simp add: not_less) |
|
288 |
then show "a$1 = b$1 \<or> a$2 = b$2 \<or> (a$1 < b$1 \<and> a$2 < b$2)" |
|
289 |
unfolding less_eq_vec_def forall_2 by auto |
|
290 |
next |
|
291 |
assume as: "a$1 = b$1" |
|
292 |
have "\<exists>z\<in>path_image g. z$2 = (pathstart f)$2" |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
293 |
proof (rule connected_ivt_component_cart) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
294 |
show "pathstart g $ 2 \<le> pathstart f $ 2" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
295 |
by (metis assms(3) assms(7) mem_box_cart(2) pathstart_in_path_image subset_iff) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
296 |
show "pathstart f $ 2 \<le> pathfinish g $ 2" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
297 |
by (metis assms(3) assms(8) in_mono mem_box_cart(2) pathstart_in_path_image) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
298 |
show "connected (path_image g)" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
299 |
using assms(2) by blast |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
300 |
qed (auto simp: path_defs) |
55675 | 301 |
then obtain z :: "real^2" where z: "z \<in> path_image g" "z $ 2 = pathstart f $ 2" .. |
56188 | 302 |
have "z \<in> cbox a b" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
303 |
using assms(4) z(1) by blast |
53627 | 304 |
then have "z = f 0" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
305 |
by (smt (verit) as assms(5) exhaust_2 mem_box_cart(2) nle_le pathstart_def vec_eq_iff z(2)) |
53627 | 306 |
then show thesis |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
307 |
by (metis path_defs(2) pathstart_in_path_image that z(1)) |
53627 | 308 |
next |
309 |
assume as: "a$2 = b$2" |
|
310 |
have "\<exists>z\<in>path_image f. z$1 = (pathstart g)$1" |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
311 |
proof (rule connected_ivt_component_cart) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
312 |
show "pathstart f $ 1 \<le> pathstart g $ 1" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
313 |
using assms(4) assms(5) mem_box_cart(2) by fastforce |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
314 |
show "pathstart g $ 1 \<le> pathfinish f $ 1" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
315 |
using assms(4) assms(6) mem_box_cart(2) pathstart_in_path_image by fastforce |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
316 |
show "connected (path_image f)" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
317 |
by (simp add: assms(1) connected_path_image) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
318 |
qed (auto simp: path_defs) |
55675 | 319 |
then obtain z where z: "z \<in> path_image f" "z $ 1 = pathstart g $ 1" .. |
56188 | 320 |
have "z \<in> cbox a b" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
321 |
using assms(3) z(1) by auto |
53627 | 322 |
then have "z = g 0" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
323 |
by (smt (verit) as assms(7) exhaust_2 mem_box_cart(2) pathstart_def vec_eq_iff z(2)) |
53627 | 324 |
then show thesis |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
325 |
by (metis path_defs(2) pathstart_in_path_image that z(1)) |
53627 | 326 |
next |
327 |
assume as: "a $ 1 < b $ 1 \<and> a $ 2 < b $ 2" |
|
56188 | 328 |
have int_nem: "cbox (-1) (1::real^2) \<noteq> {}" |
53627 | 329 |
unfolding interval_eq_empty_cart by auto |
55675 | 330 |
obtain z :: "real^2" where z: |
331 |
"z \<in> (interval_bij (a, b) (- 1, 1) \<circ> f) ` {0..1}" |
|
332 |
"z \<in> (interval_bij (a, b) (- 1, 1) \<circ> g) ` {0..1}" |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
333 |
proof (rule fashoda_unit_path) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
334 |
show "path (interval_bij (a, b) (- 1, 1) \<circ> f)" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
335 |
by (meson assms(1) continuous_on_interval_bij path_continuous_image) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
336 |
show "path (interval_bij (a, b) (- 1, 1) \<circ> g)" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
337 |
by (meson assms(2) continuous_on_interval_bij path_continuous_image) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
338 |
show "path_image (interval_bij (a, b) (- 1, 1) \<circ> f) \<subseteq> cbox (- 1) 1" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
339 |
using assms(3) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
340 |
by (simp add: path_image_def in_interval_interval_bij int_nem subset_eq) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
341 |
show "path_image (interval_bij (a, b) (- 1, 1) \<circ> g) \<subseteq> cbox (- 1) 1" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
342 |
using assms(4) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
343 |
by (simp add: path_image_def in_interval_interval_bij int_nem subset_eq) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
344 |
show "pathstart (interval_bij (a, b) (- 1, 1) \<circ> f) $ 1 = - 1" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
345 |
"pathfinish (interval_bij (a, b) (- 1, 1) \<circ> f) $ 1 = 1" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
346 |
"pathstart (interval_bij (a, b) (- 1, 1) \<circ> g) $ 2 = - 1" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
347 |
"pathfinish (interval_bij (a, b) (- 1, 1) \<circ> g) $ 2 = 1" |
56188 | 348 |
using assms as |
67982
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
paulson <lp15@cam.ac.uk>
parents:
67968
diff
changeset
|
349 |
by (simp_all add: cart_eq_inner_axis pathstart_def pathfinish_def interval_bij_def) |
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
|
350 |
(simp_all add: inner_axis) |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
351 |
qed (auto simp: path_defs) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
352 |
then obtain zf zg where zf: "zf \<in> {0..1}" "z = (interval_bij (a, b) (- 1, 1) \<circ> f) zf" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
353 |
and zg: "zg \<in> {0..1}" "z = (interval_bij (a, b) (- 1, 1) \<circ> g) zg" |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
354 |
by blast |
53627 | 355 |
have *: "\<forall>i. (- 1) $ i < (1::real^2) $ i \<and> a $ i < b $ i" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
356 |
unfolding forall_2 using as by auto |
53627 | 357 |
show thesis |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
358 |
proof (rule_tac z="interval_bij (- 1,1) (a,b) z" in that) |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
359 |
show "interval_bij (- 1, 1) (a, b) z \<in> path_image f" |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
360 |
using zf by (simp add: interval_bij_bij_cart[OF *] path_image_def) |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
361 |
show "interval_bij (- 1, 1) (a, b) z \<in> path_image g" |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
362 |
using zg by (simp add: interval_bij_bij_cart[OF *] path_image_def) |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
363 |
qed |
53627 | 364 |
qed |
36432 | 365 |
|
53627 | 366 |
|
70136 | 367 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Some slightly ad hoc lemmas I use below\<close> |
36432 | 368 |
|
53627 | 369 |
lemma segment_vertical: |
370 |
fixes a :: "real^2" |
|
371 |
assumes "a$1 = b$1" |
|
372 |
shows "x \<in> closed_segment a b \<longleftrightarrow> |
|
373 |
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)" |
|
374 |
(is "_ = ?R") |
|
375 |
proof - |
|
36432 | 376 |
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 | 377 |
{ |
378 |
presume "?L \<Longrightarrow> ?R" and "?R \<Longrightarrow> ?L" |
|
379 |
then show ?thesis |
|
380 |
unfolding closed_segment_def mem_Collect_eq |
|
53628 | 381 |
unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[symmetric] vector_component_simps |
53627 | 382 |
by blast |
383 |
} |
|
384 |
{ |
|
385 |
assume ?L |
|
55675 | 386 |
then obtain u where u: |
387 |
"x $ 1 = (1 - u) * a $ 1 + u * b $ 1" |
|
388 |
"x $ 2 = (1 - u) * a $ 2 + u * b $ 2" |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
389 |
"0 \<le> u" "u \<le> 1" |
55675 | 390 |
by blast |
53627 | 391 |
{ fix b a |
392 |
assume "b + u * a > a + u * b" |
|
393 |
then have "(1 - u) * b > (1 - u) * a" |
|
394 |
by (auto simp add:field_simps) |
|
395 |
then have "b \<ge> a" |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
396 |
using not_less_iff_gr_or_eq u(4) by fastforce |
53627 | 397 |
then have "u * a \<le> u * b" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
398 |
by (simp add: mult_left_mono u(3)) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
399 |
} |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
400 |
moreover |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
401 |
{ fix a b |
53627 | 402 |
assume "u * b > u * a" |
403 |
then have "(1 - u) * a \<le> (1 - u) * b" |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
404 |
using less_eq_real_def u(3) u(4) by force |
53627 | 405 |
then have "a + u * b \<le> b + u * a" |
406 |
by (auto simp add: field_simps) |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
407 |
} ultimately show ?R |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
408 |
by (force simp add: u assms field_simps not_le) |
53627 | 409 |
} |
410 |
{ |
|
411 |
assume ?R |
|
412 |
then show ?L |
|
413 |
proof (cases "x$2 = b$2") |
|
414 |
case True |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
415 |
with \<open>?R\<close> show ?L |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
416 |
by (rule_tac x="(x$2 - a$2) / (b$2 - a$2)" in exI) (auto simp add: field_simps) |
53627 | 417 |
next |
418 |
case False |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
419 |
with \<open>?R\<close> show ?L |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
420 |
by (rule_tac x="1 - (x$2 - b$2) / (a$2 - b$2)" in exI) (auto simp add: field_simps) |
53627 | 421 |
qed |
422 |
} |
|
423 |
qed |
|
36432 | 424 |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
425 |
text \<open>Essentially duplicate proof that could be done by swapping co-ordinates\<close> |
53627 | 426 |
lemma segment_horizontal: |
427 |
fixes a :: "real^2" |
|
428 |
assumes "a$2 = b$2" |
|
429 |
shows "x \<in> closed_segment a b \<longleftrightarrow> |
|
430 |
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)" |
|
431 |
(is "_ = ?R") |
|
432 |
proof - |
|
36432 | 433 |
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 | 434 |
{ |
435 |
presume "?L \<Longrightarrow> ?R" and "?R \<Longrightarrow> ?L" |
|
436 |
then show ?thesis |
|
437 |
unfolding closed_segment_def mem_Collect_eq |
|
53628 | 438 |
unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[symmetric] vector_component_simps |
53627 | 439 |
by blast |
440 |
} |
|
441 |
{ |
|
442 |
assume ?L |
|
55675 | 443 |
then obtain u where u: |
444 |
"x $ 1 = (1 - u) * a $ 1 + u * b $ 1" |
|
445 |
"x $ 2 = (1 - u) * a $ 2 + u * b $ 2" |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
446 |
"0 \<le> u" "u \<le> 1" |
55675 | 447 |
by blast |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
448 |
{ fix b a |
53627 | 449 |
assume "b + u * a > a + u * b" |
450 |
then have "(1 - u) * b > (1 - u) * a" |
|
53628 | 451 |
by (auto simp add: field_simps) |
53627 | 452 |
then have "b \<ge> a" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
453 |
by (smt (verit, best) mult_left_mono u(4)) |
53627 | 454 |
then have "u * a \<le> u * b" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
455 |
by (simp add: mult_left_mono u(3)) |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
456 |
} |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
457 |
moreover |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
458 |
{ fix a b |
53627 | 459 |
assume "u * b > u * a" |
460 |
then have "(1 - u) * a \<le> (1 - u) * b" |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
461 |
using less_eq_real_def u(3) u(4) by force |
53627 | 462 |
then have "a + u * b \<le> b + u * a" |
463 |
by (auto simp add: field_simps) |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
464 |
} |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
465 |
ultimately show ?R |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
466 |
by (force simp add: u assms field_simps not_le intro: ) |
53627 | 467 |
} |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
468 |
{ assume ?R |
53627 | 469 |
then show ?L |
470 |
proof (cases "x$1 = b$1") |
|
471 |
case True |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
472 |
with \<open>?R\<close> show ?L |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
473 |
by (rule_tac x="(x$1 - a$1) / (b$1 - a$1)" in exI) (auto simp add: field_simps) |
53627 | 474 |
next |
475 |
case False |
|
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
476 |
with \<open>?R\<close> show ?L |
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
477 |
by (rule_tac x="1 - (x$1 - b$1) / (a$1 - b$1)" in exI) (auto simp add: field_simps) |
53627 | 478 |
qed |
479 |
} |
|
480 |
qed |
|
36432 | 481 |
|
53627 | 482 |
|
69683 | 483 |
subsection \<open>Useful Fashoda corollary pointed out to me by Tom Hales\<close>(*FIXME change title? *) |
36432 | 484 |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
485 |
corollary fashoda_interlace: |
53627 | 486 |
fixes a :: "real^2" |
487 |
assumes "path f" |
|
488 |
and "path g" |
|
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
489 |
and paf: "path_image f \<subseteq> cbox a b" |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
490 |
and pag: "path_image g \<subseteq> cbox a b" |
53627 | 491 |
and "(pathstart f)$2 = a$2" |
492 |
and "(pathfinish f)$2 = a$2" |
|
493 |
and "(pathstart g)$2 = a$2" |
|
494 |
and "(pathfinish g)$2 = a$2" |
|
495 |
and "(pathstart f)$1 < (pathstart g)$1" |
|
496 |
and "(pathstart g)$1 < (pathfinish f)$1" |
|
497 |
and "(pathfinish f)$1 < (pathfinish g)$1" |
|
498 |
obtains z where "z \<in> path_image f" and "z \<in> path_image g" |
|
69681
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents:
69680
diff
changeset
|
499 |
proof - |
56188 | 500 |
have "cbox a b \<noteq> {}" |
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
53628
diff
changeset
|
501 |
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
|
502 |
note ab=this[unfolded interval_eq_empty_cart not_ex forall_2 not_less] |
56188 | 503 |
have "pathstart f \<in> cbox a b" |
504 |
and "pathfinish f \<in> cbox a b" |
|
505 |
and "pathstart g \<in> cbox a b" |
|
506 |
and "pathfinish g \<in> cbox a b" |
|
53628 | 507 |
using pathstart_in_path_image pathfinish_in_path_image |
508 |
using assms(3-4) |
|
509 |
by auto |
|
67673
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
510 |
note startfin = this[unfolded mem_box_cart forall_2] |
36432 | 511 |
let ?P1 = "linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2]) +++ |
512 |
linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f) +++ f +++ |
|
513 |
linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2]) +++ |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63040
diff
changeset
|
514 |
linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2])" |
36432 | 515 |
let ?P2 = "linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g) +++ g +++ |
516 |
linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1]) +++ |
|
517 |
linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1]) +++ |
|
518 |
linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3])" |
|
519 |
let ?a = "vector[a$1 - 2, a$2 - 3]" |
|
520 |
let ?b = "vector[b$1 + 2, b$2 + 3]" |
|
53627 | 521 |
have P1P2: "path_image ?P1 = path_image (linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2])) \<union> |
36432 | 522 |
path_image (linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f)) \<union> path_image f \<union> |
523 |
path_image (linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2])) \<union> |
|
524 |
path_image (linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2]))" |
|
525 |
"path_image ?P2 = path_image(linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g)) \<union> path_image g \<union> |
|
526 |
path_image(linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1])) \<union> |
|
527 |
path_image(linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1])) \<union> |
|
528 |
path_image(linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3]))" using assms(1-2) |
|
71633 | 529 |
by(auto simp add: path_image_join) |
56188 | 530 |
have abab: "cbox a b \<subseteq> cbox ?a ?b" |
531 |
unfolding interval_cbox_cart[symmetric] |
|
71633 | 532 |
by (auto simp add:less_eq_vec_def forall_2) |
55675 | 533 |
obtain z where |
534 |
"z \<in> path_image |
|
535 |
(linepath (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) +++ |
|
536 |
linepath (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f) +++ |
|
537 |
f +++ |
|
538 |
linepath (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) +++ |
|
539 |
linepath (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]))" |
|
540 |
"z \<in> path_image |
|
541 |
(linepath (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g) +++ |
|
542 |
g +++ |
|
543 |
linepath (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1]) +++ |
|
544 |
linepath (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1]) +++ |
|
545 |
linepath (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]))" |
|
53627 | 546 |
apply (rule fashoda[of ?P1 ?P2 ?a ?b]) |
547 |
unfolding pathstart_join pathfinish_join pathstart_linepath pathfinish_linepath vector_2 |
|
548 |
proof - |
|
53628 | 549 |
show "path ?P1" and "path ?P2" |
53627 | 550 |
using assms by auto |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
551 |
show "path_image ?P1 \<subseteq> cbox ?a ?b" "path_image ?P2 \<subseteq> cbox ?a ?b" |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
552 |
unfolding P1P2 path_image_linepath using startfin paf pag |
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
553 |
by (auto simp: mem_box_cart segment_horizontal segment_vertical forall_2) |
53627 | 554 |
show "a $ 1 - 2 = a $ 1 - 2" |
555 |
and "b $ 1 + 2 = b $ 1 + 2" |
|
556 |
and "pathstart g $ 2 - 3 = a $ 2 - 3" |
|
557 |
and "b $ 2 + 3 = b $ 2 + 3" |
|
558 |
by (auto simp add: assms) |
|
53628 | 559 |
qed |
560 |
note z=this[unfolded P1P2 path_image_linepath] |
|
53627 | 561 |
show thesis |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
562 |
proof (rule that[of z]) |
36432 | 563 |
have "(z \<in> closed_segment (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) \<or> |
53627 | 564 |
z \<in> closed_segment (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f)) \<or> |
565 |
z \<in> closed_segment (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) \<or> |
|
566 |
z \<in> closed_segment (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]) \<Longrightarrow> |
|
567 |
(((z \<in> closed_segment (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g)) \<or> |
|
568 |
z \<in> closed_segment (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1])) \<or> |
|
569 |
z \<in> closed_segment (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1])) \<or> |
|
570 |
z \<in> closed_segment (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]) \<Longrightarrow> False" |
|
61166
5976fe402824
renamed method "goals" to "goal_cases" to emphasize its meaning;
wenzelm
parents:
61165
diff
changeset
|
571 |
proof (simp only: segment_vertical segment_horizontal vector_2, goal_cases) |
61167 | 572 |
case prems: 1 |
56188 | 573 |
have "pathfinish f \<in> cbox a b" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63040
diff
changeset
|
574 |
using assms(3) pathfinish_in_path_image[of f] by auto |
53628 | 575 |
then have "1 + b $ 1 \<le> pathfinish f $ 1 \<Longrightarrow> False" |
67673
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
576 |
unfolding mem_box_cart forall_2 by auto |
53627 | 577 |
then have "z$1 \<noteq> pathfinish f$1" |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
578 |
using assms(10) assms(11) prems(2) by auto |
56188 | 579 |
moreover have "pathstart f \<in> cbox a b" |
53628 | 580 |
using assms(3) pathstart_in_path_image[of f] |
581 |
by auto |
|
53627 | 582 |
then have "1 + b $ 1 \<le> pathstart f $ 1 \<Longrightarrow> False" |
67673
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
583 |
unfolding mem_box_cart forall_2 |
53628 | 584 |
by auto |
53627 | 585 |
then have "z$1 \<noteq> pathstart f$1" |
61167 | 586 |
using prems(2) using assms ab |
53628 | 587 |
by (auto simp add: field_simps) |
53627 | 588 |
ultimately have *: "z$2 = a$2 - 2" |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
589 |
using prems(1) by auto |
53627 | 590 |
have "z$1 \<noteq> pathfinish g$1" |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
591 |
using prems(2) assms ab |
53628 | 592 |
by (auto simp add: field_simps *) |
56188 | 593 |
moreover have "pathstart g \<in> cbox a b" |
53628 | 594 |
using assms(4) pathstart_in_path_image[of g] |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63040
diff
changeset
|
595 |
by auto |
67673
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
596 |
note this[unfolded mem_box_cart forall_2] |
53627 | 597 |
then have "z$1 \<noteq> pathstart g$1" |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
598 |
using prems(1) assms ab |
53628 | 599 |
by (auto simp add: field_simps *) |
36432 | 600 |
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" |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
601 |
using prems(2) unfolding * assms by (auto simp add: field_simps) |
53627 | 602 |
then show False |
603 |
unfolding * using ab by auto |
|
604 |
qed |
|
605 |
then have "z \<in> path_image f \<or> z \<in> path_image g" |
|
606 |
using z unfolding Un_iff by blast |
|
56188 | 607 |
then have z': "z \<in> cbox a b" |
68054
ebd179b82e20
getting rid of more "defer", etc.
paulson <lp15@cam.ac.uk>
parents:
68004
diff
changeset
|
608 |
using assms(3-4) by auto |
53627 | 609 |
have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart f $ 1 \<or> z $ 1 = pathfinish f $ 1) \<Longrightarrow> |
610 |
z = pathstart f \<or> z = pathfinish f" |
|
53628 | 611 |
unfolding vec_eq_iff forall_2 assms |
612 |
by auto |
|
53627 | 613 |
with z' show "z \<in> path_image f" |
614 |
using z(1) |
|
67673
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
615 |
unfolding Un_iff mem_box_cart forall_2 |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
616 |
using assms(5) assms(6) segment_horizontal segment_vertical by auto |
53627 | 617 |
have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart g $ 1 \<or> z $ 1 = pathfinish g $ 1) \<Longrightarrow> |
618 |
z = pathstart g \<or> z = pathfinish g" |
|
53628 | 619 |
unfolding vec_eq_iff forall_2 assms |
620 |
by auto |
|
53627 | 621 |
with z' show "z \<in> path_image g" |
622 |
using z(2) |
|
67673
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
623 |
unfolding Un_iff mem_box_cart forall_2 |
78456
57f5127d2ff2
partly tidied some truly horrible proofs
paulson <lp15@cam.ac.uk>
parents:
78248
diff
changeset
|
624 |
using assms(7) assms(8) segment_horizontal segment_vertical by auto |
53627 | 625 |
qed |
626 |
qed |
|
36432 | 627 |
|
628 |
end |