author | wenzelm |
Wed, 23 Jan 2019 23:12:40 +0100 | |
changeset 69729 | 4591221824f6 |
parent 69712 | dc85b5b3a532 |
child 69738 | c558fef62915 |
permissions | -rw-r--r-- |
63938 | 1 |
(* Author: L C Paulson, University of Cambridge |
33175 | 2 |
Author: Amine Chaieb, University of Cambridge |
3 |
Author: Robert Himmelmann, TU Muenchen |
|
44075 | 4 |
Author: Brian Huffman, Portland State University |
33175 | 5 |
*) |
6 |
||
69676 | 7 |
chapter \<open>Vector Analysis\<close> |
33175 | 8 |
|
9 |
theory Topology_Euclidean_Space |
|
69516
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents:
69508
diff
changeset
|
10 |
imports |
69544
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
immler
parents:
69516
diff
changeset
|
11 |
Elementary_Normed_Spaces |
69516
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents:
69508
diff
changeset
|
12 |
Linear_Algebra |
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents:
69508
diff
changeset
|
13 |
Norm_Arith |
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
14 |
begin |
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
15 |
|
69676 | 16 |
section \<open>Elementary Topology in Euclidean Space\<close> |
17 |
||
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
18 |
lemma euclidean_dist_l2: |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
19 |
fixes x y :: "'a :: euclidean_space" |
67155 | 20 |
shows "dist x y = L2_set (\<lambda>i. dist (x \<bullet> i) (y \<bullet> i)) Basis" |
21 |
unfolding dist_norm norm_eq_sqrt_inner L2_set_def |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
22 |
by (subst euclidean_inner) (simp add: power2_eq_square inner_diff_left) |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
23 |
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
24 |
lemma norm_nth_le: "norm (x \<bullet> i) \<le> norm x" if "i \<in> Basis" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
25 |
proof - |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
26 |
have "(x \<bullet> i)\<^sup>2 = (\<Sum>i\<in>{i}. (x \<bullet> i)\<^sup>2)" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
27 |
by simp |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
28 |
also have "\<dots> \<le> (\<Sum>i\<in>Basis. (x \<bullet> i)\<^sup>2)" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
29 |
by (intro sum_mono2) (auto simp: that) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
30 |
finally show ?thesis |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
31 |
unfolding norm_conv_dist euclidean_dist_l2[of x] L2_set_def |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
32 |
by (auto intro!: real_le_rsqrt) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
33 |
qed |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
34 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
35 |
|
69613
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
36 |
subsection%unimportant\<open>Balls in Euclidean Space\<close> |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
37 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
38 |
lemma cball_subset_cball_iff: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
39 |
fixes a :: "'a :: euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
40 |
shows "cball a r \<subseteq> cball a' r' \<longleftrightarrow> dist a a' + r \<le> r' \<or> r < 0" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
41 |
(is "?lhs \<longleftrightarrow> ?rhs") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
42 |
proof |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
43 |
assume ?lhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
44 |
then show ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
45 |
proof (cases "r < 0") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
46 |
case True |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
47 |
then show ?rhs by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
48 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
49 |
case False |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
50 |
then have [simp]: "r \<ge> 0" by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
51 |
have "norm (a - a') + r \<le> r'" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
52 |
proof (cases "a = a'") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
53 |
case True |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
54 |
then show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
55 |
using subsetD [where c = "a + r *\<^sub>R (SOME i. i \<in> Basis)", OF \<open>?lhs\<close>] subsetD [where c = a, OF \<open>?lhs\<close>] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
56 |
by (force simp: SOME_Basis dist_norm) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
57 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
58 |
case False |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
59 |
have "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = norm (a' - a - (r / norm (a - a')) *\<^sub>R (a - a'))" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
60 |
by (simp add: algebra_simps) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
61 |
also have "... = norm ((-1 - (r / norm (a - a'))) *\<^sub>R (a - a'))" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
62 |
by (simp add: algebra_simps) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
63 |
also from \<open>a \<noteq> a'\<close> have "... = \<bar>- norm (a - a') - r\<bar>" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
64 |
by (simp add: abs_mult_pos field_simps) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
65 |
finally have [simp]: "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = \<bar>norm (a - a') + r\<bar>" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
66 |
by linarith |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
67 |
from \<open>a \<noteq> a'\<close> show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
68 |
using subsetD [where c = "a' + (1 + r / norm(a - a')) *\<^sub>R (a - a')", OF \<open>?lhs\<close>] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
69 |
by (simp add: dist_norm scaleR_add_left) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
70 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
71 |
then show ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
72 |
by (simp add: dist_norm) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
73 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
74 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
75 |
assume ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
76 |
then show ?lhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
77 |
by (auto simp: ball_def dist_norm) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
78 |
(metis add.commute add_le_cancel_right dist_norm dist_triangle3 order_trans) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
79 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
80 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
81 |
lemma cball_subset_ball_iff: "cball a r \<subseteq> ball a' r' \<longleftrightarrow> dist a a' + r < r' \<or> r < 0" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
82 |
(is "?lhs \<longleftrightarrow> ?rhs") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
83 |
for a :: "'a::euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
84 |
proof |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
85 |
assume ?lhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
86 |
then show ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
87 |
proof (cases "r < 0") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
88 |
case True then |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
89 |
show ?rhs by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
90 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
91 |
case False |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
92 |
then have [simp]: "r \<ge> 0" by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
93 |
have "norm (a - a') + r < r'" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
94 |
proof (cases "a = a'") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
95 |
case True |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
96 |
then show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
97 |
using subsetD [where c = "a + r *\<^sub>R (SOME i. i \<in> Basis)", OF \<open>?lhs\<close>] subsetD [where c = a, OF \<open>?lhs\<close>] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
98 |
by (force simp: SOME_Basis dist_norm) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
99 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
100 |
case False |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
101 |
have False if "norm (a - a') + r \<ge> r'" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
102 |
proof - |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
103 |
from that have "\<bar>r' - norm (a - a')\<bar> \<le> r" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
104 |
by (simp split: abs_split) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
105 |
(metis \<open>0 \<le> r\<close> \<open>?lhs\<close> centre_in_cball dist_commute dist_norm less_asym mem_ball subset_eq) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
106 |
then show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
107 |
using subsetD [where c = "a + (r' / norm(a - a') - 1) *\<^sub>R (a - a')", OF \<open>?lhs\<close>] \<open>a \<noteq> a'\<close> |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
108 |
by (simp add: dist_norm field_simps) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
109 |
(simp add: diff_divide_distrib scaleR_left_diff_distrib) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
110 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
111 |
then show ?thesis by force |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
112 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
113 |
then show ?rhs by (simp add: dist_norm) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
114 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
115 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
116 |
assume ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
117 |
then show ?lhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
118 |
by (auto simp: ball_def dist_norm) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
119 |
(metis add.commute add_le_cancel_right dist_norm dist_triangle3 le_less_trans) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
120 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
121 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
122 |
lemma ball_subset_cball_iff: "ball a r \<subseteq> cball a' r' \<longleftrightarrow> dist a a' + r \<le> r' \<or> r \<le> 0" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
123 |
(is "?lhs = ?rhs") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
124 |
for a :: "'a::euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
125 |
proof (cases "r \<le> 0") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
126 |
case True |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
127 |
then show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
128 |
using dist_not_less_zero less_le_trans by force |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
129 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
130 |
case False |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
131 |
show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
132 |
proof |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
133 |
assume ?lhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
134 |
then have "(cball a r \<subseteq> cball a' r')" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
135 |
by (metis False closed_cball closure_ball closure_closed closure_mono not_less) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
136 |
with False show ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
137 |
by (fastforce iff: cball_subset_cball_iff) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
138 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
139 |
assume ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
140 |
with False show ?lhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
141 |
using ball_subset_cball cball_subset_cball_iff by blast |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
142 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
143 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
144 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
145 |
lemma ball_subset_ball_iff: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
146 |
fixes a :: "'a :: euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
147 |
shows "ball a r \<subseteq> ball a' r' \<longleftrightarrow> dist a a' + r \<le> r' \<or> r \<le> 0" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
148 |
(is "?lhs = ?rhs") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
149 |
proof (cases "r \<le> 0") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
150 |
case True then show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
151 |
using dist_not_less_zero less_le_trans by force |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
152 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
153 |
case False show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
154 |
proof |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
155 |
assume ?lhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
156 |
then have "0 < r'" |
69712 | 157 |
by (metis (no_types) False \<open>?lhs\<close> centre_in_ball dist_norm le_less_trans mem_ball norm_ge_zero not_less subsetD) |
69613
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
158 |
then have "(cball a r \<subseteq> cball a' r')" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
159 |
by (metis False\<open>?lhs\<close> closure_ball closure_mono not_less) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
160 |
then show ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
161 |
using False cball_subset_cball_iff by fastforce |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
162 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
163 |
assume ?rhs then show ?lhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
164 |
apply (auto simp: ball_def) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
165 |
apply (metis add.commute add_le_cancel_right dist_commute dist_triangle_lt not_le order_trans) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
166 |
using dist_not_less_zero order.strict_trans2 apply blast |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
167 |
done |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
168 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
169 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
170 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
171 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
172 |
lemma ball_eq_ball_iff: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
173 |
fixes x :: "'a :: euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
174 |
shows "ball x d = ball y e \<longleftrightarrow> d \<le> 0 \<and> e \<le> 0 \<or> x=y \<and> d=e" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
175 |
(is "?lhs = ?rhs") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
176 |
proof |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
177 |
assume ?lhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
178 |
then show ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
179 |
proof (cases "d \<le> 0 \<or> e \<le> 0") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
180 |
case True |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
181 |
with \<open>?lhs\<close> show ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
182 |
by safe (simp_all only: ball_eq_empty [of y e, symmetric] ball_eq_empty [of x d, symmetric]) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
183 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
184 |
case False |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
185 |
with \<open>?lhs\<close> show ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
186 |
apply (auto simp: set_eq_subset ball_subset_ball_iff dist_norm norm_minus_commute algebra_simps) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
187 |
apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
188 |
apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
189 |
done |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
190 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
191 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
192 |
assume ?rhs then show ?lhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
193 |
by (auto simp: set_eq_subset ball_subset_ball_iff) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
194 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
195 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
196 |
lemma cball_eq_cball_iff: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
197 |
fixes x :: "'a :: euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
198 |
shows "cball x d = cball y e \<longleftrightarrow> d < 0 \<and> e < 0 \<or> x=y \<and> d=e" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
199 |
(is "?lhs = ?rhs") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
200 |
proof |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
201 |
assume ?lhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
202 |
then show ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
203 |
proof (cases "d < 0 \<or> e < 0") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
204 |
case True |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
205 |
with \<open>?lhs\<close> show ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
206 |
by safe (simp_all only: cball_eq_empty [of y e, symmetric] cball_eq_empty [of x d, symmetric]) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
207 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
208 |
case False |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
209 |
with \<open>?lhs\<close> show ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
210 |
apply (auto simp: set_eq_subset cball_subset_cball_iff dist_norm norm_minus_commute algebra_simps) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
211 |
apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
212 |
apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
213 |
done |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
214 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
215 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
216 |
assume ?rhs then show ?lhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
217 |
by (auto simp: set_eq_subset cball_subset_cball_iff) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
218 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
219 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
220 |
lemma ball_eq_cball_iff: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
221 |
fixes x :: "'a :: euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
222 |
shows "ball x d = cball y e \<longleftrightarrow> d \<le> 0 \<and> e < 0" (is "?lhs = ?rhs") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
223 |
proof |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
224 |
assume ?lhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
225 |
then show ?rhs |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
226 |
apply (auto simp: set_eq_subset ball_subset_cball_iff cball_subset_ball_iff algebra_simps) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
227 |
apply (metis add_increasing2 add_le_cancel_right add_less_same_cancel1 dist_not_less_zero less_le_trans zero_le_dist) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
228 |
apply (metis add_less_same_cancel1 dist_not_less_zero less_le_trans not_le) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
229 |
using \<open>?lhs\<close> ball_eq_empty cball_eq_empty apply blast+ |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
230 |
done |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
231 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
232 |
assume ?rhs then show ?lhs by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
233 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
234 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
235 |
lemma cball_eq_ball_iff: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
236 |
fixes x :: "'a :: euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
237 |
shows "cball x d = ball y e \<longleftrightarrow> d < 0 \<and> e \<le> 0" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
238 |
using ball_eq_cball_iff by blast |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
239 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
240 |
lemma finite_ball_avoid: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
241 |
fixes S :: "'a :: euclidean_space set" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
242 |
assumes "open S" "finite X" "p \<in> S" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
243 |
shows "\<exists>e>0. \<forall>w\<in>ball p e. w\<in>S \<and> (w\<noteq>p \<longrightarrow> w\<notin>X)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
244 |
proof - |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
245 |
obtain e1 where "0 < e1" and e1_b:"ball p e1 \<subseteq> S" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
246 |
using open_contains_ball_eq[OF \<open>open S\<close>] assms by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
247 |
obtain e2 where "0 < e2" and "\<forall>x\<in>X. x \<noteq> p \<longrightarrow> e2 \<le> dist p x" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
248 |
using finite_set_avoid[OF \<open>finite X\<close>,of p] by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
249 |
hence "\<forall>w\<in>ball p (min e1 e2). w\<in>S \<and> (w\<noteq>p \<longrightarrow> w\<notin>X)" using e1_b by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
250 |
thus "\<exists>e>0. \<forall>w\<in>ball p e. w \<in> S \<and> (w \<noteq> p \<longrightarrow> w \<notin> X)" using \<open>e2>0\<close> \<open>e1>0\<close> |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
251 |
apply (rule_tac x="min e1 e2" in exI) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
252 |
by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
253 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
254 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
255 |
lemma finite_cball_avoid: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
256 |
fixes S :: "'a :: euclidean_space set" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
257 |
assumes "open S" "finite X" "p \<in> S" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
258 |
shows "\<exists>e>0. \<forall>w\<in>cball p e. w\<in>S \<and> (w\<noteq>p \<longrightarrow> w\<notin>X)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
259 |
proof - |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
260 |
obtain e1 where "e1>0" and e1: "\<forall>w\<in>ball p e1. w\<in>S \<and> (w\<noteq>p \<longrightarrow> w\<notin>X)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
261 |
using finite_ball_avoid[OF assms] by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
262 |
define e2 where "e2 \<equiv> e1/2" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
263 |
have "e2>0" and "e2 < e1" unfolding e2_def using \<open>e1>0\<close> by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
264 |
then have "cball p e2 \<subseteq> ball p e1" by (subst cball_subset_ball_iff,auto) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
265 |
then show "\<exists>e>0. \<forall>w\<in>cball p e. w \<in> S \<and> (w \<noteq> p \<longrightarrow> w \<notin> X)" using \<open>e2>0\<close> e1 by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
266 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
267 |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
268 |
lemma dim_cball: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
269 |
assumes "e > 0" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
270 |
shows "dim (cball (0 :: 'n::euclidean_space) e) = DIM('n)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
271 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
272 |
{ |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
273 |
fix x :: "'n::euclidean_space" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
274 |
define y where "y = (e / norm x) *\<^sub>R x" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
275 |
then have "y \<in> cball 0 e" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
276 |
using assms by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
277 |
moreover have *: "x = (norm x / e) *\<^sub>R y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
278 |
using y_def assms by simp |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
279 |
moreover from * have "x = (norm x/e) *\<^sub>R y" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
280 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
281 |
ultimately have "x \<in> span (cball 0 e)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
282 |
using span_scale[of y "cball 0 e" "norm x/e"] |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
283 |
span_superset[of "cball 0 e"] |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
284 |
by (simp add: span_base) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
285 |
} |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
286 |
then have "span (cball 0 e) = (UNIV :: 'n::euclidean_space set)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
287 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
288 |
then show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
289 |
using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto simp: dim_UNIV) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
290 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
291 |
|
69613
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
292 |
|
60420 | 293 |
subsection \<open>Boxes\<close> |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
294 |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
295 |
abbreviation One :: "'a::euclidean_space" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
296 |
where "One \<equiv> \<Sum>Basis" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
297 |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
298 |
lemma One_non_0: assumes "One = (0::'a::euclidean_space)" shows False |
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
299 |
proof - |
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
300 |
have "dependent (Basis :: 'a set)" |
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
301 |
apply (simp add: dependent_finite) |
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
302 |
apply (rule_tac x="\<lambda>i. 1" in exI) |
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
303 |
using SOME_Basis apply (auto simp: assms) |
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
304 |
done |
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
305 |
with independent_Basis show False by force |
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
306 |
qed |
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
307 |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
308 |
corollary One_neq_0[iff]: "One \<noteq> 0" |
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
309 |
by (metis One_non_0) |
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
310 |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
311 |
corollary Zero_neq_One[iff]: "0 \<noteq> One" |
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
312 |
by (metis One_non_0) |
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
313 |
|
67962 | 314 |
definition%important (in euclidean_space) eucl_less (infix "<e" 50) |
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
315 |
where "eucl_less a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i < b \<bullet> i)" |
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
316 |
|
67962 | 317 |
definition%important box_eucl_less: "box a b = {x. a <e x \<and> x <e b}" |
318 |
definition%important "cbox a b = {x. \<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i}" |
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
319 |
|
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
320 |
lemma box_def: "box a b = {x. \<forall>i\<in>Basis. a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i}" |
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
321 |
and in_box_eucl_less: "x \<in> box a b \<longleftrightarrow> a <e x \<and> x <e b" |
56188 | 322 |
and mem_box: "x \<in> box a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i)" |
323 |
"x \<in> cbox a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i)" |
|
324 |
by (auto simp: box_eucl_less eucl_less_def cbox_def) |
|
325 |
||
60615
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
326 |
lemma cbox_Pair_eq: "cbox (a, c) (b, d) = cbox a b \<times> cbox c d" |
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
327 |
by (force simp: cbox_def Basis_prod_def) |
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
328 |
|
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
329 |
lemma cbox_Pair_iff [iff]: "(x, y) \<in> cbox (a, c) (b, d) \<longleftrightarrow> x \<in> cbox a b \<and> y \<in> cbox c d" |
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
330 |
by (force simp: cbox_Pair_eq) |
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
331 |
|
65587
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
332 |
lemma cbox_Complex_eq: "cbox (Complex a c) (Complex b d) = (\<lambda>(x,y). Complex x y) ` (cbox a b \<times> cbox c d)" |
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
333 |
apply (auto simp: cbox_def Basis_complex_def) |
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
334 |
apply (rule_tac x = "(Re x, Im x)" in image_eqI) |
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
335 |
using complex_eq by auto |
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
336 |
|
60615
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
337 |
lemma cbox_Pair_eq_0: "cbox (a, c) (b, d) = {} \<longleftrightarrow> cbox a b = {} \<or> cbox c d = {}" |
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
338 |
by (force simp: cbox_Pair_eq) |
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
339 |
|
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
340 |
lemma swap_cbox_Pair [simp]: "prod.swap ` cbox (c, a) (d, b) = cbox (a,c) (b,d)" |
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
341 |
by auto |
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
342 |
|
56188 | 343 |
lemma mem_box_real[simp]: |
344 |
"(x::real) \<in> box a b \<longleftrightarrow> a < x \<and> x < b" |
|
345 |
"(x::real) \<in> cbox a b \<longleftrightarrow> a \<le> x \<and> x \<le> b" |
|
346 |
by (auto simp: mem_box) |
|
347 |
||
348 |
lemma box_real[simp]: |
|
349 |
fixes a b:: real |
|
350 |
shows "box a b = {a <..< b}" "cbox a b = {a .. b}" |
|
351 |
by auto |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
352 |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
353 |
lemma box_Int_box: |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
354 |
fixes a :: "'a::euclidean_space" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
355 |
shows "box a b \<inter> box c d = |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
356 |
box (\<Sum>i\<in>Basis. max (a\<bullet>i) (c\<bullet>i) *\<^sub>R i) (\<Sum>i\<in>Basis. min (b\<bullet>i) (d\<bullet>i) *\<^sub>R i)" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
357 |
unfolding set_eq_iff and Int_iff and mem_box by auto |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
358 |
|
50087 | 359 |
lemma rational_boxes: |
61076 | 360 |
fixes x :: "'a::euclidean_space" |
53291 | 361 |
assumes "e > 0" |
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
362 |
shows "\<exists>a b. (\<forall>i\<in>Basis. a \<bullet> i \<in> \<rat> \<and> b \<bullet> i \<in> \<rat>) \<and> x \<in> box a b \<and> box a b \<subseteq> ball x e" |
50087 | 363 |
proof - |
63040 | 364 |
define e' where "e' = e / (2 * sqrt (real (DIM ('a))))" |
53291 | 365 |
then have e: "e' > 0" |
56541 | 366 |
using assms by (auto simp: DIM_positive) |
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
367 |
have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x \<bullet> i \<and> x \<bullet> i - y < e'" (is "\<forall>i. ?th i") |
50087 | 368 |
proof |
53255 | 369 |
fix i |
370 |
from Rats_dense_in_real[of "x \<bullet> i - e'" "x \<bullet> i"] e |
|
371 |
show "?th i" by auto |
|
50087 | 372 |
qed |
55522 | 373 |
from choice[OF this] obtain a where |
374 |
a: "\<forall>xa. a xa \<in> \<rat> \<and> a xa < x \<bullet> xa \<and> x \<bullet> xa - a xa < e'" .. |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
375 |
have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x \<bullet> i < y \<and> y - x \<bullet> i < e'" (is "\<forall>i. ?th i") |
50087 | 376 |
proof |
53255 | 377 |
fix i |
378 |
from Rats_dense_in_real[of "x \<bullet> i" "x \<bullet> i + e'"] e |
|
379 |
show "?th i" by auto |
|
50087 | 380 |
qed |
55522 | 381 |
from choice[OF this] obtain b where |
382 |
b: "\<forall>xa. b xa \<in> \<rat> \<and> x \<bullet> xa < b xa \<and> b xa - x \<bullet> xa < e'" .. |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
383 |
let ?a = "\<Sum>i\<in>Basis. a i *\<^sub>R i" and ?b = "\<Sum>i\<in>Basis. b i *\<^sub>R i" |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
384 |
show ?thesis |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
385 |
proof (rule exI[of _ ?a], rule exI[of _ ?b], safe) |
53255 | 386 |
fix y :: 'a |
387 |
assume *: "y \<in> box ?a ?b" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52625
diff
changeset
|
388 |
have "dist x y = sqrt (\<Sum>i\<in>Basis. (dist (x \<bullet> i) (y \<bullet> i))\<^sup>2)" |
67155 | 389 |
unfolding L2_set_def[symmetric] by (rule euclidean_dist_l2) |
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
390 |
also have "\<dots> < sqrt (\<Sum>(i::'a)\<in>Basis. e^2 / real (DIM('a)))" |
64267 | 391 |
proof (rule real_sqrt_less_mono, rule sum_strict_mono) |
53255 | 392 |
fix i :: "'a" |
393 |
assume i: "i \<in> Basis" |
|
394 |
have "a i < y\<bullet>i \<and> y\<bullet>i < b i" |
|
395 |
using * i by (auto simp: box_def) |
|
396 |
moreover have "a i < x\<bullet>i" "x\<bullet>i - a i < e'" |
|
397 |
using a by auto |
|
398 |
moreover have "x\<bullet>i < b i" "b i - x\<bullet>i < e'" |
|
399 |
using b by auto |
|
400 |
ultimately have "\<bar>x\<bullet>i - y\<bullet>i\<bar> < 2 * e'" |
|
401 |
by auto |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
402 |
then have "dist (x \<bullet> i) (y \<bullet> i) < e/sqrt (real (DIM('a)))" |
50087 | 403 |
unfolding e'_def by (auto simp: dist_real_def) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52625
diff
changeset
|
404 |
then have "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2" |
50087 | 405 |
by (rule power_strict_mono) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52625
diff
changeset
|
406 |
then show "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < e\<^sup>2 / real DIM('a)" |
50087 | 407 |
by (simp add: power_divide) |
408 |
qed auto |
|
53255 | 409 |
also have "\<dots> = e" |
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
410 |
using \<open>0 < e\<close> by simp |
53255 | 411 |
finally show "y \<in> ball x e" |
412 |
by (auto simp: ball_def) |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
413 |
qed (insert a b, auto simp: box_def) |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
414 |
qed |
51103 | 415 |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
416 |
lemma open_UNION_box: |
61076 | 417 |
fixes M :: "'a::euclidean_space set" |
53282 | 418 |
assumes "open M" |
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
419 |
defines "a' \<equiv> \<lambda>f :: 'a \<Rightarrow> real \<times> real. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)" |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
420 |
defines "b' \<equiv> \<lambda>f :: 'a \<Rightarrow> real \<times> real. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52625
diff
changeset
|
421 |
defines "I \<equiv> {f\<in>Basis \<rightarrow>\<^sub>E \<rat> \<times> \<rat>. box (a' f) (b' f) \<subseteq> M}" |
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
422 |
shows "M = (\<Union>f\<in>I. box (a' f) (b' f))" |
52624 | 423 |
proof - |
60462 | 424 |
have "x \<in> (\<Union>f\<in>I. box (a' f) (b' f))" if "x \<in> M" for x |
425 |
proof - |
|
52624 | 426 |
obtain e where e: "e > 0" "ball x e \<subseteq> M" |
60420 | 427 |
using openE[OF \<open>open M\<close> \<open>x \<in> M\<close>] by auto |
53282 | 428 |
moreover obtain a b where ab: |
429 |
"x \<in> box a b" |
|
430 |
"\<forall>i \<in> Basis. a \<bullet> i \<in> \<rat>" |
|
431 |
"\<forall>i\<in>Basis. b \<bullet> i \<in> \<rat>" |
|
432 |
"box a b \<subseteq> ball x e" |
|
52624 | 433 |
using rational_boxes[OF e(1)] by metis |
60462 | 434 |
ultimately show ?thesis |
52624 | 435 |
by (intro UN_I[of "\<lambda>i\<in>Basis. (a \<bullet> i, b \<bullet> i)"]) |
436 |
(auto simp: euclidean_representation I_def a'_def b'_def) |
|
60462 | 437 |
qed |
52624 | 438 |
then show ?thesis by (auto simp: I_def) |
439 |
qed |
|
440 |
||
66154
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
441 |
corollary open_countable_Union_open_box: |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
442 |
fixes S :: "'a :: euclidean_space set" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
443 |
assumes "open S" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
444 |
obtains \<D> where "countable \<D>" "\<D> \<subseteq> Pow S" "\<And>X. X \<in> \<D> \<Longrightarrow> \<exists>a b. X = box a b" "\<Union>\<D> = S" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
445 |
proof - |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
446 |
let ?a = "\<lambda>f. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
447 |
let ?b = "\<lambda>f. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
448 |
let ?I = "{f\<in>Basis \<rightarrow>\<^sub>E \<rat> \<times> \<rat>. box (?a f) (?b f) \<subseteq> S}" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
449 |
let ?\<D> = "(\<lambda>f. box (?a f) (?b f)) ` ?I" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
450 |
show ?thesis |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
451 |
proof |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
452 |
have "countable ?I" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
453 |
by (simp add: countable_PiE countable_rat) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
454 |
then show "countable ?\<D>" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
455 |
by blast |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
456 |
show "\<Union>?\<D> = S" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
457 |
using open_UNION_box [OF assms] by metis |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
458 |
qed auto |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
459 |
qed |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
460 |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
461 |
lemma rational_cboxes: |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
462 |
fixes x :: "'a::euclidean_space" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
463 |
assumes "e > 0" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
464 |
shows "\<exists>a b. (\<forall>i\<in>Basis. a \<bullet> i \<in> \<rat> \<and> b \<bullet> i \<in> \<rat>) \<and> x \<in> cbox a b \<and> cbox a b \<subseteq> ball x e" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
465 |
proof - |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
466 |
define e' where "e' = e / (2 * sqrt (real (DIM ('a))))" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
467 |
then have e: "e' > 0" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
468 |
using assms by auto |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
469 |
have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x \<bullet> i \<and> x \<bullet> i - y < e'" (is "\<forall>i. ?th i") |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
470 |
proof |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
471 |
fix i |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
472 |
from Rats_dense_in_real[of "x \<bullet> i - e'" "x \<bullet> i"] e |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
473 |
show "?th i" by auto |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
474 |
qed |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
475 |
from choice[OF this] obtain a where |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
476 |
a: "\<forall>u. a u \<in> \<rat> \<and> a u < x \<bullet> u \<and> x \<bullet> u - a u < e'" .. |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
477 |
have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x \<bullet> i < y \<and> y - x \<bullet> i < e'" (is "\<forall>i. ?th i") |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
478 |
proof |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
479 |
fix i |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
480 |
from Rats_dense_in_real[of "x \<bullet> i" "x \<bullet> i + e'"] e |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
481 |
show "?th i" by auto |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
482 |
qed |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
483 |
from choice[OF this] obtain b where |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
484 |
b: "\<forall>u. b u \<in> \<rat> \<and> x \<bullet> u < b u \<and> b u - x \<bullet> u < e'" .. |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
485 |
let ?a = "\<Sum>i\<in>Basis. a i *\<^sub>R i" and ?b = "\<Sum>i\<in>Basis. b i *\<^sub>R i" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
486 |
show ?thesis |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
487 |
proof (rule exI[of _ ?a], rule exI[of _ ?b], safe) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
488 |
fix y :: 'a |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
489 |
assume *: "y \<in> cbox ?a ?b" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
490 |
have "dist x y = sqrt (\<Sum>i\<in>Basis. (dist (x \<bullet> i) (y \<bullet> i))\<^sup>2)" |
67155 | 491 |
unfolding L2_set_def[symmetric] by (rule euclidean_dist_l2) |
66154
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
492 |
also have "\<dots> < sqrt (\<Sum>(i::'a)\<in>Basis. e^2 / real (DIM('a)))" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
493 |
proof (rule real_sqrt_less_mono, rule sum_strict_mono) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
494 |
fix i :: "'a" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
495 |
assume i: "i \<in> Basis" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
496 |
have "a i \<le> y\<bullet>i \<and> y\<bullet>i \<le> b i" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
497 |
using * i by (auto simp: cbox_def) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
498 |
moreover have "a i < x\<bullet>i" "x\<bullet>i - a i < e'" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
499 |
using a by auto |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
500 |
moreover have "x\<bullet>i < b i" "b i - x\<bullet>i < e'" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
501 |
using b by auto |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
502 |
ultimately have "\<bar>x\<bullet>i - y\<bullet>i\<bar> < 2 * e'" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
503 |
by auto |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
504 |
then have "dist (x \<bullet> i) (y \<bullet> i) < e/sqrt (real (DIM('a)))" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
505 |
unfolding e'_def by (auto simp: dist_real_def) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
506 |
then have "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
507 |
by (rule power_strict_mono) auto |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
508 |
then show "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < e\<^sup>2 / real DIM('a)" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
509 |
by (simp add: power_divide) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
510 |
qed auto |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
511 |
also have "\<dots> = e" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
512 |
using \<open>0 < e\<close> by simp |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
513 |
finally show "y \<in> ball x e" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
514 |
by (auto simp: ball_def) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
515 |
next |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
516 |
show "x \<in> cbox (\<Sum>i\<in>Basis. a i *\<^sub>R i) (\<Sum>i\<in>Basis. b i *\<^sub>R i)" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
517 |
using a b less_imp_le by (auto simp: cbox_def) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
518 |
qed (use a b cbox_def in auto) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
519 |
qed |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
520 |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
521 |
lemma open_UNION_cbox: |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
522 |
fixes M :: "'a::euclidean_space set" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
523 |
assumes "open M" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
524 |
defines "a' \<equiv> \<lambda>f. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
525 |
defines "b' \<equiv> \<lambda>f. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
526 |
defines "I \<equiv> {f\<in>Basis \<rightarrow>\<^sub>E \<rat> \<times> \<rat>. cbox (a' f) (b' f) \<subseteq> M}" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
527 |
shows "M = (\<Union>f\<in>I. cbox (a' f) (b' f))" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
528 |
proof - |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
529 |
have "x \<in> (\<Union>f\<in>I. cbox (a' f) (b' f))" if "x \<in> M" for x |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
530 |
proof - |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
531 |
obtain e where e: "e > 0" "ball x e \<subseteq> M" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
532 |
using openE[OF \<open>open M\<close> \<open>x \<in> M\<close>] by auto |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
533 |
moreover obtain a b where ab: "x \<in> cbox a b" "\<forall>i \<in> Basis. a \<bullet> i \<in> \<rat>" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
534 |
"\<forall>i \<in> Basis. b \<bullet> i \<in> \<rat>" "cbox a b \<subseteq> ball x e" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
535 |
using rational_cboxes[OF e(1)] by metis |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
536 |
ultimately show ?thesis |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
537 |
by (intro UN_I[of "\<lambda>i\<in>Basis. (a \<bullet> i, b \<bullet> i)"]) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
538 |
(auto simp: euclidean_representation I_def a'_def b'_def) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
539 |
qed |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
540 |
then show ?thesis by (auto simp: I_def) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
541 |
qed |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
542 |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
543 |
corollary open_countable_Union_open_cbox: |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
544 |
fixes S :: "'a :: euclidean_space set" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
545 |
assumes "open S" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
546 |
obtains \<D> where "countable \<D>" "\<D> \<subseteq> Pow S" "\<And>X. X \<in> \<D> \<Longrightarrow> \<exists>a b. X = cbox a b" "\<Union>\<D> = S" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
547 |
proof - |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
548 |
let ?a = "\<lambda>f. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
549 |
let ?b = "\<lambda>f. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
550 |
let ?I = "{f\<in>Basis \<rightarrow>\<^sub>E \<rat> \<times> \<rat>. cbox (?a f) (?b f) \<subseteq> S}" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
551 |
let ?\<D> = "(\<lambda>f. cbox (?a f) (?b f)) ` ?I" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
552 |
show ?thesis |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
553 |
proof |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
554 |
have "countable ?I" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
555 |
by (simp add: countable_PiE countable_rat) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
556 |
then show "countable ?\<D>" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
557 |
by blast |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
558 |
show "\<Union>?\<D> = S" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
559 |
using open_UNION_cbox [OF assms] by metis |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
560 |
qed auto |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
561 |
qed |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
562 |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
563 |
lemma box_eq_empty: |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
564 |
fixes a :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
565 |
shows "(box a b = {} \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i \<le> a\<bullet>i))" (is ?th1) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
566 |
and "(cbox a b = {} \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i < a\<bullet>i))" (is ?th2) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
567 |
proof - |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
568 |
{ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
569 |
fix i x |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
570 |
assume i: "i\<in>Basis" and as:"b\<bullet>i \<le> a\<bullet>i" and x:"x\<in>box a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
571 |
then have "a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
572 |
unfolding mem_box by (auto simp: box_def) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
573 |
then have "a\<bullet>i < b\<bullet>i" by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
574 |
then have False using as by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
575 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
576 |
moreover |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
577 |
{ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
578 |
assume as: "\<forall>i\<in>Basis. \<not> (b\<bullet>i \<le> a\<bullet>i)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
579 |
let ?x = "(1/2) *\<^sub>R (a + b)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
580 |
{ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
581 |
fix i :: 'a |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
582 |
assume i: "i \<in> Basis" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
583 |
have "a\<bullet>i < b\<bullet>i" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
584 |
using as[THEN bspec[where x=i]] i by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
585 |
then have "a\<bullet>i < ((1/2) *\<^sub>R (a+b)) \<bullet> i" "((1/2) *\<^sub>R (a+b)) \<bullet> i < b\<bullet>i" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
586 |
by (auto simp: inner_add_left) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
587 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
588 |
then have "box a b \<noteq> {}" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
589 |
using mem_box(1)[of "?x" a b] by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
590 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
591 |
ultimately show ?th1 by blast |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
592 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
593 |
{ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
594 |
fix i x |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
595 |
assume i: "i \<in> Basis" and as:"b\<bullet>i < a\<bullet>i" and x:"x\<in>cbox a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
596 |
then have "a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
597 |
unfolding mem_box by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
598 |
then have "a\<bullet>i \<le> b\<bullet>i" by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
599 |
then have False using as by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
600 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
601 |
moreover |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
602 |
{ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
603 |
assume as:"\<forall>i\<in>Basis. \<not> (b\<bullet>i < a\<bullet>i)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
604 |
let ?x = "(1/2) *\<^sub>R (a + b)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
605 |
{ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
606 |
fix i :: 'a |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
607 |
assume i:"i \<in> Basis" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
608 |
have "a\<bullet>i \<le> b\<bullet>i" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
609 |
using as[THEN bspec[where x=i]] i by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
610 |
then have "a\<bullet>i \<le> ((1/2) *\<^sub>R (a+b)) \<bullet> i" "((1/2) *\<^sub>R (a+b)) \<bullet> i \<le> b\<bullet>i" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
611 |
by (auto simp: inner_add_left) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
612 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
613 |
then have "cbox a b \<noteq> {}" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
614 |
using mem_box(2)[of "?x" a b] by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
615 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
616 |
ultimately show ?th2 by blast |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
617 |
qed |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
618 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
619 |
lemma box_ne_empty: |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
620 |
fixes a :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
621 |
shows "cbox a b \<noteq> {} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
622 |
and "box a b \<noteq> {} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
623 |
unfolding box_eq_empty[of a b] by fastforce+ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
624 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
625 |
lemma |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
626 |
fixes a :: "'a::euclidean_space" |
66112
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents:
66089
diff
changeset
|
627 |
shows cbox_sing [simp]: "cbox a a = {a}" |
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents:
66089
diff
changeset
|
628 |
and box_sing [simp]: "box a a = {}" |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
629 |
unfolding set_eq_iff mem_box eq_iff [symmetric] |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
630 |
by (auto intro!: euclidean_eqI[where 'a='a]) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
631 |
(metis all_not_in_conv nonempty_Basis) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
632 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
633 |
lemma subset_box_imp: |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
634 |
fixes a :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
635 |
shows "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> cbox c d \<subseteq> cbox a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
636 |
and "(\<forall>i\<in>Basis. a\<bullet>i < c\<bullet>i \<and> d\<bullet>i < b\<bullet>i) \<Longrightarrow> cbox c d \<subseteq> box a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
637 |
and "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> box c d \<subseteq> cbox a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
638 |
and "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> box c d \<subseteq> box a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
639 |
unfolding subset_eq[unfolded Ball_def] unfolding mem_box |
58757 | 640 |
by (best intro: order_trans less_le_trans le_less_trans less_imp_le)+ |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
641 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
642 |
lemma box_subset_cbox: |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
643 |
fixes a :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
644 |
shows "box a b \<subseteq> cbox a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
645 |
unfolding subset_eq [unfolded Ball_def] mem_box |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
646 |
by (fast intro: less_imp_le) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
647 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
648 |
lemma subset_box: |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
649 |
fixes a :: "'a::euclidean_space" |
64539 | 650 |
shows "cbox c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i \<le> d\<bullet>i) \<longrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th1) |
651 |
and "cbox c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i \<le> d\<bullet>i) \<longrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < c\<bullet>i \<and> d\<bullet>i < b\<bullet>i)" (is ?th2) |
|
652 |
and "box c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i) \<longrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th3) |
|
653 |
and "box c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i) \<longrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th4) |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
654 |
proof - |
68120 | 655 |
let ?lesscd = "\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i" |
656 |
let ?lerhs = "\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" |
|
657 |
show ?th1 ?th2 |
|
658 |
by (fastforce simp: mem_box)+ |
|
659 |
have acdb: "a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" |
|
660 |
if i: "i \<in> Basis" and box: "box c d \<subseteq> cbox a b" and cd: "\<And>i. i \<in> Basis \<Longrightarrow> c\<bullet>i < d\<bullet>i" for i |
|
661 |
proof - |
|
662 |
have "box c d \<noteq> {}" |
|
663 |
using that |
|
664 |
unfolding box_eq_empty by force |
|
665 |
{ let ?x = "(\<Sum>j\<in>Basis. (if j=i then ((min (a\<bullet>j) (d\<bullet>j))+c\<bullet>j)/2 else (c\<bullet>j+d\<bullet>j)/2) *\<^sub>R j)::'a" |
|
666 |
assume *: "a\<bullet>i > c\<bullet>i" |
|
667 |
then have "c \<bullet> j < ?x \<bullet> j \<and> ?x \<bullet> j < d \<bullet> j" if "j \<in> Basis" for j |
|
668 |
using cd that by (fastforce simp add: i *) |
|
669 |
then have "?x \<in> box c d" |
|
670 |
unfolding mem_box by auto |
|
671 |
moreover have "?x \<notin> cbox a b" |
|
672 |
using i cd * by (force simp: mem_box) |
|
673 |
ultimately have False using box by auto |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
674 |
} |
68120 | 675 |
then have "a\<bullet>i \<le> c\<bullet>i" by force |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
676 |
moreover |
68120 | 677 |
{ let ?x = "(\<Sum>j\<in>Basis. (if j=i then ((max (b\<bullet>j) (c\<bullet>j))+d\<bullet>j)/2 else (c\<bullet>j+d\<bullet>j)/2) *\<^sub>R j)::'a" |
678 |
assume *: "b\<bullet>i < d\<bullet>i" |
|
679 |
then have "d \<bullet> j > ?x \<bullet> j \<and> ?x \<bullet> j > c \<bullet> j" if "j \<in> Basis" for j |
|
680 |
using cd that by (fastforce simp add: i *) |
|
681 |
then have "?x \<in> box c d" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
682 |
unfolding mem_box by auto |
68120 | 683 |
moreover have "?x \<notin> cbox a b" |
684 |
using i cd * by (force simp: mem_box) |
|
685 |
ultimately have False using box by auto |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
686 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
687 |
then have "b\<bullet>i \<ge> d\<bullet>i" by (rule ccontr) auto |
68120 | 688 |
ultimately show ?thesis by auto |
689 |
qed |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
690 |
show ?th3 |
68120 | 691 |
using acdb by (fastforce simp add: mem_box) |
692 |
have acdb': "a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" |
|
693 |
if "i \<in> Basis" "box c d \<subseteq> box a b" "\<And>i. i \<in> Basis \<Longrightarrow> c\<bullet>i < d\<bullet>i" for i |
|
694 |
using box_subset_cbox[of a b] that acdb by auto |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
695 |
show ?th4 |
68120 | 696 |
using acdb' by (fastforce simp add: mem_box) |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
697 |
qed |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
698 |
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
699 |
lemma eq_cbox: "cbox a b = cbox c d \<longleftrightarrow> cbox a b = {} \<and> cbox c d = {} \<or> a = c \<and> b = d" |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
700 |
(is "?lhs = ?rhs") |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
701 |
proof |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
702 |
assume ?lhs |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
703 |
then have "cbox a b \<subseteq> cbox c d" "cbox c d \<subseteq> cbox a b" |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
704 |
by auto |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
705 |
then show ?rhs |
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
706 |
by (force simp: subset_box box_eq_empty intro: antisym euclidean_eqI) |
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
707 |
next |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
708 |
assume ?rhs |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
709 |
then show ?lhs |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
710 |
by force |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
711 |
qed |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
712 |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
713 |
lemma eq_cbox_box [simp]: "cbox a b = box c d \<longleftrightarrow> cbox a b = {} \<and> box c d = {}" |
64539 | 714 |
(is "?lhs \<longleftrightarrow> ?rhs") |
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
715 |
proof |
68120 | 716 |
assume L: ?lhs |
717 |
then have "cbox a b \<subseteq> box c d" "box c d \<subseteq> cbox a b" |
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
718 |
by auto |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
719 |
then show ?rhs |
63957
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
720 |
apply (simp add: subset_box) |
68120 | 721 |
using L box_ne_empty box_sing apply (fastforce simp add:) |
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
722 |
done |
68120 | 723 |
qed force |
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
724 |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
725 |
lemma eq_box_cbox [simp]: "box a b = cbox c d \<longleftrightarrow> box a b = {} \<and> cbox c d = {}" |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
726 |
by (metis eq_cbox_box) |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
727 |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
728 |
lemma eq_box: "box a b = box c d \<longleftrightarrow> box a b = {} \<and> box c d = {} \<or> a = c \<and> b = d" |
64539 | 729 |
(is "?lhs \<longleftrightarrow> ?rhs") |
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
730 |
proof |
68120 | 731 |
assume L: ?lhs |
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
732 |
then have "box a b \<subseteq> box c d" "box c d \<subseteq> box a b" |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
733 |
by auto |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
734 |
then show ?rhs |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
735 |
apply (simp add: subset_box) |
68120 | 736 |
using box_ne_empty(2) L |
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
737 |
apply auto |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
738 |
apply (meson euclidean_eqI less_eq_real_def not_less)+ |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
739 |
done |
68120 | 740 |
qed force |
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
741 |
|
66466
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
742 |
lemma subset_box_complex: |
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
743 |
"cbox a b \<subseteq> cbox c d \<longleftrightarrow> |
66466
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
744 |
(Re a \<le> Re b \<and> Im a \<le> Im b) \<longrightarrow> Re a \<ge> Re c \<and> Im a \<ge> Im c \<and> Re b \<le> Re d \<and> Im b \<le> Im d" |
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
745 |
"cbox a b \<subseteq> box c d \<longleftrightarrow> |
66466
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
746 |
(Re a \<le> Re b \<and> Im a \<le> Im b) \<longrightarrow> Re a > Re c \<and> Im a > Im c \<and> Re b < Re d \<and> Im b < Im d" |
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
747 |
"box a b \<subseteq> cbox c d \<longleftrightarrow> |
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
748 |
(Re a < Re b \<and> Im a < Im b) \<longrightarrow> Re a \<ge> Re c \<and> Im a \<ge> Im c \<and> Re b \<le> Re d \<and> Im b \<le> Im d" |
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
749 |
"box a b \<subseteq> box c d \<longleftrightarrow> |
66466
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
750 |
(Re a < Re b \<and> Im a < Im b) \<longrightarrow> Re a \<ge> Re c \<and> Im a \<ge> Im c \<and> Re b \<le> Re d \<and> Im b \<le> Im d" |
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
751 |
by (subst subset_box; force simp: Basis_complex_def)+ |
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
752 |
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
753 |
lemma Int_interval: |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
754 |
fixes a :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
755 |
shows "cbox a b \<inter> cbox c d = |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
756 |
cbox (\<Sum>i\<in>Basis. max (a\<bullet>i) (c\<bullet>i) *\<^sub>R i) (\<Sum>i\<in>Basis. min (b\<bullet>i) (d\<bullet>i) *\<^sub>R i)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
757 |
unfolding set_eq_iff and Int_iff and mem_box |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
758 |
by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
759 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
760 |
lemma disjoint_interval: |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
761 |
fixes a::"'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
762 |
shows "cbox a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i < a\<bullet>i \<or> d\<bullet>i < c\<bullet>i \<or> b\<bullet>i < c\<bullet>i \<or> d\<bullet>i < a\<bullet>i))" (is ?th1) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
763 |
and "cbox a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i < a\<bullet>i \<or> d\<bullet>i \<le> c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th2) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
764 |
and "box a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i \<le> a\<bullet>i \<or> d\<bullet>i < c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th3) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
765 |
and "box a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i \<le> a\<bullet>i \<or> d\<bullet>i \<le> c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th4) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
766 |
proof - |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
767 |
let ?z = "(\<Sum>i\<in>Basis. (((max (a\<bullet>i) (c\<bullet>i)) + (min (b\<bullet>i) (d\<bullet>i))) / 2) *\<^sub>R i)::'a" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
768 |
have **: "\<And>P Q. (\<And>i :: 'a. i \<in> Basis \<Longrightarrow> Q ?z i \<Longrightarrow> P i) \<Longrightarrow> |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
769 |
(\<And>i x :: 'a. i \<in> Basis \<Longrightarrow> P i \<Longrightarrow> Q x i) \<Longrightarrow> (\<forall>x. \<exists>i\<in>Basis. Q x i) \<longleftrightarrow> (\<exists>i\<in>Basis. P i)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
770 |
by blast |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
771 |
note * = set_eq_iff Int_iff empty_iff mem_box ball_conj_distrib[symmetric] eq_False ball_simps(10) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
772 |
show ?th1 unfolding * by (intro **) auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
773 |
show ?th2 unfolding * by (intro **) auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
774 |
show ?th3 unfolding * by (intro **) auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
775 |
show ?th4 unfolding * by (intro **) auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
776 |
qed |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
777 |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
778 |
lemma UN_box_eq_UNIV: "(\<Union>i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One)) = UNIV" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
779 |
proof - |
61942 | 780 |
have "\<bar>x \<bullet> b\<bar> < real_of_int (\<lceil>Max ((\<lambda>b. \<bar>x \<bullet> b\<bar>)`Basis)\<rceil> + 1)" |
60462 | 781 |
if [simp]: "b \<in> Basis" for x b :: 'a |
782 |
proof - |
|
61942 | 783 |
have "\<bar>x \<bullet> b\<bar> \<le> real_of_int \<lceil>\<bar>x \<bullet> b\<bar>\<rceil>" |
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
784 |
by (rule le_of_int_ceiling) |
61942 | 785 |
also have "\<dots> \<le> real_of_int \<lceil>Max ((\<lambda>b. \<bar>x \<bullet> b\<bar>)`Basis)\<rceil>" |
59587
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
nipkow
parents:
58877
diff
changeset
|
786 |
by (auto intro!: ceiling_mono) |
61942 | 787 |
also have "\<dots> < real_of_int (\<lceil>Max ((\<lambda>b. \<bar>x \<bullet> b\<bar>)`Basis)\<rceil> + 1)" |
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
788 |
by simp |
60462 | 789 |
finally show ?thesis . |
790 |
qed |
|
791 |
then have "\<exists>n::nat. \<forall>b\<in>Basis. \<bar>x \<bullet> b\<bar> < real n" for x :: 'a |
|
59587
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
nipkow
parents:
58877
diff
changeset
|
792 |
by (metis order.strict_trans reals_Archimedean2) |
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
793 |
moreover have "\<And>x b::'a. \<And>n::nat. \<bar>x \<bullet> b\<bar> < real n \<longleftrightarrow> - real n < x \<bullet> b \<and> x \<bullet> b < real n" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
794 |
by auto |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
795 |
ultimately show ?thesis |
64267 | 796 |
by (auto simp: box_def inner_sum_left inner_Basis sum.If_cases) |
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
797 |
qed |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
798 |
|
69613
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
799 |
lemma image_affinity_cbox: fixes m::real |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
800 |
fixes a b c :: "'a::euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
801 |
shows "(\<lambda>x. m *\<^sub>R x + c) ` cbox a b = |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
802 |
(if cbox a b = {} then {} |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
803 |
else (if 0 \<le> m then cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
804 |
else cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c)))" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
805 |
proof (cases "m = 0") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
806 |
case True |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
807 |
{ |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
808 |
fix x |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
809 |
assume "\<forall>i\<in>Basis. x \<bullet> i \<le> c \<bullet> i" "\<forall>i\<in>Basis. c \<bullet> i \<le> x \<bullet> i" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
810 |
then have "x = c" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
811 |
by (simp add: dual_order.antisym euclidean_eqI) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
812 |
} |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
813 |
moreover have "c \<in> cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
814 |
unfolding True by (auto simp: cbox_sing) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
815 |
ultimately show ?thesis using True by (auto simp: cbox_def) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
816 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
817 |
case False |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
818 |
{ |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
819 |
fix y |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
820 |
assume "\<forall>i\<in>Basis. a \<bullet> i \<le> y \<bullet> i" "\<forall>i\<in>Basis. y \<bullet> i \<le> b \<bullet> i" "m > 0" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
821 |
then have "\<forall>i\<in>Basis. (m *\<^sub>R a + c) \<bullet> i \<le> (m *\<^sub>R y + c) \<bullet> i" and "\<forall>i\<in>Basis. (m *\<^sub>R y + c) \<bullet> i \<le> (m *\<^sub>R b + c) \<bullet> i" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
822 |
by (auto simp: inner_distrib) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
823 |
} |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
824 |
moreover |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
825 |
{ |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
826 |
fix y |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
827 |
assume "\<forall>i\<in>Basis. a \<bullet> i \<le> y \<bullet> i" "\<forall>i\<in>Basis. y \<bullet> i \<le> b \<bullet> i" "m < 0" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
828 |
then have "\<forall>i\<in>Basis. (m *\<^sub>R b + c) \<bullet> i \<le> (m *\<^sub>R y + c) \<bullet> i" and "\<forall>i\<in>Basis. (m *\<^sub>R y + c) \<bullet> i \<le> (m *\<^sub>R a + c) \<bullet> i" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
829 |
by (auto simp: mult_left_mono_neg inner_distrib) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
830 |
} |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
831 |
moreover |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
832 |
{ |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
833 |
fix y |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
834 |
assume "m > 0" and "\<forall>i\<in>Basis. (m *\<^sub>R a + c) \<bullet> i \<le> y \<bullet> i" and "\<forall>i\<in>Basis. y \<bullet> i \<le> (m *\<^sub>R b + c) \<bullet> i" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
835 |
then have "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` cbox a b" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
836 |
unfolding image_iff Bex_def mem_box |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
837 |
apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
838 |
apply (auto simp: pos_le_divide_eq pos_divide_le_eq mult.commute inner_distrib inner_diff_left) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
839 |
done |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
840 |
} |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
841 |
moreover |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
842 |
{ |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
843 |
fix y |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
844 |
assume "\<forall>i\<in>Basis. (m *\<^sub>R b + c) \<bullet> i \<le> y \<bullet> i" "\<forall>i\<in>Basis. y \<bullet> i \<le> (m *\<^sub>R a + c) \<bullet> i" "m < 0" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
845 |
then have "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` cbox a b" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
846 |
unfolding image_iff Bex_def mem_box |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
847 |
apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
848 |
apply (auto simp: neg_le_divide_eq neg_divide_le_eq mult.commute inner_distrib inner_diff_left) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
849 |
done |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
850 |
} |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
851 |
ultimately show ?thesis using False by (auto simp: cbox_def) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
852 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
853 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
854 |
lemma image_smult_cbox:"(\<lambda>x. m *\<^sub>R (x::_::euclidean_space)) ` cbox a b = |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
855 |
(if cbox a b = {} then {} else if 0 \<le> m then cbox (m *\<^sub>R a) (m *\<^sub>R b) else cbox (m *\<^sub>R b) (m *\<^sub>R a))" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
856 |
using image_affinity_cbox[of m 0 a b] by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
857 |
|
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
858 |
lemma swap_continuous: |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
859 |
assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
860 |
shows "continuous_on (cbox (c,a) (d,b)) (\<lambda>(x, y). f y x)" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
861 |
proof - |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
862 |
have "(\<lambda>(x, y). f y x) = (\<lambda>(x, y). f x y) \<circ> prod.swap" |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
863 |
by auto |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
864 |
then show ?thesis |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
865 |
apply (rule ssubst) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
866 |
apply (rule continuous_on_compose) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
867 |
apply (simp add: split_def) |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
868 |
apply (rule continuous_intros | simp add: assms)+ |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
869 |
done |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
870 |
qed |
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
871 |
|
69516
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents:
69508
diff
changeset
|
872 |
|
69508 | 873 |
subsection \<open>General Intervals\<close> |
67962 | 874 |
|
875 |
definition%important "is_interval (s::('a::euclidean_space) set) \<longleftrightarrow> |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
876 |
(\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i\<in>Basis. ((a\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> b\<bullet>i) \<or> (b\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> a\<bullet>i))) \<longrightarrow> x \<in> s)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
877 |
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
878 |
lemma is_interval_1: |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
879 |
"is_interval (s::real set) \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. \<forall> x. a \<le> x \<and> x \<le> b \<longrightarrow> x \<in> s)" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
880 |
unfolding is_interval_def by auto |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
881 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
882 |
lemma is_interval_inter: "is_interval X \<Longrightarrow> is_interval Y \<Longrightarrow> is_interval (X \<inter> Y)" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
883 |
unfolding is_interval_def |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
884 |
by blast |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
885 |
|
66089
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
886 |
lemma is_interval_cbox [simp]: "is_interval (cbox a (b::'a::euclidean_space))" (is ?th1) |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
887 |
and is_interval_box [simp]: "is_interval (box a b)" (is ?th2) |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
888 |
unfolding is_interval_def mem_box Ball_def atLeastAtMost_iff |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
889 |
by (meson order_trans le_less_trans less_le_trans less_trans)+ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
890 |
|
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
891 |
lemma is_interval_empty [iff]: "is_interval {}" |
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
892 |
unfolding is_interval_def by simp |
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
893 |
|
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
894 |
lemma is_interval_univ [iff]: "is_interval UNIV" |
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
895 |
unfolding is_interval_def by simp |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
896 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
897 |
lemma mem_is_intervalI: |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
898 |
assumes "is_interval s" |
64539 | 899 |
and "a \<in> s" "b \<in> s" |
900 |
and "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i \<or> b \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> a \<bullet> i" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
901 |
shows "x \<in> s" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
902 |
by (rule assms(1)[simplified is_interval_def, rule_format, OF assms(2,3,4)]) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
903 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
904 |
lemma interval_subst: |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
905 |
fixes S::"'a::euclidean_space set" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
906 |
assumes "is_interval S" |
64539 | 907 |
and "x \<in> S" "y j \<in> S" |
908 |
and "j \<in> Basis" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
909 |
shows "(\<Sum>i\<in>Basis. (if i = j then y i \<bullet> i else x \<bullet> i) *\<^sub>R i) \<in> S" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
910 |
by (rule mem_is_intervalI[OF assms(1,2)]) (auto simp: assms) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
911 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
912 |
lemma mem_box_componentwiseI: |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
913 |
fixes S::"'a::euclidean_space set" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
914 |
assumes "is_interval S" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
915 |
assumes "\<And>i. i \<in> Basis \<Longrightarrow> x \<bullet> i \<in> ((\<lambda>x. x \<bullet> i) ` S)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
916 |
shows "x \<in> S" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
917 |
proof - |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
918 |
from assms have "\<forall>i \<in> Basis. \<exists>s \<in> S. x \<bullet> i = s \<bullet> i" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
919 |
by auto |
64539 | 920 |
with finite_Basis obtain s and bs::"'a list" |
921 |
where s: "\<And>i. i \<in> Basis \<Longrightarrow> x \<bullet> i = s i \<bullet> i" "\<And>i. i \<in> Basis \<Longrightarrow> s i \<in> S" |
|
922 |
and bs: "set bs = Basis" "distinct bs" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
923 |
by (metis finite_distinct_list) |
64539 | 924 |
from nonempty_Basis s obtain j where j: "j \<in> Basis" "s j \<in> S" |
925 |
by blast |
|
63040 | 926 |
define y where |
927 |
"y = rec_list (s j) (\<lambda>j _ Y. (\<Sum>i\<in>Basis. (if i = j then s i \<bullet> i else Y \<bullet> i) *\<^sub>R i))" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
928 |
have "x = (\<Sum>i\<in>Basis. (if i \<in> set bs then s i \<bullet> i else s j \<bullet> i) *\<^sub>R i)" |
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
929 |
using bs by (auto simp: s(1)[symmetric] euclidean_representation) |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
930 |
also have [symmetric]: "y bs = \<dots>" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
931 |
using bs(2) bs(1)[THEN equalityD1] |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
932 |
by (induct bs) (auto simp: y_def euclidean_representation intro!: euclidean_eqI[where 'a='a]) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
933 |
also have "y bs \<in> S" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
934 |
using bs(1)[THEN equalityD1] |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
935 |
apply (induct bs) |
64539 | 936 |
apply (auto simp: y_def j) |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
937 |
apply (rule interval_subst[OF assms(1)]) |
64539 | 938 |
apply (auto simp: s) |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
939 |
done |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
940 |
finally show ?thesis . |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
941 |
qed |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
942 |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62948
diff
changeset
|
943 |
lemma cbox01_nonempty [simp]: "cbox 0 One \<noteq> {}" |
64267 | 944 |
by (simp add: box_ne_empty inner_Basis inner_sum_left sum_nonneg) |
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62948
diff
changeset
|
945 |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62948
diff
changeset
|
946 |
lemma box01_nonempty [simp]: "box 0 One \<noteq> {}" |
66089
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
947 |
by (simp add: box_ne_empty inner_Basis inner_sum_left) |
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
948 |
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
949 |
lemma empty_as_interval: "{} = cbox One (0::'a::euclidean_space)" |
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
950 |
using nonempty_Basis box01_nonempty box_eq_empty(1) box_ne_empty(1) by blast |
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
951 |
|
66089
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
952 |
lemma interval_subset_is_interval: |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
953 |
assumes "is_interval S" |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
954 |
shows "cbox a b \<subseteq> S \<longleftrightarrow> cbox a b = {} \<or> a \<in> S \<and> b \<in> S" (is "?lhs = ?rhs") |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
955 |
proof |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
956 |
assume ?lhs |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
957 |
then show ?rhs using box_ne_empty(1) mem_box(2) by fastforce |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
958 |
next |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
959 |
assume ?rhs |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
960 |
have "cbox a b \<subseteq> S" if "a \<in> S" "b \<in> S" |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
961 |
using assms unfolding is_interval_def |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
962 |
apply (clarsimp simp add: mem_box) |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
963 |
using that by blast |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
964 |
with \<open>?rhs\<close> show ?lhs |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
965 |
by blast |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
966 |
qed |
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
967 |
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
968 |
lemma is_real_interval_union: |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
969 |
"is_interval (X \<union> Y)" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
970 |
if X: "is_interval X" and Y: "is_interval Y" and I: "(X \<noteq> {} \<Longrightarrow> Y \<noteq> {} \<Longrightarrow> X \<inter> Y \<noteq> {})" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
971 |
for X Y::"real set" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
972 |
proof - |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
973 |
consider "X \<noteq> {}" "Y \<noteq> {}" | "X = {}" | "Y = {}" by blast |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
974 |
then show ?thesis |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
975 |
proof cases |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
976 |
case 1 |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
977 |
then obtain r where "r \<in> X \<or> X \<inter> Y = {}" "r \<in> Y \<or> X \<inter> Y = {}" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
978 |
by blast |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
979 |
then show ?thesis |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
980 |
using I 1 X Y unfolding is_interval_1 |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
981 |
by (metis (full_types) Un_iff le_cases) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
982 |
qed (use that in auto) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
983 |
qed |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
984 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
985 |
lemma is_interval_translationI: |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
986 |
assumes "is_interval X" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
987 |
shows "is_interval ((+) x ` X)" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
988 |
unfolding is_interval_def |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
989 |
proof safe |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
990 |
fix b d e |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
991 |
assume "b \<in> X" "d \<in> X" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
992 |
"\<forall>i\<in>Basis. (x + b) \<bullet> i \<le> e \<bullet> i \<and> e \<bullet> i \<le> (x + d) \<bullet> i \<or> |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
993 |
(x + d) \<bullet> i \<le> e \<bullet> i \<and> e \<bullet> i \<le> (x + b) \<bullet> i" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
994 |
hence "e - x \<in> X" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
995 |
by (intro mem_is_intervalI[OF assms \<open>b \<in> X\<close> \<open>d \<in> X\<close>, of "e - x"]) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
996 |
(auto simp: algebra_simps) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
997 |
thus "e \<in> (+) x ` X" by force |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
998 |
qed |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
999 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1000 |
lemma is_interval_uminusI: |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1001 |
assumes "is_interval X" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1002 |
shows "is_interval (uminus ` X)" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1003 |
unfolding is_interval_def |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1004 |
proof safe |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1005 |
fix b d e |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1006 |
assume "b \<in> X" "d \<in> X" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1007 |
"\<forall>i\<in>Basis. (- b) \<bullet> i \<le> e \<bullet> i \<and> e \<bullet> i \<le> (- d) \<bullet> i \<or> |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1008 |
(- d) \<bullet> i \<le> e \<bullet> i \<and> e \<bullet> i \<le> (- b) \<bullet> i" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1009 |
hence "- e \<in> X" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1010 |
by (intro mem_is_intervalI[OF assms \<open>b \<in> X\<close> \<open>d \<in> X\<close>, of "- e"]) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1011 |
(auto simp: algebra_simps) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1012 |
thus "e \<in> uminus ` X" by force |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1013 |
qed |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1014 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1015 |
lemma is_interval_uminus[simp]: "is_interval (uminus ` x) = is_interval x" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1016 |
using is_interval_uminusI[of x] is_interval_uminusI[of "uminus ` x"] |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1017 |
by (auto simp: image_image) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1018 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1019 |
lemma is_interval_neg_translationI: |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1020 |
assumes "is_interval X" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1021 |
shows "is_interval ((-) x ` X)" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1022 |
proof - |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1023 |
have "(-) x ` X = (+) x ` uminus ` X" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1024 |
by (force simp: algebra_simps) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1025 |
also have "is_interval \<dots>" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1026 |
by (metis is_interval_uminusI is_interval_translationI assms) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1027 |
finally show ?thesis . |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1028 |
qed |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1029 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1030 |
lemma is_interval_translation[simp]: |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1031 |
"is_interval ((+) x ` X) = is_interval X" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1032 |
using is_interval_neg_translationI[of "(+) x ` X" x] |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1033 |
by (auto intro!: is_interval_translationI simp: image_image) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1034 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1035 |
lemma is_interval_minus_translation[simp]: |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1036 |
shows "is_interval ((-) x ` X) = is_interval X" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1037 |
proof - |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1038 |
have "(-) x ` X = (+) x ` uminus ` X" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1039 |
by (force simp: algebra_simps) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1040 |
also have "is_interval \<dots> = is_interval X" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1041 |
by simp |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1042 |
finally show ?thesis . |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1043 |
qed |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1044 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1045 |
lemma is_interval_minus_translation'[simp]: |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1046 |
shows "is_interval ((\<lambda>x. x - c) ` X) = is_interval X" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1047 |
using is_interval_translation[of "-c" X] |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1048 |
by (metis image_cong uminus_add_conv_diff) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1049 |
|
69611
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1050 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1051 |
subsection%unimportant \<open>Bounded Projections\<close> |
62127 | 1052 |
|
69611
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1053 |
lemma bounded_inner_imp_bdd_above: |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1054 |
assumes "bounded s" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1055 |
shows "bdd_above ((\<lambda>x. x \<bullet> a) ` s)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1056 |
by (simp add: assms bounded_imp_bdd_above bounded_linear_image bounded_linear_inner_left) |
33175 | 1057 |
|
69611
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1058 |
lemma bounded_inner_imp_bdd_below: |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1059 |
assumes "bounded s" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1060 |
shows "bdd_below ((\<lambda>x. x \<bullet> a) ` s)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1061 |
by (simp add: assms bounded_imp_bdd_below bounded_linear_image bounded_linear_inner_left) |
44632 | 1062 |
|
53282 | 1063 |
|
69611
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1064 |
subsection%unimportant \<open>Structural rules for pointwise continuity\<close> |
33175 | 1065 |
|
51361
21e5b6efb317
changed continuous_intros into a named theorems collection
hoelzl
parents:
51351
diff
changeset
|
1066 |
lemma continuous_infnorm[continuous_intros]: |
53282 | 1067 |
"continuous F f \<Longrightarrow> continuous F (\<lambda>x. infnorm (f x))" |
44647
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
1068 |
unfolding continuous_def by (rule tendsto_infnorm) |
33175 | 1069 |
|
51361
21e5b6efb317
changed continuous_intros into a named theorems collection
hoelzl
parents:
51351
diff
changeset
|
1070 |
lemma continuous_inner[continuous_intros]: |
53282 | 1071 |
assumes "continuous F f" |
1072 |
and "continuous F g" |
|
44647
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
1073 |
shows "continuous F (\<lambda>x. inner (f x) (g x))" |
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
1074 |
using assms unfolding continuous_def by (rule tendsto_inner) |
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
1075 |
|
69516
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents:
69508
diff
changeset
|
1076 |
|
69611
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1077 |
subsection%unimportant \<open>Structural rules for setwise continuity\<close> |
33175 | 1078 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
1079 |
lemma continuous_on_infnorm[continuous_intros]: |
53282 | 1080 |
"continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. infnorm (f x))" |
44647
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
1081 |
unfolding continuous_on by (fast intro: tendsto_infnorm) |
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
1082 |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
1083 |
lemma continuous_on_inner[continuous_intros]: |
44531
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44530
diff
changeset
|
1084 |
fixes g :: "'a::topological_space \<Rightarrow> 'b::real_inner" |
53282 | 1085 |
assumes "continuous_on s f" |
1086 |
and "continuous_on s g" |
|
44531
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44530
diff
changeset
|
1087 |
shows "continuous_on s (\<lambda>x. inner (f x) (g x))" |
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44530
diff
changeset
|
1088 |
using bounded_bilinear_inner assms |
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44530
diff
changeset
|
1089 |
by (rule bounded_bilinear.continuous_on) |
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44530
diff
changeset
|
1090 |
|
69613
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1091 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1092 |
subsection%unimportant \<open>Openness of halfspaces.\<close> |
33175 | 1093 |
|
1094 |
lemma open_halfspace_lt: "open {x. inner a x < b}" |
|
63332 | 1095 |
by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) |
33175 | 1096 |
|
1097 |
lemma open_halfspace_gt: "open {x. inner a x > b}" |
|
63332 | 1098 |
by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) |
33175 | 1099 |
|
53282 | 1100 |
lemma open_halfspace_component_lt: "open {x::'a::euclidean_space. x\<bullet>i < a}" |
63332 | 1101 |
by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) |
33175 | 1102 |
|
53282 | 1103 |
lemma open_halfspace_component_gt: "open {x::'a::euclidean_space. x\<bullet>i > a}" |
63332 | 1104 |
by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) |
33175 | 1105 |
|
69611
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1106 |
lemma eucl_less_eq_halfspaces: |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1107 |
fixes a :: "'a::euclidean_space" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1108 |
shows "{x. x <e a} = (\<Inter>i\<in>Basis. {x. x \<bullet> i < a \<bullet> i})" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1109 |
"{x. a <e x} = (\<Inter>i\<in>Basis. {x. a \<bullet> i < x \<bullet> i})" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1110 |
by (auto simp: eucl_less_def) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1111 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1112 |
lemma open_Collect_eucl_less[simp, intro]: |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1113 |
fixes a :: "'a::euclidean_space" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1114 |
shows "open {x. x <e a}" "open {x. a <e x}" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1115 |
by (auto simp: eucl_less_eq_halfspaces open_halfspace_component_lt open_halfspace_component_gt) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1116 |
|
69613
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1117 |
subsection%unimportant \<open>Closure of halfspaces and hyperplanes\<close> |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1118 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1119 |
lemma continuous_at_inner: "continuous (at x) (inner a)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1120 |
unfolding continuous_at by (intro tendsto_intros) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1121 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1122 |
lemma closed_halfspace_le: "closed {x. inner a x \<le> b}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1123 |
by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1124 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1125 |
lemma closed_halfspace_ge: "closed {x. inner a x \<ge> b}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1126 |
by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1127 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1128 |
lemma closed_hyperplane: "closed {x. inner a x = b}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1129 |
by (simp add: closed_Collect_eq continuous_on_inner continuous_on_const continuous_on_id) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1130 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1131 |
lemma closed_halfspace_component_le: "closed {x::'a::euclidean_space. x\<bullet>i \<le> a}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1132 |
by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1133 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1134 |
lemma closed_halfspace_component_ge: "closed {x::'a::euclidean_space. x\<bullet>i \<ge> a}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1135 |
by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1136 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1137 |
lemma closed_interval_left: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1138 |
fixes b :: "'a::euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1139 |
shows "closed {x::'a. \<forall>i\<in>Basis. x\<bullet>i \<le> b\<bullet>i}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1140 |
by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1141 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1142 |
lemma closed_interval_right: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1143 |
fixes a :: "'a::euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1144 |
shows "closed {x::'a. \<forall>i\<in>Basis. a\<bullet>i \<le> x\<bullet>i}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1145 |
by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1146 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1147 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1148 |
subsection%unimportant\<open>Some more convenient intermediate-value theorem formulations\<close> |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1149 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1150 |
lemma connected_ivt_hyperplane: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1151 |
assumes "connected S" and xy: "x \<in> S" "y \<in> S" and b: "inner a x \<le> b" "b \<le> inner a y" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1152 |
shows "\<exists>z \<in> S. inner a z = b" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1153 |
proof (rule ccontr) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1154 |
assume as:"\<not> (\<exists>z\<in>S. inner a z = b)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1155 |
let ?A = "{x. inner a x < b}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1156 |
let ?B = "{x. inner a x > b}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1157 |
have "open ?A" "open ?B" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1158 |
using open_halfspace_lt and open_halfspace_gt by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1159 |
moreover have "?A \<inter> ?B = {}" by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1160 |
moreover have "S \<subseteq> ?A \<union> ?B" using as by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1161 |
ultimately show False |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1162 |
using \<open>connected S\<close>[unfolded connected_def not_ex, |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1163 |
THEN spec[where x="?A"], THEN spec[where x="?B"]] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1164 |
using xy b by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1165 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1166 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1167 |
lemma connected_ivt_component: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1168 |
fixes x::"'a::euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1169 |
shows "connected S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x\<bullet>k \<le> a \<Longrightarrow> a \<le> y\<bullet>k \<Longrightarrow> (\<exists>z\<in>S. z\<bullet>k = a)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1170 |
using connected_ivt_hyperplane[of S x y "k::'a" a] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1171 |
by (auto simp: inner_commute) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1172 |
|
69611
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1173 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1174 |
subsection \<open>Limit Component Bounds\<close> |
33175 | 1175 |
|
69613
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1176 |
lemma Lim_component_le: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1177 |
fixes f :: "'a \<Rightarrow> 'b::euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1178 |
assumes "(f \<longlongrightarrow> l) net" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1179 |
and "\<not> (trivial_limit net)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1180 |
and "eventually (\<lambda>x. f(x)\<bullet>i \<le> b) net" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1181 |
shows "l\<bullet>i \<le> b" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1182 |
by (rule tendsto_le[OF assms(2) tendsto_const tendsto_inner[OF assms(1) tendsto_const] assms(3)]) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1183 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1184 |
lemma Lim_component_ge: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1185 |
fixes f :: "'a \<Rightarrow> 'b::euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1186 |
assumes "(f \<longlongrightarrow> l) net" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1187 |
and "\<not> (trivial_limit net)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1188 |
and "eventually (\<lambda>x. b \<le> (f x)\<bullet>i) net" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1189 |
shows "b \<le> l\<bullet>i" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1190 |
by (rule tendsto_le[OF assms(2) tendsto_inner[OF assms(1) tendsto_const] tendsto_const assms(3)]) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1191 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1192 |
lemma Lim_component_eq: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1193 |
fixes f :: "'a \<Rightarrow> 'b::euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1194 |
assumes net: "(f \<longlongrightarrow> l) net" "\<not> trivial_limit net" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1195 |
and ev:"eventually (\<lambda>x. f(x)\<bullet>i = b) net" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1196 |
shows "l\<bullet>i = b" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1197 |
using ev[unfolded order_eq_iff eventually_conj_iff] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1198 |
using Lim_component_ge[OF net, of b i] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1199 |
using Lim_component_le[OF net, of i b] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1200 |
by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1201 |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1202 |
lemma open_box[intro]: "open (box a b)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1203 |
proof - |
67399 | 1204 |
have "open (\<Inter>i\<in>Basis. ((\<bullet>) i) -` {a \<bullet> i <..< b \<bullet> i})" |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1205 |
by (auto intro!: continuous_open_vimage continuous_inner continuous_ident continuous_const) |
67399 | 1206 |
also have "(\<Inter>i\<in>Basis. ((\<bullet>) i) -` {a \<bullet> i <..< b \<bullet> i}) = box a b" |
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1207 |
by (auto simp: box_def inner_commute) |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1208 |
finally show ?thesis . |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1209 |
qed |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1210 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1211 |
lemma closed_cbox[intro]: |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1212 |
fixes a b :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1213 |
shows "closed (cbox a b)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1214 |
proof - |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1215 |
have "closed (\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i .. b\<bullet>i})" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1216 |
by (intro closed_INT ballI continuous_closed_vimage allI |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1217 |
linear_continuous_at closed_real_atLeastAtMost finite_Basis bounded_linear_inner_left) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1218 |
also have "(\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i .. b\<bullet>i}) = cbox a b" |
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1219 |
by (auto simp: cbox_def) |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1220 |
finally show "closed (cbox a b)" . |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1221 |
qed |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1222 |
|
62618
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1223 |
lemma interior_cbox [simp]: |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1224 |
fixes a b :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1225 |
shows "interior (cbox a b) = box a b" (is "?L = ?R") |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1226 |
proof(rule subset_antisym) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1227 |
show "?R \<subseteq> ?L" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1228 |
using box_subset_cbox open_box |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1229 |
by (rule interior_maximal) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1230 |
{ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1231 |
fix x |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1232 |
assume "x \<in> interior (cbox a b)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1233 |
then obtain s where s: "open s" "x \<in> s" "s \<subseteq> cbox a b" .. |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1234 |
then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> cbox a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1235 |
unfolding open_dist and subset_eq by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1236 |
{ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1237 |
fix i :: 'a |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1238 |
assume i: "i \<in> Basis" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1239 |
have "dist (x - (e / 2) *\<^sub>R i) x < e" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1240 |
and "dist (x + (e / 2) *\<^sub>R i) x < e" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1241 |
unfolding dist_norm |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1242 |
apply auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1243 |
unfolding norm_minus_cancel |
60420 | 1244 |
using norm_Basis[OF i] \<open>e>0\<close> |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1245 |
apply auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1246 |
done |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1247 |
then have "a \<bullet> i \<le> (x - (e / 2) *\<^sub>R i) \<bullet> i" and "(x + (e / 2) *\<^sub>R i) \<bullet> i \<le> b \<bullet> i" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1248 |
using e[THEN spec[where x="x - (e/2) *\<^sub>R i"]] |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1249 |
and e[THEN spec[where x="x + (e/2) *\<^sub>R i"]] |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1250 |
unfolding mem_box |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1251 |
using i |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1252 |
by blast+ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1253 |
then have "a \<bullet> i < x \<bullet> i" and "x \<bullet> i < b \<bullet> i" |
60420 | 1254 |
using \<open>e>0\<close> i |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1255 |
by (auto simp: inner_diff_left inner_Basis inner_add_left) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1256 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1257 |
then have "x \<in> box a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1258 |
unfolding mem_box by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1259 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1260 |
then show "?L \<subseteq> ?R" .. |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1261 |
qed |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1262 |
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63881
diff
changeset
|
1263 |
lemma bounded_cbox [simp]: |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1264 |
fixes a :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1265 |
shows "bounded (cbox a b)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1266 |
proof - |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1267 |
let ?b = "\<Sum>i\<in>Basis. \<bar>a\<bullet>i\<bar> + \<bar>b\<bullet>i\<bar>" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1268 |
{ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1269 |
fix x :: "'a" |
68120 | 1270 |
assume "\<And>i. i\<in>Basis \<Longrightarrow> a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i" |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1271 |
then have "(\<Sum>i\<in>Basis. \<bar>x \<bullet> i\<bar>) \<le> ?b" |
68120 | 1272 |
by (force simp: intro!: sum_mono) |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1273 |
then have "norm x \<le> ?b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1274 |
using norm_le_l1[of x] by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1275 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1276 |
then show ?thesis |
68120 | 1277 |
unfolding cbox_def bounded_iff by force |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1278 |
qed |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1279 |
|
62618
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1280 |
lemma bounded_box [simp]: |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1281 |
fixes a :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1282 |
shows "bounded (box a b)" |
68120 | 1283 |
using bounded_cbox[of a b] box_subset_cbox[of a b] bounded_subset[of "cbox a b" "box a b"] |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1284 |
by simp |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1285 |
|
62618
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1286 |
lemma not_interval_UNIV [simp]: |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1287 |
fixes a :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1288 |
shows "cbox a b \<noteq> UNIV" "box a b \<noteq> UNIV" |
62618
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1289 |
using bounded_box[of a b] bounded_cbox[of a b] by force+ |
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1290 |
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1291 |
lemma not_interval_UNIV2 [simp]: |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1292 |
fixes a :: "'a::euclidean_space" |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1293 |
shows "UNIV \<noteq> cbox a b" "UNIV \<noteq> box a b" |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1294 |
using bounded_box[of a b] bounded_cbox[of a b] by force+ |
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1295 |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1296 |
lemma box_midpoint: |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1297 |
fixes a :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1298 |
assumes "box a b \<noteq> {}" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1299 |
shows "((1/2) *\<^sub>R (a + b)) \<in> box a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1300 |
proof - |
68120 | 1301 |
have "a \<bullet> i < ((1 / 2) *\<^sub>R (a + b)) \<bullet> i \<and> ((1 / 2) *\<^sub>R (a + b)) \<bullet> i < b \<bullet> i" if "i \<in> Basis" for i |
1302 |
using assms that by (auto simp: inner_add_left box_ne_empty) |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1303 |
then show ?thesis unfolding mem_box by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1304 |
qed |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1305 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1306 |
lemma open_cbox_convex: |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1307 |
fixes x :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1308 |
assumes x: "x \<in> box a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1309 |
and y: "y \<in> cbox a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1310 |
and e: "0 < e" "e \<le> 1" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1311 |
shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<in> box a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1312 |
proof - |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1313 |
{ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1314 |
fix i :: 'a |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1315 |
assume i: "i \<in> Basis" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1316 |
have "a \<bullet> i = e * (a \<bullet> i) + (1 - e) * (a \<bullet> i)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1317 |
unfolding left_diff_distrib by simp |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1318 |
also have "\<dots> < e * (x \<bullet> i) + (1 - e) * (y \<bullet> i)" |
68120 | 1319 |
proof (rule add_less_le_mono) |
1320 |
show "e * (a \<bullet> i) < e * (x \<bullet> i)" |
|
1321 |
using \<open>0 < e\<close> i mem_box(1) x by auto |
|
1322 |
show "(1 - e) * (a \<bullet> i) \<le> (1 - e) * (y \<bullet> i)" |
|
1323 |
by (meson diff_ge_0_iff_ge \<open>e \<le> 1\<close> i mem_box(2) mult_left_mono y) |
|
1324 |
qed |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1325 |
finally have "a \<bullet> i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1326 |
unfolding inner_simps by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1327 |
moreover |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1328 |
{ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1329 |
have "b \<bullet> i = e * (b\<bullet>i) + (1 - e) * (b\<bullet>i)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1330 |
unfolding left_diff_distrib by simp |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1331 |
also have "\<dots> > e * (x \<bullet> i) + (1 - e) * (y \<bullet> i)" |
66827
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1332 |
proof (rule add_less_le_mono) |
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1333 |
show "e * (x \<bullet> i) < e * (b \<bullet> i)" |
68120 | 1334 |
using \<open>0 < e\<close> i mem_box(1) x by auto |
66827
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1335 |
show "(1 - e) * (y \<bullet> i) \<le> (1 - e) * (b \<bullet> i)" |
68120 | 1336 |
by (meson diff_ge_0_iff_ge \<open>e \<le> 1\<close> i mem_box(2) mult_left_mono y) |
66827
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1337 |
qed |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1338 |
finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i < b \<bullet> i" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1339 |
unfolding inner_simps by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1340 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1341 |
ultimately have "a \<bullet> i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i \<and> (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i < b \<bullet> i" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1342 |
by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1343 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1344 |
then show ?thesis |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1345 |
unfolding mem_box by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1346 |
qed |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1347 |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1348 |
lemma closure_cbox [simp]: "closure (cbox a b) = cbox a b" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1349 |
by (simp add: closed_cbox) |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1350 |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1351 |
lemma closure_box [simp]: |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1352 |
fixes a :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1353 |
assumes "box a b \<noteq> {}" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1354 |
shows "closure (box a b) = cbox a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1355 |
proof - |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1356 |
have ab: "a <e b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1357 |
using assms by (simp add: eucl_less_def box_ne_empty) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1358 |
let ?c = "(1 / 2) *\<^sub>R (a + b)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1359 |
{ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1360 |
fix x |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1361 |
assume as:"x \<in> cbox a b" |
63040 | 1362 |
define f where [abs_def]: "f n = x + (inverse (real n + 1)) *\<^sub>R (?c - x)" for n |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1363 |
{ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1364 |
fix n |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1365 |
assume fn: "f n <e b \<longrightarrow> a <e f n \<longrightarrow> f n = x" and xc: "x \<noteq> ?c" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1366 |
have *: "0 < inverse (real n + 1)" "inverse (real n + 1) \<le> 1" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1367 |
unfolding inverse_le_1_iff by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1368 |
have "(inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b)) + (1 - inverse (real n + 1)) *\<^sub>R x = |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1369 |
x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)" |
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1370 |
by (auto simp: algebra_simps) |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1371 |
then have "f n <e b" and "a <e f n" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1372 |
using open_cbox_convex[OF box_midpoint[OF assms] as *] |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1373 |
unfolding f_def by (auto simp: box_def eucl_less_def) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1374 |
then have False |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1375 |
using fn unfolding f_def using xc by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1376 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1377 |
moreover |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1378 |
{ |
61973 | 1379 |
assume "\<not> (f \<longlongrightarrow> x) sequentially" |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1380 |
{ |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1381 |
fix e :: real |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1382 |
assume "e > 0" |
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
1383 |
then obtain N :: nat where N: "inverse (real (N + 1)) < e" |
68120 | 1384 |
using reals_Archimedean by auto |
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
1385 |
have "inverse (real n + 1) < e" if "N \<le> n" for n |
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
1386 |
by (auto intro!: that le_less_trans [OF _ N]) |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1387 |
then have "\<exists>N::nat. \<forall>n\<ge>N. inverse (real n + 1) < e" by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1388 |
} |
61973 | 1389 |
then have "((\<lambda>n. inverse (real n + 1)) \<longlongrightarrow> 0) sequentially" |
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1390 |
unfolding lim_sequentially by(auto simp: dist_norm) |
61973 | 1391 |
then have "(f \<longlongrightarrow> x) sequentially" |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1392 |
unfolding f_def |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1393 |
using tendsto_add[OF tendsto_const, of "\<lambda>n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x] |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1394 |
using tendsto_scaleR [OF _ tendsto_const, of "\<lambda>n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1395 |
by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1396 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1397 |
ultimately have "x \<in> closure (box a b)" |
68120 | 1398 |
using as box_midpoint[OF assms] |
1399 |
unfolding closure_def islimpt_sequential |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1400 |
by (cases "x=?c") (auto simp: in_box_eucl_less) |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1401 |
} |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1402 |
then show ?thesis |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1403 |
using closure_minimal[OF box_subset_cbox, of a b] by blast |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1404 |
qed |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1405 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1406 |
lemma bounded_subset_box_symmetric: |
68120 | 1407 |
fixes S :: "('a::euclidean_space) set" |
1408 |
assumes "bounded S" |
|
1409 |
obtains a where "S \<subseteq> box (-a) a" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1410 |
proof - |
68120 | 1411 |
obtain b where "b>0" and b: "\<forall>x\<in>S. norm x \<le> b" |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1412 |
using assms[unfolded bounded_pos] by auto |
63040 | 1413 |
define a :: 'a where "a = (\<Sum>i\<in>Basis. (b + 1) *\<^sub>R i)" |
68120 | 1414 |
have "(-a)\<bullet>i < x\<bullet>i" and "x\<bullet>i < a\<bullet>i" if "x \<in> S" and i: "i \<in> Basis" for x i |
1415 |
using b Basis_le_norm[OF i, of x] that by (auto simp: a_def) |
|
1416 |
then have "S \<subseteq> box (-a) a" |
|
1417 |
by (auto simp: simp add: box_def) |
|
1418 |
then show ?thesis .. |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1419 |
qed |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1420 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1421 |
lemma bounded_subset_cbox_symmetric: |
68120 | 1422 |
fixes S :: "('a::euclidean_space) set" |
1423 |
assumes "bounded S" |
|
1424 |
obtains a where "S \<subseteq> cbox (-a) a" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1425 |
proof - |
68120 | 1426 |
obtain a where "S \<subseteq> box (-a) a" |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1427 |
using bounded_subset_box_symmetric[OF assms] by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1428 |
then show ?thesis |
68120 | 1429 |
by (meson box_subset_cbox dual_order.trans that) |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1430 |
qed |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1431 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1432 |
lemma frontier_cbox: |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1433 |
fixes a b :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1434 |
shows "frontier (cbox a b) = cbox a b - box a b" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1435 |
unfolding frontier_def unfolding interior_cbox and closure_closed[OF closed_cbox] .. |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1436 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1437 |
lemma frontier_box: |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1438 |
fixes a b :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1439 |
shows "frontier (box a b) = (if box a b = {} then {} else cbox a b - box a b)" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1440 |
proof (cases "box a b = {}") |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1441 |
case True |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1442 |
then show ?thesis |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1443 |
using frontier_empty by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1444 |
next |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1445 |
case False |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1446 |
then show ?thesis |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1447 |
unfolding frontier_def and closure_box[OF False] and interior_open[OF open_box] |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1448 |
by auto |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1449 |
qed |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1450 |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
1451 |
lemma Int_interval_mixed_eq_empty: |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1452 |
fixes a :: "'a::euclidean_space" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1453 |
assumes "box c d \<noteq> {}" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1454 |
shows "box a b \<inter> cbox c d = {} \<longleftrightarrow> box a b \<inter> box c d = {}" |
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1455 |
unfolding closure_box[OF assms, symmetric] |
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1456 |
unfolding open_Int_closure_eq_empty[OF open_box] .. |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1457 |
|
69611
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1458 |
subsection \<open>Class Instances\<close> |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1459 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1460 |
lemma compact_lemma: |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1461 |
fixes f :: "nat \<Rightarrow> 'a::euclidean_space" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1462 |
assumes "bounded (range f)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1463 |
shows "\<forall>d\<subseteq>Basis. \<exists>l::'a. \<exists> r. |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1464 |
strict_mono r \<and> (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1465 |
by (rule compact_lemma_general[where unproj="\<lambda>e. \<Sum>i\<in>Basis. e i *\<^sub>R i"]) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1466 |
(auto intro!: assms bounded_linear_inner_left bounded_linear_image |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1467 |
simp: euclidean_representation) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1468 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1469 |
instance%important euclidean_space \<subseteq> heine_borel |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1470 |
proof%unimportant |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1471 |
fix f :: "nat \<Rightarrow> 'a" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1472 |
assume f: "bounded (range f)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1473 |
then obtain l::'a and r where r: "strict_mono r" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1474 |
and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1475 |
using compact_lemma [OF f] by blast |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1476 |
{ |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1477 |
fix e::real |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1478 |
assume "e > 0" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1479 |
hence "e / real_of_nat DIM('a) > 0" by (simp add: DIM_positive) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1480 |
with l have "eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e / (real_of_nat DIM('a))) sequentially" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1481 |
by simp |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1482 |
moreover |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1483 |
{ |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1484 |
fix n |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1485 |
assume n: "\<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e / (real_of_nat DIM('a))" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1486 |
have "dist (f (r n)) l \<le> (\<Sum>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i))" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1487 |
apply (subst euclidean_dist_l2) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1488 |
using zero_le_dist |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1489 |
apply (rule L2_set_le_sum) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1490 |
done |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1491 |
also have "\<dots> < (\<Sum>i\<in>(Basis::'a set). e / (real_of_nat DIM('a)))" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1492 |
apply (rule sum_strict_mono) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1493 |
using n |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1494 |
apply auto |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1495 |
done |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1496 |
finally have "dist (f (r n)) l < e" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1497 |
by auto |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1498 |
} |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1499 |
ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1500 |
by (rule eventually_mono) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1501 |
} |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1502 |
then have *: "((f \<circ> r) \<longlongrightarrow> l) sequentially" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1503 |
unfolding o_def tendsto_iff by simp |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1504 |
with r show "\<exists>l r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1505 |
by auto |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1506 |
qed |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1507 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1508 |
instance%important euclidean_space \<subseteq> banach .. |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1509 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1510 |
instance euclidean_space \<subseteq> second_countable_topology |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1511 |
proof |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1512 |
define a where "a f = (\<Sum>i\<in>Basis. fst (f i) *\<^sub>R i)" for f :: "'a \<Rightarrow> real \<times> real" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1513 |
then have a: "\<And>f. (\<Sum>i\<in>Basis. fst (f i) *\<^sub>R i) = a f" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1514 |
by simp |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1515 |
define b where "b f = (\<Sum>i\<in>Basis. snd (f i) *\<^sub>R i)" for f :: "'a \<Rightarrow> real \<times> real" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1516 |
then have b: "\<And>f. (\<Sum>i\<in>Basis. snd (f i) *\<^sub>R i) = b f" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1517 |
by simp |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1518 |
define B where "B = (\<lambda>f. box (a f) (b f)) ` (Basis \<rightarrow>\<^sub>E (\<rat> \<times> \<rat>))" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1519 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1520 |
have "Ball B open" by (simp add: B_def open_box) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1521 |
moreover have "(\<forall>A. open A \<longrightarrow> (\<exists>B'\<subseteq>B. \<Union>B' = A))" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1522 |
proof safe |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1523 |
fix A::"'a set" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1524 |
assume "open A" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1525 |
show "\<exists>B'\<subseteq>B. \<Union>B' = A" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1526 |
apply (rule exI[of _ "{b\<in>B. b \<subseteq> A}"]) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1527 |
apply (subst (3) open_UNION_box[OF \<open>open A\<close>]) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1528 |
apply (auto simp: a b B_def) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1529 |
done |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1530 |
qed |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1531 |
ultimately |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1532 |
have "topological_basis B" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1533 |
unfolding topological_basis_def by blast |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1534 |
moreover |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1535 |
have "countable B" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1536 |
unfolding B_def |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1537 |
by (intro countable_image countable_PiE finite_Basis countable_SIGMA countable_rat) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1538 |
ultimately show "\<exists>B::'a set set. countable B \<and> open = generate_topology B" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1539 |
by (blast intro: topological_basis_imp_subbasis) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1540 |
qed |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1541 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1542 |
instance euclidean_space \<subseteq> polish_space .. |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1543 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1544 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1545 |
subsection \<open>Compact Boxes\<close> |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1546 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1547 |
lemma compact_cbox [simp]: |
61076 | 1548 |
fixes a :: "'a::euclidean_space" |
69611
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1549 |
shows "compact (cbox a b)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1550 |
using bounded_closed_imp_seq_compact[of "cbox a b"] using bounded_cbox[of a b] |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1551 |
by (auto simp: compact_eq_seq_compact_metric) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1552 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1553 |
proposition is_interval_compact: |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1554 |
"is_interval S \<and> compact S \<longleftrightarrow> (\<exists>a b. S = cbox a b)" (is "?lhs = ?rhs") |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1555 |
proof (cases "S = {}") |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1556 |
case True |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1557 |
with empty_as_interval show ?thesis by auto |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1558 |
next |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1559 |
case False |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1560 |
show ?thesis |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1561 |
proof |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1562 |
assume L: ?lhs |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1563 |
then have "is_interval S" "compact S" by auto |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1564 |
define a where "a \<equiv> \<Sum>i\<in>Basis. (INF x\<in>S. x \<bullet> i) *\<^sub>R i" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1565 |
define b where "b \<equiv> \<Sum>i\<in>Basis. (SUP x\<in>S. x \<bullet> i) *\<^sub>R i" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1566 |
have 1: "\<And>x i. \<lbrakk>x \<in> S; i \<in> Basis\<rbrakk> \<Longrightarrow> (INF x\<in>S. x \<bullet> i) \<le> x \<bullet> i" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1567 |
by (simp add: cInf_lower bounded_inner_imp_bdd_below compact_imp_bounded L) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1568 |
have 2: "\<And>x i. \<lbrakk>x \<in> S; i \<in> Basis\<rbrakk> \<Longrightarrow> x \<bullet> i \<le> (SUP x\<in>S. x \<bullet> i)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1569 |
by (simp add: cSup_upper bounded_inner_imp_bdd_above compact_imp_bounded L) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1570 |
have 3: "x \<in> S" if inf: "\<And>i. i \<in> Basis \<Longrightarrow> (INF x\<in>S. x \<bullet> i) \<le> x \<bullet> i" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1571 |
and sup: "\<And>i. i \<in> Basis \<Longrightarrow> x \<bullet> i \<le> (SUP x\<in>S. x \<bullet> i)" for x |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1572 |
proof (rule mem_box_componentwiseI [OF \<open>is_interval S\<close>]) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1573 |
fix i::'a |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1574 |
assume i: "i \<in> Basis" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1575 |
have cont: "continuous_on S (\<lambda>x. x \<bullet> i)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1576 |
by (intro continuous_intros) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1577 |
obtain a where "a \<in> S" and a: "\<And>y. y\<in>S \<Longrightarrow> a \<bullet> i \<le> y \<bullet> i" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1578 |
using continuous_attains_inf [OF \<open>compact S\<close> False cont] by blast |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1579 |
obtain b where "b \<in> S" and b: "\<And>y. y\<in>S \<Longrightarrow> y \<bullet> i \<le> b \<bullet> i" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1580 |
using continuous_attains_sup [OF \<open>compact S\<close> False cont] by blast |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1581 |
have "a \<bullet> i \<le> (INF x\<in>S. x \<bullet> i)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1582 |
by (simp add: False a cINF_greatest) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1583 |
also have "\<dots> \<le> x \<bullet> i" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1584 |
by (simp add: i inf) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1585 |
finally have ai: "a \<bullet> i \<le> x \<bullet> i" . |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1586 |
have "x \<bullet> i \<le> (SUP x\<in>S. x \<bullet> i)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1587 |
by (simp add: i sup) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1588 |
also have "(SUP x\<in>S. x \<bullet> i) \<le> b \<bullet> i" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1589 |
by (simp add: False b cSUP_least) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1590 |
finally have bi: "x \<bullet> i \<le> b \<bullet> i" . |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1591 |
show "x \<bullet> i \<in> (\<lambda>x. x \<bullet> i) ` S" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1592 |
apply (rule_tac x="\<Sum>j\<in>Basis. (if j = i then x \<bullet> i else a \<bullet> j) *\<^sub>R j" in image_eqI) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1593 |
apply (simp add: i) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1594 |
apply (rule mem_is_intervalI [OF \<open>is_interval S\<close> \<open>a \<in> S\<close> \<open>b \<in> S\<close>]) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1595 |
using i ai bi apply force |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1596 |
done |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1597 |
qed |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1598 |
have "S = cbox a b" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1599 |
by (auto simp: a_def b_def mem_box intro: 1 2 3) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1600 |
then show ?rhs |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1601 |
by blast |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1602 |
next |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1603 |
assume R: ?rhs |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1604 |
then show ?lhs |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1605 |
using compact_cbox is_interval_cbox by blast |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1606 |
qed |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1607 |
qed |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1608 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1609 |
|
69615
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1610 |
subsection%unimportant\<open>Componentwise limits and continuity\<close> |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1611 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1612 |
text\<open>But is the premise really necessary? Need to generalise @{thm euclidean_dist_l2}\<close> |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1613 |
lemma Euclidean_dist_upper: "i \<in> Basis \<Longrightarrow> dist (x \<bullet> i) (y \<bullet> i) \<le> dist x y" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1614 |
by (metis (no_types) member_le_L2_set euclidean_dist_l2 finite_Basis) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1615 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1616 |
text\<open>But is the premise \<^term>\<open>i \<in> Basis\<close> really necessary?\<close> |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1617 |
lemma open_preimage_inner: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1618 |
assumes "open S" "i \<in> Basis" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1619 |
shows "open {x. x \<bullet> i \<in> S}" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1620 |
proof (rule openI, simp) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1621 |
fix x |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1622 |
assume x: "x \<bullet> i \<in> S" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1623 |
with assms obtain e where "0 < e" and e: "ball (x \<bullet> i) e \<subseteq> S" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1624 |
by (auto simp: open_contains_ball_eq) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1625 |
have "\<exists>e>0. ball (y \<bullet> i) e \<subseteq> S" if dxy: "dist x y < e / 2" for y |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1626 |
proof (intro exI conjI) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1627 |
have "dist (x \<bullet> i) (y \<bullet> i) < e / 2" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1628 |
by (meson \<open>i \<in> Basis\<close> dual_order.trans Euclidean_dist_upper not_le that) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1629 |
then have "dist (x \<bullet> i) z < e" if "dist (y \<bullet> i) z < e / 2" for z |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1630 |
by (metis dist_commute dist_triangle_half_l that) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1631 |
then have "ball (y \<bullet> i) (e / 2) \<subseteq> ball (x \<bullet> i) e" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1632 |
using mem_ball by blast |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1633 |
with e show "ball (y \<bullet> i) (e / 2) \<subseteq> S" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1634 |
by (metis order_trans) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1635 |
qed (simp add: \<open>0 < e\<close>) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1636 |
then show "\<exists>e>0. ball x e \<subseteq> {s. s \<bullet> i \<in> S}" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1637 |
by (metis (no_types, lifting) \<open>0 < e\<close> \<open>open S\<close> half_gt_zero_iff mem_Collect_eq mem_ball open_contains_ball_eq subsetI) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1638 |
qed |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1639 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1640 |
proposition tendsto_componentwise_iff: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1641 |
fixes f :: "_ \<Rightarrow> 'b::euclidean_space" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1642 |
shows "(f \<longlongrightarrow> l) F \<longleftrightarrow> (\<forall>i \<in> Basis. ((\<lambda>x. (f x \<bullet> i)) \<longlongrightarrow> (l \<bullet> i)) F)" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1643 |
(is "?lhs = ?rhs") |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1644 |
proof |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1645 |
assume ?lhs |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1646 |
then show ?rhs |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1647 |
unfolding tendsto_def |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1648 |
apply clarify |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1649 |
apply (drule_tac x="{s. s \<bullet> i \<in> S}" in spec) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1650 |
apply (auto simp: open_preimage_inner) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1651 |
done |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1652 |
next |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1653 |
assume R: ?rhs |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1654 |
then have "\<And>e. e > 0 \<Longrightarrow> \<forall>i\<in>Basis. \<forall>\<^sub>F x in F. dist (f x \<bullet> i) (l \<bullet> i) < e" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1655 |
unfolding tendsto_iff by blast |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1656 |
then have R': "\<And>e. e > 0 \<Longrightarrow> \<forall>\<^sub>F x in F. \<forall>i\<in>Basis. dist (f x \<bullet> i) (l \<bullet> i) < e" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1657 |
by (simp add: eventually_ball_finite_distrib [symmetric]) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1658 |
show ?lhs |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1659 |
unfolding tendsto_iff |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1660 |
proof clarify |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1661 |
fix e::real |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1662 |
assume "0 < e" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1663 |
have *: "L2_set (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis < e" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1664 |
if "\<forall>i\<in>Basis. dist (f x \<bullet> i) (l \<bullet> i) < e / real DIM('b)" for x |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1665 |
proof - |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1666 |
have "L2_set (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis \<le> sum (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1667 |
by (simp add: L2_set_le_sum) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1668 |
also have "... < DIM('b) * (e / real DIM('b))" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1669 |
apply (rule sum_bounded_above_strict) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1670 |
using that by auto |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1671 |
also have "... = e" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1672 |
by (simp add: field_simps) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1673 |
finally show "L2_set (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis < e" . |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1674 |
qed |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1675 |
have "\<forall>\<^sub>F x in F. \<forall>i\<in>Basis. dist (f x \<bullet> i) (l \<bullet> i) < e / DIM('b)" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1676 |
apply (rule R') |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1677 |
using \<open>0 < e\<close> by simp |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1678 |
then show "\<forall>\<^sub>F x in F. dist (f x) l < e" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1679 |
apply (rule eventually_mono) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1680 |
apply (subst euclidean_dist_l2) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1681 |
using * by blast |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1682 |
qed |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1683 |
qed |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1684 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1685 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1686 |
corollary continuous_componentwise: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1687 |
"continuous F f \<longleftrightarrow> (\<forall>i \<in> Basis. continuous F (\<lambda>x. (f x \<bullet> i)))" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1688 |
by (simp add: continuous_def tendsto_componentwise_iff [symmetric]) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1689 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1690 |
corollary continuous_on_componentwise: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1691 |
fixes S :: "'a :: t2_space set" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1692 |
shows "continuous_on S f \<longleftrightarrow> (\<forall>i \<in> Basis. continuous_on S (\<lambda>x. (f x \<bullet> i)))" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1693 |
apply (simp add: continuous_on_eq_continuous_within) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1694 |
using continuous_componentwise by blast |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1695 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1696 |
lemma linear_componentwise_iff: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1697 |
"(linear f') \<longleftrightarrow> (\<forall>i\<in>Basis. linear (\<lambda>x. f' x \<bullet> i))" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1698 |
apply (auto simp: linear_iff inner_left_distrib) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1699 |
apply (metis inner_left_distrib euclidean_eq_iff) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1700 |
by (metis euclidean_eqI inner_scaleR_left) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1701 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1702 |
lemma bounded_linear_componentwise_iff: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1703 |
"(bounded_linear f') \<longleftrightarrow> (\<forall>i\<in>Basis. bounded_linear (\<lambda>x. f' x \<bullet> i))" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1704 |
(is "?lhs = ?rhs") |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1705 |
proof |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1706 |
assume ?lhs then show ?rhs |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1707 |
by (simp add: bounded_linear_inner_left_comp) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1708 |
next |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1709 |
assume ?rhs |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1710 |
then have "(\<forall>i\<in>Basis. \<exists>K. \<forall>x. \<bar>f' x \<bullet> i\<bar> \<le> norm x * K)" "linear f'" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1711 |
by (auto simp: bounded_linear_def bounded_linear_axioms_def linear_componentwise_iff [symmetric] ball_conj_distrib) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1712 |
then obtain F where F: "\<And>i x. i \<in> Basis \<Longrightarrow> \<bar>f' x \<bullet> i\<bar> \<le> norm x * F i" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1713 |
by metis |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1714 |
have "norm (f' x) \<le> norm x * sum F Basis" for x |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1715 |
proof - |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1716 |
have "norm (f' x) \<le> (\<Sum>i\<in>Basis. \<bar>f' x \<bullet> i\<bar>)" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1717 |
by (rule norm_le_l1) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1718 |
also have "... \<le> (\<Sum>i\<in>Basis. norm x * F i)" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1719 |
by (metis F sum_mono) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1720 |
also have "... = norm x * sum F Basis" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1721 |
by (simp add: sum_distrib_left) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1722 |
finally show ?thesis . |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1723 |
qed |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1724 |
then show ?lhs |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1725 |
by (force simp: bounded_linear_def bounded_linear_axioms_def \<open>linear f'\<close>) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1726 |
qed |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1727 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1728 |
subsection%unimportant \<open>Continuous Extension\<close> |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1729 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1730 |
definition clamp :: "'a::euclidean_space \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1731 |
"clamp a b x = (if (\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1732 |
then (\<Sum>i\<in>Basis. (if x\<bullet>i < a\<bullet>i then a\<bullet>i else if x\<bullet>i \<le> b\<bullet>i then x\<bullet>i else b\<bullet>i) *\<^sub>R i) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1733 |
else a)" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1734 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1735 |
lemma clamp_in_interval[simp]: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1736 |
assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1737 |
shows "clamp a b x \<in> cbox a b" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1738 |
unfolding clamp_def |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1739 |
using box_ne_empty(1)[of a b] assms by (auto simp: cbox_def) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1740 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1741 |
lemma clamp_cancel_cbox[simp]: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1742 |
fixes x a b :: "'a::euclidean_space" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1743 |
assumes x: "x \<in> cbox a b" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1744 |
shows "clamp a b x = x" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1745 |
using assms |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1746 |
by (auto simp: clamp_def mem_box intro!: euclidean_eqI[where 'a='a]) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1747 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1748 |
lemma clamp_empty_interval: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1749 |
assumes "i \<in> Basis" "a \<bullet> i > b \<bullet> i" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1750 |
shows "clamp a b = (\<lambda>_. a)" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1751 |
using assms |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1752 |
by (force simp: clamp_def[abs_def] split: if_splits intro!: ext) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1753 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1754 |
lemma dist_clamps_le_dist_args: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1755 |
fixes x :: "'a::euclidean_space" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1756 |
shows "dist (clamp a b y) (clamp a b x) \<le> dist y x" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1757 |
proof cases |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1758 |
assume le: "(\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1759 |
then have "(\<Sum>i\<in>Basis. (dist (clamp a b y \<bullet> i) (clamp a b x \<bullet> i))\<^sup>2) \<le> |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1760 |
(\<Sum>i\<in>Basis. (dist (y \<bullet> i) (x \<bullet> i))\<^sup>2)" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1761 |
by (auto intro!: sum_mono simp: clamp_def dist_real_def abs_le_square_iff[symmetric]) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1762 |
then show ?thesis |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1763 |
by (auto intro: real_sqrt_le_mono |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1764 |
simp: euclidean_dist_l2[where y=x] euclidean_dist_l2[where y="clamp a b x"] L2_set_def) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1765 |
qed (auto simp: clamp_def) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1766 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1767 |
lemma clamp_continuous_at: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1768 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::metric_space" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1769 |
and x :: 'a |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1770 |
assumes f_cont: "continuous_on (cbox a b) f" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1771 |
shows "continuous (at x) (\<lambda>x. f (clamp a b x))" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1772 |
proof cases |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1773 |
assume le: "(\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1774 |
show ?thesis |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1775 |
unfolding continuous_at_eps_delta |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1776 |
proof safe |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1777 |
fix x :: 'a |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1778 |
fix e :: real |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1779 |
assume "e > 0" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1780 |
moreover have "clamp a b x \<in> cbox a b" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1781 |
by (simp add: clamp_in_interval le) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1782 |
moreover note f_cont[simplified continuous_on_iff] |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1783 |
ultimately |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1784 |
obtain d where d: "0 < d" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1785 |
"\<And>x'. x' \<in> cbox a b \<Longrightarrow> dist x' (clamp a b x) < d \<Longrightarrow> dist (f x') (f (clamp a b x)) < e" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1786 |
by force |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1787 |
show "\<exists>d>0. \<forall>x'. dist x' x < d \<longrightarrow> |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1788 |
dist (f (clamp a b x')) (f (clamp a b x)) < e" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1789 |
using le |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1790 |
by (auto intro!: d clamp_in_interval dist_clamps_le_dist_args[THEN le_less_trans]) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1791 |
qed |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1792 |
qed (auto simp: clamp_empty_interval) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1793 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1794 |
lemma clamp_continuous_on: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1795 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::metric_space" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1796 |
assumes f_cont: "continuous_on (cbox a b) f" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1797 |
shows "continuous_on S (\<lambda>x. f (clamp a b x))" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1798 |
using assms |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1799 |
by (auto intro: continuous_at_imp_continuous_on clamp_continuous_at) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1800 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1801 |
lemma clamp_bounded: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1802 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::metric_space" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1803 |
assumes bounded: "bounded (f ` (cbox a b))" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1804 |
shows "bounded (range (\<lambda>x. f (clamp a b x)))" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1805 |
proof cases |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1806 |
assume le: "(\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1807 |
from bounded obtain c where f_bound: "\<forall>x\<in>f ` cbox a b. dist undefined x \<le> c" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1808 |
by (auto simp: bounded_any_center[where a=undefined]) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1809 |
then show ?thesis |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1810 |
by (auto intro!: exI[where x=c] clamp_in_interval[OF le[rule_format]] |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1811 |
simp: bounded_any_center[where a=undefined]) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1812 |
qed (auto simp: clamp_empty_interval image_def) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1813 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1814 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1815 |
definition ext_cont :: "('a::euclidean_space \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1816 |
where "ext_cont f a b = (\<lambda>x. f (clamp a b x))" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1817 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1818 |
lemma ext_cont_cancel_cbox[simp]: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1819 |
fixes x a b :: "'a::euclidean_space" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1820 |
assumes x: "x \<in> cbox a b" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1821 |
shows "ext_cont f a b x = f x" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1822 |
using assms |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1823 |
unfolding ext_cont_def |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1824 |
by (auto simp: clamp_def mem_box intro!: euclidean_eqI[where 'a='a] arg_cong[where f=f]) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1825 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1826 |
lemma continuous_on_ext_cont[continuous_intros]: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1827 |
"continuous_on (cbox a b) f \<Longrightarrow> continuous_on S (ext_cont f a b)" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1828 |
by (auto intro!: clamp_continuous_on simp: ext_cont_def) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1829 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1830 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1831 |
subsection \<open>Separability\<close> |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1832 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1833 |
lemma univ_second_countable_sequence: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1834 |
obtains B :: "nat \<Rightarrow> 'a::euclidean_space set" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1835 |
where "inj B" "\<And>n. open(B n)" "\<And>S. open S \<Longrightarrow> \<exists>k. S = \<Union>{B n |n. n \<in> k}" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1836 |
proof - |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1837 |
obtain \<B> :: "'a set set" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1838 |
where "countable \<B>" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1839 |
and opn: "\<And>C. C \<in> \<B> \<Longrightarrow> open C" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1840 |
and Un: "\<And>S. open S \<Longrightarrow> \<exists>U. U \<subseteq> \<B> \<and> S = \<Union>U" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1841 |
using univ_second_countable by blast |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1842 |
have *: "infinite (range (\<lambda>n. ball (0::'a) (inverse(Suc n))))" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1843 |
apply (rule Infinite_Set.range_inj_infinite) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1844 |
apply (simp add: inj_on_def ball_eq_ball_iff) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1845 |
done |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1846 |
have "infinite \<B>" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1847 |
proof |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1848 |
assume "finite \<B>" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1849 |
then have "finite (Union ` (Pow \<B>))" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1850 |
by simp |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1851 |
then have "finite (range (\<lambda>n. ball (0::'a) (inverse(Suc n))))" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1852 |
apply (rule rev_finite_subset) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1853 |
by (metis (no_types, lifting) PowI image_eqI image_subset_iff Un [OF open_ball]) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1854 |
with * show False by simp |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1855 |
qed |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1856 |
obtain f :: "nat \<Rightarrow> 'a set" where "\<B> = range f" "inj f" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1857 |
by (blast intro: countable_as_injective_image [OF \<open>countable \<B>\<close> \<open>infinite \<B>\<close>]) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1858 |
have *: "\<exists>k. S = \<Union>{f n |n. n \<in> k}" if "open S" for S |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1859 |
using Un [OF that] |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1860 |
apply clarify |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1861 |
apply (rule_tac x="f-`U" in exI) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1862 |
using \<open>inj f\<close> \<open>\<B> = range f\<close> apply force |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1863 |
done |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1864 |
show ?thesis |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1865 |
apply (rule that [OF \<open>inj f\<close> _ *]) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1866 |
apply (auto simp: \<open>\<B> = range f\<close> opn) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1867 |
done |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1868 |
qed |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1869 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1870 |
proposition separable: |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1871 |
fixes S :: "'a::{metric_space, second_countable_topology} set" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1872 |
obtains T where "countable T" "T \<subseteq> S" "S \<subseteq> closure T" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1873 |
proof - |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1874 |
obtain \<B> :: "'a set set" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1875 |
where "countable \<B>" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1876 |
and "{} \<notin> \<B>" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1877 |
and ope: "\<And>C. C \<in> \<B> \<Longrightarrow> openin(subtopology euclidean S) C" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1878 |
and if_ope: "\<And>T. openin(subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1879 |
by (meson subset_second_countable) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1880 |
then obtain f where f: "\<And>C. C \<in> \<B> \<Longrightarrow> f C \<in> C" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1881 |
by (metis equals0I) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1882 |
show ?thesis |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1883 |
proof |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1884 |
show "countable (f ` \<B>)" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1885 |
by (simp add: \<open>countable \<B>\<close>) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1886 |
show "f ` \<B> \<subseteq> S" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1887 |
using ope f openin_imp_subset by blast |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1888 |
show "S \<subseteq> closure (f ` \<B>)" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1889 |
proof (clarsimp simp: closure_approachable) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1890 |
fix x and e::real |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1891 |
assume "x \<in> S" "0 < e" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1892 |
have "openin (subtopology euclidean S) (S \<inter> ball x e)" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1893 |
by (simp add: openin_Int_open) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1894 |
with if_ope obtain \<U> where \<U>: "\<U> \<subseteq> \<B>" "S \<inter> ball x e = \<Union>\<U>" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1895 |
by meson |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1896 |
show "\<exists>C \<in> \<B>. dist (f C) x < e" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1897 |
proof (cases "\<U> = {}") |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1898 |
case True |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1899 |
then show ?thesis |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1900 |
using \<open>0 < e\<close> \<U> \<open>x \<in> S\<close> by auto |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1901 |
next |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1902 |
case False |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1903 |
then obtain C where "C \<in> \<U>" by blast |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1904 |
show ?thesis |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1905 |
proof |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1906 |
show "dist (f C) x < e" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1907 |
by (metis Int_iff Union_iff \<U> \<open>C \<in> \<U>\<close> dist_commute f mem_ball subsetCE) |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1908 |
show "C \<in> \<B>" |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1909 |
using \<open>\<U> \<subseteq> \<B>\<close> \<open>C \<in> \<U>\<close> by blast |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1910 |
qed |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1911 |
qed |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1912 |
qed |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1913 |
qed |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1914 |
qed |
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1915 |
|
e502cd4d7062
moved material from Connected.thy to more appropriate places
immler
parents:
69613
diff
changeset
|
1916 |
|
69613
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1917 |
subsection%unimportant \<open>Diameter\<close> |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1918 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1919 |
lemma diameter_cball [simp]: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1920 |
fixes a :: "'a::euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1921 |
shows "diameter(cball a r) = (if r < 0 then 0 else 2*r)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1922 |
proof - |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1923 |
have "diameter(cball a r) = 2*r" if "r \<ge> 0" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1924 |
proof (rule order_antisym) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1925 |
show "diameter (cball a r) \<le> 2*r" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1926 |
proof (rule diameter_le) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1927 |
fix x y assume "x \<in> cball a r" "y \<in> cball a r" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1928 |
then have "norm (x - a) \<le> r" "norm (a - y) \<le> r" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1929 |
by (auto simp: dist_norm norm_minus_commute) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1930 |
then have "norm (x - y) \<le> r+r" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1931 |
using norm_diff_triangle_le by blast |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1932 |
then show "norm (x - y) \<le> 2*r" by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1933 |
qed (simp add: that) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1934 |
have "2*r = dist (a + r *\<^sub>R (SOME i. i \<in> Basis)) (a - r *\<^sub>R (SOME i. i \<in> Basis))" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1935 |
apply (simp add: dist_norm) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1936 |
by (metis abs_of_nonneg mult.right_neutral norm_numeral norm_scaleR norm_some_Basis real_norm_def scaleR_2 that) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1937 |
also have "... \<le> diameter (cball a r)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1938 |
apply (rule diameter_bounded_bound) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1939 |
using that by (auto simp: dist_norm) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1940 |
finally show "2*r \<le> diameter (cball a r)" . |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1941 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1942 |
then show ?thesis by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1943 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1944 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1945 |
lemma diameter_ball [simp]: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1946 |
fixes a :: "'a::euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1947 |
shows "diameter(ball a r) = (if r < 0 then 0 else 2*r)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1948 |
proof - |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1949 |
have "diameter(ball a r) = 2*r" if "r > 0" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1950 |
by (metis bounded_ball diameter_closure closure_ball diameter_cball less_eq_real_def linorder_not_less that) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1951 |
then show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1952 |
by (simp add: diameter_def) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1953 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1954 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1955 |
lemma diameter_closed_interval [simp]: "diameter {a..b} = (if b < a then 0 else b-a)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1956 |
proof - |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1957 |
have "{a .. b} = cball ((a+b)/2) ((b-a)/2)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1958 |
by (auto simp: dist_norm abs_if divide_simps split: if_split_asm) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1959 |
then show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1960 |
by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1961 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1962 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1963 |
lemma diameter_open_interval [simp]: "diameter {a<..<b} = (if b < a then 0 else b-a)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1964 |
proof - |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1965 |
have "{a <..< b} = ball ((a+b)/2) ((b-a)/2)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1966 |
by (auto simp: dist_norm abs_if divide_simps split: if_split_asm) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1967 |
then show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1968 |
by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1969 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1970 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1971 |
lemma diameter_cbox: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1972 |
fixes a b::"'a::euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1973 |
shows "(\<forall>i \<in> Basis. a \<bullet> i \<le> b \<bullet> i) \<Longrightarrow> diameter (cbox a b) = dist a b" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1974 |
by (force simp: diameter_def intro!: cSup_eq_maximum L2_set_mono |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1975 |
simp: euclidean_dist_l2[where 'a='a] cbox_def dist_norm) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1976 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
1977 |
|
69617 | 1978 |
subsection%unimportant\<open>Relating linear images to open/closed/interior/closure/connected\<close> |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1979 |
|
69611
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1980 |
proposition open_surjective_linear_image: |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1981 |
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1982 |
assumes "open A" "linear f" "surj f" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1983 |
shows "open(f ` A)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1984 |
unfolding open_dist |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1985 |
proof clarify |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1986 |
fix x |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1987 |
assume "x \<in> A" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1988 |
have "bounded (inv f ` Basis)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1989 |
by (simp add: finite_imp_bounded) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1990 |
with bounded_pos obtain B where "B > 0" and B: "\<And>x. x \<in> inv f ` Basis \<Longrightarrow> norm x \<le> B" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1991 |
by metis |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1992 |
obtain e where "e > 0" and e: "\<And>z. dist z x < e \<Longrightarrow> z \<in> A" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1993 |
by (metis open_dist \<open>x \<in> A\<close> \<open>open A\<close>) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1994 |
define \<delta> where "\<delta> \<equiv> e / B / DIM('b)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1995 |
show "\<exists>e>0. \<forall>y. dist y (f x) < e \<longrightarrow> y \<in> f ` A" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1996 |
proof (intro exI conjI) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1997 |
show "\<delta> > 0" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1998 |
using \<open>e > 0\<close> \<open>B > 0\<close> by (simp add: \<delta>_def divide_simps) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
1999 |
have "y \<in> f ` A" if "dist y (f x) * (B * real DIM('b)) < e" for y |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2000 |
proof - |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2001 |
define u where "u \<equiv> y - f x" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2002 |
show ?thesis |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2003 |
proof (rule image_eqI) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2004 |
show "y = f (x + (\<Sum>i\<in>Basis. (u \<bullet> i) *\<^sub>R inv f i))" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2005 |
apply (simp add: linear_add linear_sum linear.scaleR \<open>linear f\<close> surj_f_inv_f \<open>surj f\<close>) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2006 |
apply (simp add: euclidean_representation u_def) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2007 |
done |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2008 |
have "dist (x + (\<Sum>i\<in>Basis. (u \<bullet> i) *\<^sub>R inv f i)) x \<le> (\<Sum>i\<in>Basis. norm ((u \<bullet> i) *\<^sub>R inv f i))" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2009 |
by (simp add: dist_norm sum_norm_le) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2010 |
also have "... = (\<Sum>i\<in>Basis. \<bar>u \<bullet> i\<bar> * norm (inv f i))" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2011 |
by simp |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2012 |
also have "... \<le> (\<Sum>i\<in>Basis. \<bar>u \<bullet> i\<bar>) * B" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2013 |
by (simp add: B sum_distrib_right sum_mono mult_left_mono) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2014 |
also have "... \<le> DIM('b) * dist y (f x) * B" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2015 |
apply (rule mult_right_mono [OF sum_bounded_above]) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2016 |
using \<open>0 < B\<close> by (auto simp: Basis_le_norm dist_norm u_def) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2017 |
also have "... < e" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2018 |
by (metis mult.commute mult.left_commute that) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2019 |
finally show "x + (\<Sum>i\<in>Basis. (u \<bullet> i) *\<^sub>R inv f i) \<in> A" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2020 |
by (rule e) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2021 |
qed |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2022 |
qed |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2023 |
then show "\<forall>y. dist y (f x) < \<delta> \<longrightarrow> y \<in> f ` A" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2024 |
using \<open>e > 0\<close> \<open>B > 0\<close> |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2025 |
by (auto simp: \<delta>_def divide_simps mult_less_0_iff) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2026 |
qed |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2027 |
qed |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2028 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2029 |
corollary open_bijective_linear_image_eq: |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2030 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2031 |
assumes "linear f" "bij f" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2032 |
shows "open(f ` A) \<longleftrightarrow> open A" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2033 |
proof |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2034 |
assume "open(f ` A)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2035 |
then have "open(f -` (f ` A))" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2036 |
using assms by (force simp: linear_continuous_at linear_conv_bounded_linear continuous_open_vimage) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2037 |
then show "open A" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2038 |
by (simp add: assms bij_is_inj inj_vimage_image_eq) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2039 |
next |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2040 |
assume "open A" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2041 |
then show "open(f ` A)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2042 |
by (simp add: assms bij_is_surj open_surjective_linear_image) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2043 |
qed |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2044 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2045 |
corollary interior_bijective_linear_image: |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2046 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2047 |
assumes "linear f" "bij f" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2048 |
shows "interior (f ` S) = f ` interior S" (is "?lhs = ?rhs") |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2049 |
proof safe |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2050 |
fix x |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2051 |
assume x: "x \<in> ?lhs" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2052 |
then obtain T where "open T" and "x \<in> T" and "T \<subseteq> f ` S" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2053 |
by (metis interiorE) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2054 |
then show "x \<in> ?rhs" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2055 |
by (metis (no_types, hide_lams) assms subsetD interior_maximal open_bijective_linear_image_eq subset_image_iff) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2056 |
next |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2057 |
fix x |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2058 |
assume x: "x \<in> interior S" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2059 |
then show "f x \<in> interior (f ` S)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2060 |
by (meson assms imageI image_mono interiorI interior_subset open_bijective_linear_image_eq open_interior) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2061 |
qed |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2062 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2063 |
lemma interior_injective_linear_image: |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2064 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2065 |
assumes "linear f" "inj f" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2066 |
shows "interior(f ` S) = f ` (interior S)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2067 |
by (simp add: linear_injective_imp_surjective assms bijI interior_bijective_linear_image) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2068 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2069 |
lemma interior_surjective_linear_image: |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2070 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2071 |
assumes "linear f" "surj f" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2072 |
shows "interior(f ` S) = f ` (interior S)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2073 |
by (simp add: assms interior_injective_linear_image linear_surjective_imp_injective) |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2074 |
|
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2075 |
lemma interior_negations: |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2076 |
fixes S :: "'a::euclidean_space set" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2077 |
shows "interior(uminus ` S) = image uminus (interior S)" |
42cc3609fedf
moved some material from Connected.thy to more appropriate places
immler
parents:
69544
diff
changeset
|
2078 |
by (simp add: bij_uminus interior_bijective_linear_image linear_uminus) |
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
2079 |
|
69617 | 2080 |
lemma connected_linear_image: |
2081 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" |
|
2082 |
assumes "linear f" and "connected s" |
|
2083 |
shows "connected (f ` s)" |
|
2084 |
using connected_continuous_image assms linear_continuous_on linear_conv_bounded_linear by blast |
|
2085 |
||
2086 |
||
69613
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2087 |
subsection%unimportant \<open>"Isometry" (up to constant bounds) of Injective Linear Map\<close> |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2088 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2089 |
proposition injective_imp_isometric: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2090 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2091 |
assumes s: "closed s" "subspace s" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2092 |
and f: "bounded_linear f" "\<forall>x\<in>s. f x = 0 \<longrightarrow> x = 0" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2093 |
shows "\<exists>e>0. \<forall>x\<in>s. norm (f x) \<ge> e * norm x" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2094 |
proof (cases "s \<subseteq> {0::'a}") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2095 |
case True |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2096 |
have "norm x \<le> norm (f x)" if "x \<in> s" for x |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2097 |
proof - |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2098 |
from True that have "x = 0" by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2099 |
then show ?thesis by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2100 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2101 |
then show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2102 |
by (auto intro!: exI[where x=1]) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2103 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2104 |
case False |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2105 |
interpret f: bounded_linear f by fact |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2106 |
from False obtain a where a: "a \<noteq> 0" "a \<in> s" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2107 |
by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2108 |
from False have "s \<noteq> {}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2109 |
by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2110 |
let ?S = "{f x| x. x \<in> s \<and> norm x = norm a}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2111 |
let ?S' = "{x::'a. x\<in>s \<and> norm x = norm a}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2112 |
let ?S'' = "{x::'a. norm x = norm a}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2113 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2114 |
have "?S'' = frontier (cball 0 (norm a))" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2115 |
by (simp add: sphere_def dist_norm) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2116 |
then have "compact ?S''" by (metis compact_cball compact_frontier) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2117 |
moreover have "?S' = s \<inter> ?S''" by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2118 |
ultimately have "compact ?S'" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2119 |
using closed_Int_compact[of s ?S''] using s(1) by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2120 |
moreover have *:"f ` ?S' = ?S" by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2121 |
ultimately have "compact ?S" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2122 |
using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2123 |
then have "closed ?S" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2124 |
using compact_imp_closed by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2125 |
moreover from a have "?S \<noteq> {}" by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2126 |
ultimately obtain b' where "b'\<in>?S" "\<forall>y\<in>?S. norm b' \<le> norm y" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2127 |
using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2128 |
then obtain b where "b\<in>s" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2129 |
and ba: "norm b = norm a" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2130 |
and b: "\<forall>x\<in>{x \<in> s. norm x = norm a}. norm (f b) \<le> norm (f x)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2131 |
unfolding *[symmetric] unfolding image_iff by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2132 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2133 |
let ?e = "norm (f b) / norm b" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2134 |
have "norm b > 0" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2135 |
using ba and a and norm_ge_zero by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2136 |
moreover have "norm (f b) > 0" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2137 |
using f(2)[THEN bspec[where x=b], OF \<open>b\<in>s\<close>] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2138 |
using \<open>norm b >0\<close> by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2139 |
ultimately have "0 < norm (f b) / norm b" by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2140 |
moreover |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2141 |
have "norm (f b) / norm b * norm x \<le> norm (f x)" if "x\<in>s" for x |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2142 |
proof (cases "x = 0") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2143 |
case True |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2144 |
then show "norm (f b) / norm b * norm x \<le> norm (f x)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2145 |
by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2146 |
next |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2147 |
case False |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2148 |
with \<open>a \<noteq> 0\<close> have *: "0 < norm a / norm x" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2149 |
unfolding zero_less_norm_iff[symmetric] by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2150 |
have "\<forall>x\<in>s. c *\<^sub>R x \<in> s" for c |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2151 |
using s[unfolded subspace_def] by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2152 |
with \<open>x \<in> s\<close> \<open>x \<noteq> 0\<close> have "(norm a / norm x) *\<^sub>R x \<in> {x \<in> s. norm x = norm a}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2153 |
by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2154 |
with \<open>x \<noteq> 0\<close> \<open>a \<noteq> 0\<close> show "norm (f b) / norm b * norm x \<le> norm (f x)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2155 |
using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2156 |
unfolding f.scaleR and ba |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2157 |
by (auto simp: mult.commute pos_le_divide_eq pos_divide_le_eq) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2158 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2159 |
ultimately show ?thesis by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2160 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2161 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2162 |
proposition closed_injective_image_subspace: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2163 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2164 |
assumes "subspace s" "bounded_linear f" "\<forall>x\<in>s. f x = 0 \<longrightarrow> x = 0" "closed s" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2165 |
shows "closed(f ` s)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2166 |
proof - |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2167 |
obtain e where "e > 0" and e: "\<forall>x\<in>s. e * norm x \<le> norm (f x)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2168 |
using injective_imp_isometric[OF assms(4,1,2,3)] by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2169 |
show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2170 |
using complete_isometric_image[OF \<open>e>0\<close> assms(1,2) e] and assms(4) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2171 |
unfolding complete_eq_closed[symmetric] by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2172 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2173 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2174 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2175 |
subsection%unimportant \<open>Some properties of a canonical subspace\<close> |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2176 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2177 |
lemma closed_substandard: "closed {x::'a::euclidean_space. \<forall>i\<in>Basis. P i \<longrightarrow> x\<bullet>i = 0}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2178 |
(is "closed ?A") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2179 |
proof - |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2180 |
let ?D = "{i\<in>Basis. P i}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2181 |
have "closed (\<Inter>i\<in>?D. {x::'a. x\<bullet>i = 0})" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2182 |
by (simp add: closed_INT closed_Collect_eq continuous_on_inner |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2183 |
continuous_on_const continuous_on_id) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2184 |
also have "(\<Inter>i\<in>?D. {x::'a. x\<bullet>i = 0}) = ?A" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2185 |
by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2186 |
finally show "closed ?A" . |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2187 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2188 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2189 |
lemma closed_subspace: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2190 |
fixes s :: "'a::euclidean_space set" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2191 |
assumes "subspace s" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2192 |
shows "closed s" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2193 |
proof - |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2194 |
have "dim s \<le> card (Basis :: 'a set)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2195 |
using dim_subset_UNIV by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2196 |
with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \<subseteq> Basis" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2197 |
by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2198 |
let ?t = "{x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2199 |
have "\<exists>f. linear f \<and> f ` {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} = s \<and> |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2200 |
inj_on f {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2201 |
using dim_substandard[of d] t d assms |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2202 |
by (intro subspace_isomorphism[OF subspace_substandard[of "\<lambda>i. i \<notin> d"]]) (auto simp: inner_Basis) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2203 |
then obtain f where f: |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2204 |
"linear f" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2205 |
"f ` {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} = s" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2206 |
"inj_on f {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0}" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2207 |
by blast |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2208 |
interpret f: bounded_linear f |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2209 |
using f by (simp add: linear_conv_bounded_linear) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2210 |
have "x \<in> ?t \<Longrightarrow> f x = 0 \<Longrightarrow> x = 0" for x |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2211 |
using f.zero d f(3)[THEN inj_onD, of x 0] by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2212 |
moreover have "closed ?t" by (rule closed_substandard) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2213 |
moreover have "subspace ?t" by (rule subspace_substandard) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2214 |
ultimately show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2215 |
using closed_injective_image_subspace[of ?t f] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2216 |
unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2217 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2218 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2219 |
lemma complete_subspace: "subspace s \<Longrightarrow> complete s" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2220 |
for s :: "'a::euclidean_space set" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2221 |
using complete_eq_closed closed_subspace by auto |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2222 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2223 |
lemma closed_span [iff]: "closed (span s)" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2224 |
for s :: "'a::euclidean_space set" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2225 |
by (simp add: closed_subspace subspace_span) |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2226 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2227 |
lemma dim_closure [simp]: "dim (closure s) = dim s" (is "?dc = ?d") |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2228 |
for s :: "'a::euclidean_space set" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2229 |
proof - |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2230 |
have "?dc \<le> ?d" |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2231 |
using closure_minimal[OF span_superset, of s] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2232 |
using closed_subspace[OF subspace_span, of s] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2233 |
using dim_subset[of "closure s" "span s"] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2234 |
by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2235 |
then show ?thesis |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2236 |
using dim_subset[OF closure_subset, of s] |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2237 |
by simp |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2238 |
qed |
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2239 |
|
1331e57b54c6
moved material from Connected.thy to more appropriate places
immler
parents:
69611
diff
changeset
|
2240 |
|
69618 | 2241 |
subsection \<open>Set Distance\<close> |
2242 |
||
2243 |
lemma setdist_compact_closed: |
|
2244 |
fixes S :: "'a::{heine_borel,real_normed_vector} set" |
|
2245 |
assumes S: "compact S" and T: "closed T" |
|
2246 |
and "S \<noteq> {}" "T \<noteq> {}" |
|
2247 |
shows "\<exists>x \<in> S. \<exists>y \<in> T. dist x y = setdist S T" |
|
2248 |
proof - |
|
2249 |
have "(\<Union>x\<in> S. \<Union>y \<in> T. {x - y}) \<noteq> {}" |
|
2250 |
using assms by blast |
|
2251 |
then have "\<exists>x \<in> S. \<exists>y \<in> T. dist x y \<le> setdist S T" |
|
2252 |
apply (rule distance_attains_inf [where a=0, OF compact_closed_differences [OF S T]]) |
|
2253 |
apply (simp add: dist_norm le_setdist_iff) |
|
2254 |
apply blast |
|
2255 |
done |
|
2256 |
then show ?thesis |
|
2257 |
by (blast intro!: antisym [OF _ setdist_le_dist] ) |
|
2258 |
qed |
|
2259 |
||
2260 |
lemma setdist_closed_compact: |
|
2261 |
fixes S :: "'a::{heine_borel,real_normed_vector} set" |
|
2262 |
assumes S: "closed S" and T: "compact T" |
|
2263 |
and "S \<noteq> {}" "T \<noteq> {}" |
|
2264 |
shows "\<exists>x \<in> S. \<exists>y \<in> T. dist x y = setdist S T" |
|
2265 |
using setdist_compact_closed [OF T S \<open>T \<noteq> {}\<close> \<open>S \<noteq> {}\<close>] |
|
2266 |
by (metis dist_commute setdist_sym) |
|
2267 |
||
2268 |
lemma setdist_eq_0_compact_closed: |
|
2269 |
fixes S :: "'a::{heine_borel,real_normed_vector} set" |
|
2270 |
assumes S: "compact S" and T: "closed T" |
|
2271 |
shows "setdist S T = 0 \<longleftrightarrow> S = {} \<or> T = {} \<or> S \<inter> T \<noteq> {}" |
|
2272 |
apply (cases "S = {} \<or> T = {}", force) |
|
2273 |
using setdist_compact_closed [OF S T] |
|
2274 |
apply (force intro: setdist_eq_0I ) |
|
2275 |
done |
|
2276 |
||
2277 |
corollary setdist_gt_0_compact_closed: |
|
2278 |
fixes S :: "'a::{heine_borel,real_normed_vector} set" |
|
2279 |
assumes S: "compact S" and T: "closed T" |
|
2280 |
shows "setdist S T > 0 \<longleftrightarrow> (S \<noteq> {} \<and> T \<noteq> {} \<and> S \<inter> T = {})" |
|
2281 |
using setdist_pos_le [of S T] setdist_eq_0_compact_closed [OF assms] |
|
2282 |
by linarith |
|
2283 |
||
2284 |
lemma setdist_eq_0_closed_compact: |
|
2285 |
fixes S :: "'a::{heine_borel,real_normed_vector} set" |
|
2286 |
assumes S: "closed S" and T: "compact T" |
|
2287 |
shows "setdist S T = 0 \<longleftrightarrow> S = {} \<or> T = {} \<or> S \<inter> T \<noteq> {}" |
|
2288 |
using setdist_eq_0_compact_closed [OF T S] |
|
2289 |
by (metis Int_commute setdist_sym) |
|
2290 |
||
2291 |
lemma setdist_eq_0_bounded: |
|
2292 |
fixes S :: "'a::{heine_borel,real_normed_vector} set" |
|
2293 |
assumes "bounded S \<or> bounded T" |
|
2294 |
shows "setdist S T = 0 \<longleftrightarrow> S = {} \<or> T = {} \<or> closure S \<inter> closure T \<noteq> {}" |
|
2295 |
apply (cases "S = {} \<or> T = {}", force) |
|
2296 |
using setdist_eq_0_compact_closed [of "closure S" "closure T"] |
|
2297 |
setdist_eq_0_closed_compact [of "closure S" "closure T"] assms |
|
2298 |
apply (force simp add: bounded_closure compact_eq_bounded_closed) |
|
2299 |
done |
|
2300 |
||
2301 |
lemma setdist_eq_0_sing_1: |
|
2302 |
fixes S :: "'a::{heine_borel,real_normed_vector} set" |
|
2303 |
shows "setdist {x} S = 0 \<longleftrightarrow> S = {} \<or> x \<in> closure S" |
|
2304 |
by (auto simp: setdist_eq_0_bounded) |
|
2305 |
||
2306 |
lemma setdist_eq_0_sing_2: |
|
2307 |
fixes S :: "'a::{heine_borel,real_normed_vector} set" |
|
2308 |
shows "setdist S {x} = 0 \<longleftrightarrow> S = {} \<or> x \<in> closure S" |
|
2309 |
by (auto simp: setdist_eq_0_bounded) |
|
2310 |
||
2311 |
lemma setdist_neq_0_sing_1: |
|
2312 |
fixes S :: "'a::{heine_borel,real_normed_vector} set" |
|
2313 |
shows "\<lbrakk>setdist {x} S = a; a \<noteq> 0\<rbrakk> \<Longrightarrow> S \<noteq> {} \<and> x \<notin> closure S" |
|
2314 |
by (auto simp: setdist_eq_0_sing_1) |
|
2315 |
||
2316 |
lemma setdist_neq_0_sing_2: |
|
2317 |
fixes S :: "'a::{heine_borel,real_normed_vector} set" |
|
2318 |
shows "\<lbrakk>setdist S {x} = a; a \<noteq> 0\<rbrakk> \<Longrightarrow> S \<noteq> {} \<and> x \<notin> closure S" |
|
2319 |
by (auto simp: setdist_eq_0_sing_2) |
|
2320 |
||
2321 |
lemma setdist_sing_in_set: |
|
2322 |
fixes S :: "'a::{heine_borel,real_normed_vector} set" |
|
2323 |
shows "x \<in> S \<Longrightarrow> setdist {x} S = 0" |
|
2324 |
using closure_subset by (auto simp: setdist_eq_0_sing_1) |
|
2325 |
||
2326 |
lemma setdist_eq_0_closed: |
|
2327 |
fixes S :: "'a::{heine_borel,real_normed_vector} set" |
|
2328 |
shows "closed S \<Longrightarrow> (setdist {x} S = 0 \<longleftrightarrow> S = {} \<or> x \<in> S)" |
|
2329 |
by (simp add: setdist_eq_0_sing_1) |
|
2330 |
||
2331 |
lemma setdist_eq_0_closedin: |
|
2332 |
fixes S :: "'a::{heine_borel,real_normed_vector} set" |
|
2333 |
shows "\<lbrakk>closedin (subtopology euclidean U) S; x \<in> U\<rbrakk> |
|
2334 |
\<Longrightarrow> (setdist {x} S = 0 \<longleftrightarrow> S = {} \<or> x \<in> S)" |
|
2335 |
by (auto simp: closedin_limpt setdist_eq_0_sing_1 closure_def) |
|
2336 |
||
2337 |
lemma setdist_gt_0_closedin: |
|
2338 |
fixes S :: "'a::{heine_borel,real_normed_vector} set" |
|
2339 |
shows "\<lbrakk>closedin (subtopology euclidean U) S; x \<in> U; S \<noteq> {}; x \<notin> S\<rbrakk> |
|
2340 |
\<Longrightarrow> setdist {x} S > 0" |
|
2341 |
using less_eq_real_def setdist_eq_0_closedin by fastforce |
|
2342 |
||
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
2343 |
no_notation |
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
2344 |
eucl_less (infix "<e" 50) |
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
2345 |
|
33175 | 2346 |
end |