71028
|
1 |
(* Title: HOL/Analysis/Line_Segment.thy
|
|
2 |
Author: L C Paulson, University of Cambridge
|
|
3 |
Author: Robert Himmelmann, TU Muenchen
|
|
4 |
Author: Bogdan Grechuk, University of Edinburgh
|
|
5 |
Author: Armin Heller, TU Muenchen
|
|
6 |
Author: Johannes Hoelzl, TU Muenchen
|
|
7 |
*)
|
|
8 |
|
|
9 |
section \<open>Line Segment\<close>
|
|
10 |
|
|
11 |
theory Line_Segment
|
|
12 |
imports
|
|
13 |
Convex
|
|
14 |
Topology_Euclidean_Space
|
|
15 |
begin
|
|
16 |
|
|
17 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Topological Properties of Convex Sets and Functions\<close>
|
|
18 |
|
|
19 |
lemma convex_supp_sum:
|
|
20 |
assumes "convex S" and 1: "supp_sum u I = 1"
|
|
21 |
and "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> u i \<and> (u i = 0 \<or> f i \<in> S)"
|
|
22 |
shows "supp_sum (\<lambda>i. u i *\<^sub>R f i) I \<in> S"
|
|
23 |
proof -
|
|
24 |
have fin: "finite {i \<in> I. u i \<noteq> 0}"
|
|
25 |
using 1 sum.infinite by (force simp: supp_sum_def support_on_def)
|
|
26 |
then have eq: "supp_sum (\<lambda>i. u i *\<^sub>R f i) I = sum (\<lambda>i. u i *\<^sub>R f i) {i \<in> I. u i \<noteq> 0}"
|
|
27 |
by (force intro: sum.mono_neutral_left simp: supp_sum_def support_on_def)
|
|
28 |
show ?thesis
|
|
29 |
apply (simp add: eq)
|
|
30 |
apply (rule convex_sum [OF fin \<open>convex S\<close>])
|
|
31 |
using 1 assms apply (auto simp: supp_sum_def support_on_def)
|
|
32 |
done
|
|
33 |
qed
|
|
34 |
|
|
35 |
lemma closure_bounded_linear_image_subset:
|
|
36 |
assumes f: "bounded_linear f"
|
|
37 |
shows "f ` closure S \<subseteq> closure (f ` S)"
|
|
38 |
using linear_continuous_on [OF f] closed_closure closure_subset
|
|
39 |
by (rule image_closure_subset)
|
|
40 |
|
|
41 |
lemma closure_linear_image_subset:
|
|
42 |
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::real_normed_vector"
|
|
43 |
assumes "linear f"
|
|
44 |
shows "f ` (closure S) \<subseteq> closure (f ` S)"
|
|
45 |
using assms unfolding linear_conv_bounded_linear
|
|
46 |
by (rule closure_bounded_linear_image_subset)
|
|
47 |
|
|
48 |
lemma closed_injective_linear_image:
|
|
49 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
|
|
50 |
assumes S: "closed S" and f: "linear f" "inj f"
|
|
51 |
shows "closed (f ` S)"
|
|
52 |
proof -
|
|
53 |
obtain g where g: "linear g" "g \<circ> f = id"
|
|
54 |
using linear_injective_left_inverse [OF f] by blast
|
|
55 |
then have confg: "continuous_on (range f) g"
|
|
56 |
using linear_continuous_on linear_conv_bounded_linear by blast
|
|
57 |
have [simp]: "g ` f ` S = S"
|
|
58 |
using g by (simp add: image_comp)
|
|
59 |
have cgf: "closed (g ` f ` S)"
|
|
60 |
by (simp add: \<open>g \<circ> f = id\<close> S image_comp)
|
|
61 |
have [simp]: "(range f \<inter> g -` S) = f ` S"
|
|
62 |
using g unfolding o_def id_def image_def by auto metis+
|
|
63 |
show ?thesis
|
|
64 |
proof (rule closedin_closed_trans [of "range f"])
|
|
65 |
show "closedin (top_of_set (range f)) (f ` S)"
|
|
66 |
using continuous_closedin_preimage [OF confg cgf] by simp
|
|
67 |
show "closed (range f)"
|
|
68 |
apply (rule closed_injective_image_subspace)
|
|
69 |
using f apply (auto simp: linear_linear linear_injective_0)
|
|
70 |
done
|
|
71 |
qed
|
|
72 |
qed
|
|
73 |
|
|
74 |
lemma closed_injective_linear_image_eq:
|
|
75 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
|
|
76 |
assumes f: "linear f" "inj f"
|
|
77 |
shows "(closed(image f s) \<longleftrightarrow> closed s)"
|
|
78 |
by (metis closed_injective_linear_image closure_eq closure_linear_image_subset closure_subset_eq f(1) f(2) inj_image_subset_iff)
|
|
79 |
|
|
80 |
lemma closure_injective_linear_image:
|
|
81 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
|
|
82 |
shows "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)"
|
|
83 |
apply (rule subset_antisym)
|
|
84 |
apply (simp add: closure_linear_image_subset)
|
|
85 |
by (simp add: closure_minimal closed_injective_linear_image closure_subset image_mono)
|
|
86 |
|
|
87 |
lemma closure_bounded_linear_image:
|
|
88 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
|
|
89 |
shows "\<lbrakk>linear f; bounded S\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)"
|
|
90 |
apply (rule subset_antisym, simp add: closure_linear_image_subset)
|
|
91 |
apply (rule closure_minimal, simp add: closure_subset image_mono)
|
|
92 |
by (meson bounded_closure closed_closure compact_continuous_image compact_eq_bounded_closed linear_continuous_on linear_conv_bounded_linear)
|
|
93 |
|
|
94 |
lemma closure_scaleR:
|
|
95 |
fixes S :: "'a::real_normed_vector set"
|
|
96 |
shows "((*\<^sub>R) c) ` (closure S) = closure (((*\<^sub>R) c) ` S)"
|
|
97 |
proof
|
|
98 |
show "((*\<^sub>R) c) ` (closure S) \<subseteq> closure (((*\<^sub>R) c) ` S)"
|
|
99 |
using bounded_linear_scaleR_right
|
|
100 |
by (rule closure_bounded_linear_image_subset)
|
|
101 |
show "closure (((*\<^sub>R) c) ` S) \<subseteq> ((*\<^sub>R) c) ` (closure S)"
|
|
102 |
by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure)
|
|
103 |
qed
|
|
104 |
|
|
105 |
lemma sphere_eq_empty [simp]:
|
|
106 |
fixes a :: "'a::{real_normed_vector, perfect_space}"
|
|
107 |
shows "sphere a r = {} \<longleftrightarrow> r < 0"
|
|
108 |
by (auto simp: sphere_def dist_norm) (metis dist_norm le_less_linear vector_choose_dist)
|
|
109 |
|
|
110 |
lemma cone_closure:
|
|
111 |
fixes S :: "'a::real_normed_vector set"
|
|
112 |
assumes "cone S"
|
|
113 |
shows "cone (closure S)"
|
|
114 |
proof (cases "S = {}")
|
|
115 |
case True
|
|
116 |
then show ?thesis by auto
|
|
117 |
next
|
|
118 |
case False
|
|
119 |
then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (*\<^sub>R) c ` S = S)"
|
|
120 |
using cone_iff[of S] assms by auto
|
|
121 |
then have "0 \<in> closure S \<and> (\<forall>c. c > 0 \<longrightarrow> (*\<^sub>R) c ` closure S = closure S)"
|
|
122 |
using closure_subset by (auto simp: closure_scaleR)
|
|
123 |
then show ?thesis
|
|
124 |
using False cone_iff[of "closure S"] by auto
|
|
125 |
qed
|
|
126 |
|
|
127 |
|
|
128 |
corollary component_complement_connected:
|
|
129 |
fixes S :: "'a::real_normed_vector set"
|
|
130 |
assumes "connected S" "C \<in> components (-S)"
|
|
131 |
shows "connected(-C)"
|
|
132 |
using component_diff_connected [of S UNIV] assms
|
|
133 |
by (auto simp: Compl_eq_Diff_UNIV)
|
|
134 |
|
|
135 |
proposition clopen:
|
|
136 |
fixes S :: "'a :: real_normed_vector set"
|
|
137 |
shows "closed S \<and> open S \<longleftrightarrow> S = {} \<or> S = UNIV"
|
|
138 |
by (force intro!: connected_UNIV [unfolded connected_clopen, rule_format])
|
|
139 |
|
|
140 |
corollary compact_open:
|
|
141 |
fixes S :: "'a :: euclidean_space set"
|
|
142 |
shows "compact S \<and> open S \<longleftrightarrow> S = {}"
|
|
143 |
by (auto simp: compact_eq_bounded_closed clopen)
|
|
144 |
|
|
145 |
corollary finite_imp_not_open:
|
|
146 |
fixes S :: "'a::{real_normed_vector, perfect_space} set"
|
|
147 |
shows "\<lbrakk>finite S; open S\<rbrakk> \<Longrightarrow> S={}"
|
|
148 |
using clopen [of S] finite_imp_closed not_bounded_UNIV by blast
|
|
149 |
|
|
150 |
corollary empty_interior_finite:
|
|
151 |
fixes S :: "'a::{real_normed_vector, perfect_space} set"
|
|
152 |
shows "finite S \<Longrightarrow> interior S = {}"
|
|
153 |
by (metis interior_subset finite_subset open_interior [of S] finite_imp_not_open)
|
|
154 |
|
|
155 |
text \<open>Balls, being convex, are connected.\<close>
|
|
156 |
|
|
157 |
lemma convex_local_global_minimum:
|
|
158 |
fixes s :: "'a::real_normed_vector set"
|
|
159 |
assumes "e > 0"
|
|
160 |
and "convex_on s f"
|
|
161 |
and "ball x e \<subseteq> s"
|
|
162 |
and "\<forall>y\<in>ball x e. f x \<le> f y"
|
|
163 |
shows "\<forall>y\<in>s. f x \<le> f y"
|
|
164 |
proof (rule ccontr)
|
|
165 |
have "x \<in> s" using assms(1,3) by auto
|
|
166 |
assume "\<not> ?thesis"
|
|
167 |
then obtain y where "y\<in>s" and y: "f x > f y" by auto
|
|
168 |
then have xy: "0 < dist x y" by auto
|
|
169 |
then obtain u where "0 < u" "u \<le> 1" and u: "u < e / dist x y"
|
|
170 |
using field_lbound_gt_zero[of 1 "e / dist x y"] xy \<open>e>0\<close> by auto
|
|
171 |
then have "f ((1-u) *\<^sub>R x + u *\<^sub>R y) \<le> (1-u) * f x + u * f y"
|
|
172 |
using \<open>x\<in>s\<close> \<open>y\<in>s\<close>
|
|
173 |
using assms(2)[unfolded convex_on_def,
|
|
174 |
THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x="1-u"]]
|
|
175 |
by auto
|
|
176 |
moreover
|
|
177 |
have *: "x - ((1 - u) *\<^sub>R x + u *\<^sub>R y) = u *\<^sub>R (x - y)"
|
|
178 |
by (simp add: algebra_simps)
|
|
179 |
have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> ball x e"
|
|
180 |
unfolding mem_ball dist_norm
|
|
181 |
unfolding * and norm_scaleR and abs_of_pos[OF \<open>0<u\<close>]
|
|
182 |
unfolding dist_norm[symmetric]
|
|
183 |
using u
|
|
184 |
unfolding pos_less_divide_eq[OF xy]
|
|
185 |
by auto
|
|
186 |
then have "f x \<le> f ((1 - u) *\<^sub>R x + u *\<^sub>R y)"
|
|
187 |
using assms(4) by auto
|
|
188 |
ultimately show False
|
|
189 |
using mult_strict_left_mono[OF y \<open>u>0\<close>]
|
|
190 |
unfolding left_diff_distrib
|
|
191 |
by auto
|
|
192 |
qed
|
|
193 |
|
|
194 |
lemma convex_ball [iff]:
|
|
195 |
fixes x :: "'a::real_normed_vector"
|
|
196 |
shows "convex (ball x e)"
|
|
197 |
proof (auto simp: convex_def)
|
|
198 |
fix y z
|
|
199 |
assume yz: "dist x y < e" "dist x z < e"
|
|
200 |
fix u v :: real
|
|
201 |
assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1"
|
|
202 |
have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z"
|
|
203 |
using uv yz
|
|
204 |
using convex_on_dist [of "ball x e" x, unfolded convex_on_def,
|
|
205 |
THEN bspec[where x=y], THEN bspec[where x=z]]
|
|
206 |
by auto
|
|
207 |
then show "dist x (u *\<^sub>R y + v *\<^sub>R z) < e"
|
|
208 |
using convex_bound_lt[OF yz uv] by auto
|
|
209 |
qed
|
|
210 |
|
|
211 |
lemma convex_cball [iff]:
|
|
212 |
fixes x :: "'a::real_normed_vector"
|
|
213 |
shows "convex (cball x e)"
|
|
214 |
proof -
|
|
215 |
{
|
|
216 |
fix y z
|
|
217 |
assume yz: "dist x y \<le> e" "dist x z \<le> e"
|
|
218 |
fix u v :: real
|
|
219 |
assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1"
|
|
220 |
have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z"
|
|
221 |
using uv yz
|
|
222 |
using convex_on_dist [of "cball x e" x, unfolded convex_on_def,
|
|
223 |
THEN bspec[where x=y], THEN bspec[where x=z]]
|
|
224 |
by auto
|
|
225 |
then have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> e"
|
|
226 |
using convex_bound_le[OF yz uv] by auto
|
|
227 |
}
|
|
228 |
then show ?thesis by (auto simp: convex_def Ball_def)
|
|
229 |
qed
|
|
230 |
|
|
231 |
lemma connected_ball [iff]:
|
|
232 |
fixes x :: "'a::real_normed_vector"
|
|
233 |
shows "connected (ball x e)"
|
|
234 |
using convex_connected convex_ball by auto
|
|
235 |
|
|
236 |
lemma connected_cball [iff]:
|
|
237 |
fixes x :: "'a::real_normed_vector"
|
|
238 |
shows "connected (cball x e)"
|
|
239 |
using convex_connected convex_cball by auto
|
|
240 |
|
|
241 |
lemma bounded_convex_hull:
|
|
242 |
fixes s :: "'a::real_normed_vector set"
|
|
243 |
assumes "bounded s"
|
|
244 |
shows "bounded (convex hull s)"
|
|
245 |
proof -
|
|
246 |
from assms obtain B where B: "\<forall>x\<in>s. norm x \<le> B"
|
|
247 |
unfolding bounded_iff by auto
|
|
248 |
show ?thesis
|
|
249 |
apply (rule bounded_subset[OF bounded_cball, of _ 0 B])
|
|
250 |
unfolding subset_hull[of convex, OF convex_cball]
|
|
251 |
unfolding subset_eq mem_cball dist_norm using B
|
|
252 |
apply auto
|
|
253 |
done
|
|
254 |
qed
|
|
255 |
|
|
256 |
lemma finite_imp_bounded_convex_hull:
|
|
257 |
fixes s :: "'a::real_normed_vector set"
|
|
258 |
shows "finite s \<Longrightarrow> bounded (convex hull s)"
|
|
259 |
using bounded_convex_hull finite_imp_bounded
|
|
260 |
by auto
|
|
261 |
|
|
262 |
|
|
263 |
section \<open>Line Segments\<close>
|
|
264 |
|
|
265 |
subsection \<open>Midpoint\<close>
|
|
266 |
|
|
267 |
definition\<^marker>\<open>tag important\<close> midpoint :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a"
|
|
268 |
where "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)"
|
|
269 |
|
|
270 |
lemma midpoint_idem [simp]: "midpoint x x = x"
|
|
271 |
unfolding midpoint_def by simp
|
|
272 |
|
|
273 |
lemma midpoint_sym: "midpoint a b = midpoint b a"
|
|
274 |
unfolding midpoint_def by (auto simp add: scaleR_right_distrib)
|
|
275 |
|
|
276 |
lemma midpoint_eq_iff: "midpoint a b = c \<longleftrightarrow> a + b = c + c"
|
|
277 |
proof -
|
|
278 |
have "midpoint a b = c \<longleftrightarrow> scaleR 2 (midpoint a b) = scaleR 2 c"
|
|
279 |
by simp
|
|
280 |
then show ?thesis
|
|
281 |
unfolding midpoint_def scaleR_2 [symmetric] by simp
|
|
282 |
qed
|
|
283 |
|
|
284 |
lemma
|
|
285 |
fixes a::real
|
|
286 |
assumes "a \<le> b" shows ge_midpoint_1: "a \<le> midpoint a b"
|
|
287 |
and le_midpoint_1: "midpoint a b \<le> b"
|
|
288 |
by (simp_all add: midpoint_def assms)
|
|
289 |
|
|
290 |
lemma dist_midpoint:
|
|
291 |
fixes a b :: "'a::real_normed_vector" shows
|
|
292 |
"dist a (midpoint a b) = (dist a b) / 2" (is ?t1)
|
|
293 |
"dist b (midpoint a b) = (dist a b) / 2" (is ?t2)
|
|
294 |
"dist (midpoint a b) a = (dist a b) / 2" (is ?t3)
|
|
295 |
"dist (midpoint a b) b = (dist a b) / 2" (is ?t4)
|
|
296 |
proof -
|
|
297 |
have *: "\<And>x y::'a. 2 *\<^sub>R x = - y \<Longrightarrow> norm x = (norm y) / 2"
|
|
298 |
unfolding equation_minus_iff by auto
|
|
299 |
have **: "\<And>x y::'a. 2 *\<^sub>R x = y \<Longrightarrow> norm x = (norm y) / 2"
|
|
300 |
by auto
|
|
301 |
note scaleR_right_distrib [simp]
|
|
302 |
show ?t1
|
|
303 |
unfolding midpoint_def dist_norm
|
|
304 |
apply (rule **)
|
|
305 |
apply (simp add: scaleR_right_diff_distrib)
|
|
306 |
apply (simp add: scaleR_2)
|
|
307 |
done
|
|
308 |
show ?t2
|
|
309 |
unfolding midpoint_def dist_norm
|
|
310 |
apply (rule *)
|
|
311 |
apply (simp add: scaleR_right_diff_distrib)
|
|
312 |
apply (simp add: scaleR_2)
|
|
313 |
done
|
|
314 |
show ?t3
|
|
315 |
unfolding midpoint_def dist_norm
|
|
316 |
apply (rule *)
|
|
317 |
apply (simp add: scaleR_right_diff_distrib)
|
|
318 |
apply (simp add: scaleR_2)
|
|
319 |
done
|
|
320 |
show ?t4
|
|
321 |
unfolding midpoint_def dist_norm
|
|
322 |
apply (rule **)
|
|
323 |
apply (simp add: scaleR_right_diff_distrib)
|
|
324 |
apply (simp add: scaleR_2)
|
|
325 |
done
|
|
326 |
qed
|
|
327 |
|
|
328 |
lemma midpoint_eq_endpoint [simp]:
|
|
329 |
"midpoint a b = a \<longleftrightarrow> a = b"
|
|
330 |
"midpoint a b = b \<longleftrightarrow> a = b"
|
|
331 |
unfolding midpoint_eq_iff by auto
|
|
332 |
|
|
333 |
lemma midpoint_plus_self [simp]: "midpoint a b + midpoint a b = a + b"
|
|
334 |
using midpoint_eq_iff by metis
|
|
335 |
|
|
336 |
lemma midpoint_linear_image:
|
|
337 |
"linear f \<Longrightarrow> midpoint(f a)(f b) = f(midpoint a b)"
|
|
338 |
by (simp add: linear_iff midpoint_def)
|
|
339 |
|
|
340 |
|
|
341 |
subsection \<open>Line segments\<close>
|
|
342 |
|
|
343 |
definition\<^marker>\<open>tag important\<close> closed_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set"
|
|
344 |
where "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \<le> u \<and> u \<le> 1}"
|
|
345 |
|
|
346 |
definition\<^marker>\<open>tag important\<close> open_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set" where
|
|
347 |
"open_segment a b \<equiv> closed_segment a b - {a,b}"
|
|
348 |
|
|
349 |
lemmas segment = open_segment_def closed_segment_def
|
|
350 |
|
|
351 |
lemma in_segment:
|
|
352 |
"x \<in> closed_segment a b \<longleftrightarrow> (\<exists>u. 0 \<le> u \<and> u \<le> 1 \<and> x = (1 - u) *\<^sub>R a + u *\<^sub>R b)"
|
|
353 |
"x \<in> open_segment a b \<longleftrightarrow> a \<noteq> b \<and> (\<exists>u. 0 < u \<and> u < 1 \<and> x = (1 - u) *\<^sub>R a + u *\<^sub>R b)"
|
|
354 |
using less_eq_real_def by (auto simp: segment algebra_simps)
|
|
355 |
|
|
356 |
lemma closed_segment_linear_image:
|
|
357 |
"closed_segment (f a) (f b) = f ` (closed_segment a b)" if "linear f"
|
|
358 |
proof -
|
|
359 |
interpret linear f by fact
|
|
360 |
show ?thesis
|
|
361 |
by (force simp add: in_segment add scale)
|
|
362 |
qed
|
|
363 |
|
|
364 |
lemma open_segment_linear_image:
|
|
365 |
"\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> open_segment (f a) (f b) = f ` (open_segment a b)"
|
|
366 |
by (force simp: open_segment_def closed_segment_linear_image inj_on_def)
|
|
367 |
|
|
368 |
lemma closed_segment_translation:
|
|
369 |
"closed_segment (c + a) (c + b) = image (\<lambda>x. c + x) (closed_segment a b)"
|
|
370 |
apply safe
|
|
371 |
apply (rule_tac x="x-c" in image_eqI)
|
|
372 |
apply (auto simp: in_segment algebra_simps)
|
|
373 |
done
|
|
374 |
|
|
375 |
lemma open_segment_translation:
|
|
376 |
"open_segment (c + a) (c + b) = image (\<lambda>x. c + x) (open_segment a b)"
|
|
377 |
by (simp add: open_segment_def closed_segment_translation translation_diff)
|
|
378 |
|
|
379 |
lemma closed_segment_of_real:
|
|
380 |
"closed_segment (of_real x) (of_real y) = of_real ` closed_segment x y"
|
|
381 |
apply (auto simp: image_iff in_segment scaleR_conv_of_real)
|
|
382 |
apply (rule_tac x="(1-u)*x + u*y" in bexI)
|
|
383 |
apply (auto simp: in_segment)
|
|
384 |
done
|
|
385 |
|
|
386 |
lemma open_segment_of_real:
|
|
387 |
"open_segment (of_real x) (of_real y) = of_real ` open_segment x y"
|
|
388 |
apply (auto simp: image_iff in_segment scaleR_conv_of_real)
|
|
389 |
apply (rule_tac x="(1-u)*x + u*y" in bexI)
|
|
390 |
apply (auto simp: in_segment)
|
|
391 |
done
|
|
392 |
|
|
393 |
lemma closed_segment_Reals:
|
|
394 |
"\<lbrakk>x \<in> Reals; y \<in> Reals\<rbrakk> \<Longrightarrow> closed_segment x y = of_real ` closed_segment (Re x) (Re y)"
|
|
395 |
by (metis closed_segment_of_real of_real_Re)
|
|
396 |
|
|
397 |
lemma open_segment_Reals:
|
|
398 |
"\<lbrakk>x \<in> Reals; y \<in> Reals\<rbrakk> \<Longrightarrow> open_segment x y = of_real ` open_segment (Re x) (Re y)"
|
|
399 |
by (metis open_segment_of_real of_real_Re)
|
|
400 |
|
|
401 |
lemma open_segment_PairD:
|
|
402 |
"(x, x') \<in> open_segment (a, a') (b, b')
|
|
403 |
\<Longrightarrow> (x \<in> open_segment a b \<or> a = b) \<and> (x' \<in> open_segment a' b' \<or> a' = b')"
|
|
404 |
by (auto simp: in_segment)
|
|
405 |
|
|
406 |
lemma closed_segment_PairD:
|
|
407 |
"(x, x') \<in> closed_segment (a, a') (b, b') \<Longrightarrow> x \<in> closed_segment a b \<and> x' \<in> closed_segment a' b'"
|
|
408 |
by (auto simp: closed_segment_def)
|
|
409 |
|
|
410 |
lemma closed_segment_translation_eq [simp]:
|
|
411 |
"d + x \<in> closed_segment (d + a) (d + b) \<longleftrightarrow> x \<in> closed_segment a b"
|
|
412 |
proof -
|
|
413 |
have *: "\<And>d x a b. x \<in> closed_segment a b \<Longrightarrow> d + x \<in> closed_segment (d + a) (d + b)"
|
|
414 |
apply (simp add: closed_segment_def)
|
|
415 |
apply (erule ex_forward)
|
|
416 |
apply (simp add: algebra_simps)
|
|
417 |
done
|
|
418 |
show ?thesis
|
|
419 |
using * [where d = "-d"] *
|
|
420 |
by (fastforce simp add:)
|
|
421 |
qed
|
|
422 |
|
|
423 |
lemma open_segment_translation_eq [simp]:
|
|
424 |
"d + x \<in> open_segment (d + a) (d + b) \<longleftrightarrow> x \<in> open_segment a b"
|
|
425 |
by (simp add: open_segment_def)
|
|
426 |
|
|
427 |
lemma of_real_closed_segment [simp]:
|
|
428 |
"of_real x \<in> closed_segment (of_real a) (of_real b) \<longleftrightarrow> x \<in> closed_segment a b"
|
|
429 |
apply (auto simp: in_segment scaleR_conv_of_real elim!: ex_forward)
|
|
430 |
using of_real_eq_iff by fastforce
|
|
431 |
|
|
432 |
lemma of_real_open_segment [simp]:
|
|
433 |
"of_real x \<in> open_segment (of_real a) (of_real b) \<longleftrightarrow> x \<in> open_segment a b"
|
|
434 |
apply (auto simp: in_segment scaleR_conv_of_real elim!: ex_forward del: exE)
|
|
435 |
using of_real_eq_iff by fastforce
|
|
436 |
|
|
437 |
lemma convex_contains_segment:
|
|
438 |
"convex S \<longleftrightarrow> (\<forall>a\<in>S. \<forall>b\<in>S. closed_segment a b \<subseteq> S)"
|
|
439 |
unfolding convex_alt closed_segment_def by auto
|
|
440 |
|
|
441 |
lemma closed_segment_in_Reals:
|
|
442 |
"\<lbrakk>x \<in> closed_segment a b; a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> x \<in> Reals"
|
|
443 |
by (meson subsetD convex_Reals convex_contains_segment)
|
|
444 |
|
|
445 |
lemma open_segment_in_Reals:
|
|
446 |
"\<lbrakk>x \<in> open_segment a b; a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> x \<in> Reals"
|
|
447 |
by (metis Diff_iff closed_segment_in_Reals open_segment_def)
|
|
448 |
|
|
449 |
lemma closed_segment_subset: "\<lbrakk>x \<in> S; y \<in> S; convex S\<rbrakk> \<Longrightarrow> closed_segment x y \<subseteq> S"
|
|
450 |
by (simp add: convex_contains_segment)
|
|
451 |
|
|
452 |
lemma closed_segment_subset_convex_hull:
|
|
453 |
"\<lbrakk>x \<in> convex hull S; y \<in> convex hull S\<rbrakk> \<Longrightarrow> closed_segment x y \<subseteq> convex hull S"
|
|
454 |
using convex_contains_segment by blast
|
|
455 |
|
|
456 |
lemma segment_convex_hull:
|
|
457 |
"closed_segment a b = convex hull {a,b}"
|
|
458 |
proof -
|
|
459 |
have *: "\<And>x. {x} \<noteq> {}" by auto
|
|
460 |
show ?thesis
|
|
461 |
unfolding segment convex_hull_insert[OF *] convex_hull_singleton
|
|
462 |
by (safe; rule_tac x="1 - u" in exI; force)
|
|
463 |
qed
|
|
464 |
|
|
465 |
lemma open_closed_segment: "u \<in> open_segment w z \<Longrightarrow> u \<in> closed_segment w z"
|
|
466 |
by (auto simp add: closed_segment_def open_segment_def)
|
|
467 |
|
|
468 |
lemma segment_open_subset_closed:
|
|
469 |
"open_segment a b \<subseteq> closed_segment a b"
|
|
470 |
by (auto simp: closed_segment_def open_segment_def)
|
|
471 |
|
|
472 |
lemma bounded_closed_segment:
|
|
473 |
fixes a :: "'a::real_normed_vector" shows "bounded (closed_segment a b)"
|
|
474 |
by (rule boundedI[where B="max (norm a) (norm b)"])
|
|
475 |
(auto simp: closed_segment_def max_def convex_bound_le intro!: norm_triangle_le)
|
|
476 |
|
|
477 |
lemma bounded_open_segment:
|
|
478 |
fixes a :: "'a::real_normed_vector" shows "bounded (open_segment a b)"
|
|
479 |
by (rule bounded_subset [OF bounded_closed_segment segment_open_subset_closed])
|
|
480 |
|
|
481 |
lemmas bounded_segment = bounded_closed_segment open_closed_segment
|
|
482 |
|
|
483 |
lemma ends_in_segment [iff]: "a \<in> closed_segment a b" "b \<in> closed_segment a b"
|
|
484 |
unfolding segment_convex_hull
|
|
485 |
by (auto intro!: hull_subset[unfolded subset_eq, rule_format])
|
|
486 |
|
|
487 |
|
|
488 |
lemma eventually_closed_segment:
|
|
489 |
fixes x0::"'a::real_normed_vector"
|
|
490 |
assumes "open X0" "x0 \<in> X0"
|
|
491 |
shows "\<forall>\<^sub>F x in at x0 within U. closed_segment x0 x \<subseteq> X0"
|
|
492 |
proof -
|
|
493 |
from openE[OF assms]
|
|
494 |
obtain e where e: "0 < e" "ball x0 e \<subseteq> X0" .
|
|
495 |
then have "\<forall>\<^sub>F x in at x0 within U. x \<in> ball x0 e"
|
|
496 |
by (auto simp: dist_commute eventually_at)
|
|
497 |
then show ?thesis
|
|
498 |
proof eventually_elim
|
|
499 |
case (elim x)
|
|
500 |
have "x0 \<in> ball x0 e" using \<open>e > 0\<close> by simp
|
|
501 |
from convex_ball[unfolded convex_contains_segment, rule_format, OF this elim]
|
|
502 |
have "closed_segment x0 x \<subseteq> ball x0 e" .
|
|
503 |
also note \<open>\<dots> \<subseteq> X0\<close>
|
|
504 |
finally show ?case .
|
|
505 |
qed
|
|
506 |
qed
|
|
507 |
|
|
508 |
lemma closed_segment_commute: "closed_segment a b = closed_segment b a"
|
|
509 |
proof -
|
|
510 |
have "{a, b} = {b, a}" by auto
|
|
511 |
thus ?thesis
|
|
512 |
by (simp add: segment_convex_hull)
|
|
513 |
qed
|
|
514 |
|
|
515 |
lemma segment_bound1:
|
|
516 |
assumes "x \<in> closed_segment a b"
|
|
517 |
shows "norm (x - a) \<le> norm (b - a)"
|
|
518 |
proof -
|
|
519 |
obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1"
|
|
520 |
using assms by (auto simp add: closed_segment_def)
|
|
521 |
then show "norm (x - a) \<le> norm (b - a)"
|
|
522 |
apply clarify
|
|
523 |
apply (auto simp: algebra_simps)
|
|
524 |
apply (simp add: scaleR_diff_right [symmetric] mult_left_le_one_le)
|
|
525 |
done
|
|
526 |
qed
|
|
527 |
|
|
528 |
lemma segment_bound:
|
|
529 |
assumes "x \<in> closed_segment a b"
|
|
530 |
shows "norm (x - a) \<le> norm (b - a)" "norm (x - b) \<le> norm (b - a)"
|
|
531 |
apply (simp add: assms segment_bound1)
|
|
532 |
by (metis assms closed_segment_commute dist_commute dist_norm segment_bound1)
|
|
533 |
|
|
534 |
lemma open_segment_commute: "open_segment a b = open_segment b a"
|
|
535 |
proof -
|
|
536 |
have "{a, b} = {b, a}" by auto
|
|
537 |
thus ?thesis
|
|
538 |
by (simp add: closed_segment_commute open_segment_def)
|
|
539 |
qed
|
|
540 |
|
|
541 |
lemma closed_segment_idem [simp]: "closed_segment a a = {a}"
|
|
542 |
unfolding segment by (auto simp add: algebra_simps)
|
|
543 |
|
|
544 |
lemma open_segment_idem [simp]: "open_segment a a = {}"
|
|
545 |
by (simp add: open_segment_def)
|
|
546 |
|
|
547 |
lemma closed_segment_eq_open: "closed_segment a b = open_segment a b \<union> {a,b}"
|
|
548 |
using open_segment_def by auto
|
|
549 |
|
|
550 |
lemma convex_contains_open_segment:
|
|
551 |
"convex s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. open_segment a b \<subseteq> s)"
|
|
552 |
by (simp add: convex_contains_segment closed_segment_eq_open)
|
|
553 |
|
|
554 |
lemma closed_segment_eq_real_ivl1:
|
|
555 |
fixes a b::real
|
|
556 |
assumes "a \<le> b"
|
|
557 |
shows "closed_segment a b = {a .. b}"
|
|
558 |
proof safe
|
|
559 |
fix x
|
|
560 |
assume "x \<in> closed_segment a b"
|
|
561 |
then obtain u where u: "0 \<le> u" "u \<le> 1" and x_def: "x = (1 - u) * a + u * b"
|
|
562 |
by (auto simp: closed_segment_def)
|
|
563 |
have "u * a \<le> u * b" "(1 - u) * a \<le> (1 - u) * b"
|
|
564 |
by (auto intro!: mult_left_mono u assms)
|
|
565 |
then show "x \<in> {a .. b}"
|
|
566 |
unfolding x_def by (auto simp: algebra_simps)
|
|
567 |
qed (auto simp: closed_segment_def divide_simps algebra_simps
|
|
568 |
intro!: exI[where x="(x - a) / (b - a)" for x])
|
|
569 |
|
|
570 |
lemma closed_segment_eq_real_ivl:
|
|
571 |
fixes a b::real
|
|
572 |
shows "closed_segment a b = (if a \<le> b then {a .. b} else {b .. a})"
|
|
573 |
using closed_segment_eq_real_ivl1[of a b] closed_segment_eq_real_ivl1[of b a]
|
|
574 |
by (auto simp: closed_segment_commute)
|
|
575 |
|
|
576 |
lemma open_segment_eq_real_ivl:
|
|
577 |
fixes a b::real
|
|
578 |
shows "open_segment a b = (if a \<le> b then {a<..<b} else {b<..<a})"
|
|
579 |
by (auto simp: closed_segment_eq_real_ivl open_segment_def split: if_split_asm)
|
|
580 |
|
|
581 |
lemma closed_segment_real_eq:
|
|
582 |
fixes u::real shows "closed_segment u v = (\<lambda>x. (v - u) * x + u) ` {0..1}"
|
|
583 |
by (simp add: add.commute [of u] image_affinity_atLeastAtMost [where c=u] closed_segment_eq_real_ivl)
|
|
584 |
|
|
585 |
lemma dist_in_closed_segment:
|
|
586 |
fixes a :: "'a :: euclidean_space"
|
|
587 |
assumes "x \<in> closed_segment a b"
|
|
588 |
shows "dist x a \<le> dist a b \<and> dist x b \<le> dist a b"
|
|
589 |
proof (intro conjI)
|
|
590 |
obtain u where u: "0 \<le> u" "u \<le> 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b"
|
|
591 |
using assms by (force simp: in_segment algebra_simps)
|
|
592 |
have "dist x a = u * dist a b"
|
|
593 |
apply (simp add: dist_norm algebra_simps x)
|
|
594 |
by (metis \<open>0 \<le> u\<close> abs_of_nonneg norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib)
|
|
595 |
also have "... \<le> dist a b"
|
|
596 |
by (simp add: mult_left_le_one_le u)
|
|
597 |
finally show "dist x a \<le> dist a b" .
|
|
598 |
have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)"
|
|
599 |
by (simp add: dist_norm algebra_simps x)
|
|
600 |
also have "... = (1-u) * dist a b"
|
|
601 |
proof -
|
|
602 |
have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)"
|
|
603 |
using \<open>u \<le> 1\<close> by force
|
|
604 |
then show ?thesis
|
|
605 |
by (simp add: dist_norm real_vector.scale_right_diff_distrib)
|
|
606 |
qed
|
|
607 |
also have "... \<le> dist a b"
|
|
608 |
by (simp add: mult_left_le_one_le u)
|
|
609 |
finally show "dist x b \<le> dist a b" .
|
|
610 |
qed
|
|
611 |
|
|
612 |
lemma dist_in_open_segment:
|
|
613 |
fixes a :: "'a :: euclidean_space"
|
|
614 |
assumes "x \<in> open_segment a b"
|
|
615 |
shows "dist x a < dist a b \<and> dist x b < dist a b"
|
|
616 |
proof (intro conjI)
|
|
617 |
obtain u where u: "0 < u" "u < 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b"
|
|
618 |
using assms by (force simp: in_segment algebra_simps)
|
|
619 |
have "dist x a = u * dist a b"
|
|
620 |
apply (simp add: dist_norm algebra_simps x)
|
|
621 |
by (metis abs_of_nonneg less_eq_real_def norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib \<open>0 < u\<close>)
|
|
622 |
also have *: "... < dist a b"
|
|
623 |
by (metis (no_types) assms dist_eq_0_iff dist_not_less_zero in_segment(2) linorder_neqE_linordered_idom mult.left_neutral real_mult_less_iff1 \<open>u < 1\<close>)
|
|
624 |
finally show "dist x a < dist a b" .
|
|
625 |
have ab_ne0: "dist a b \<noteq> 0"
|
|
626 |
using * by fastforce
|
|
627 |
have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)"
|
|
628 |
by (simp add: dist_norm algebra_simps x)
|
|
629 |
also have "... = (1-u) * dist a b"
|
|
630 |
proof -
|
|
631 |
have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)"
|
|
632 |
using \<open>u < 1\<close> by force
|
|
633 |
then show ?thesis
|
|
634 |
by (simp add: dist_norm real_vector.scale_right_diff_distrib)
|
|
635 |
qed
|
|
636 |
also have "... < dist a b"
|
|
637 |
using ab_ne0 \<open>0 < u\<close> by simp
|
|
638 |
finally show "dist x b < dist a b" .
|
|
639 |
qed
|
|
640 |
|
|
641 |
lemma dist_decreases_open_segment_0:
|
|
642 |
fixes x :: "'a :: euclidean_space"
|
|
643 |
assumes "x \<in> open_segment 0 b"
|
|
644 |
shows "dist c x < dist c 0 \<or> dist c x < dist c b"
|
|
645 |
proof (rule ccontr, clarsimp simp: not_less)
|
|
646 |
obtain u where u: "0 \<noteq> b" "0 < u" "u < 1" and x: "x = u *\<^sub>R b"
|
|
647 |
using assms by (auto simp: in_segment)
|
|
648 |
have xb: "x \<bullet> b < b \<bullet> b"
|
|
649 |
using u x by auto
|
|
650 |
assume "norm c \<le> dist c x"
|
|
651 |
then have "c \<bullet> c \<le> (c - x) \<bullet> (c - x)"
|
|
652 |
by (simp add: dist_norm norm_le)
|
|
653 |
moreover have "0 < x \<bullet> b"
|
|
654 |
using u x by auto
|
|
655 |
ultimately have less: "c \<bullet> b < x \<bullet> b"
|
|
656 |
by (simp add: x algebra_simps inner_commute u)
|
|
657 |
assume "dist c b \<le> dist c x"
|
|
658 |
then have "(c - b) \<bullet> (c - b) \<le> (c - x) \<bullet> (c - x)"
|
|
659 |
by (simp add: dist_norm norm_le)
|
|
660 |
then have "(b \<bullet> b) * (1 - u*u) \<le> 2 * (b \<bullet> c) * (1-u)"
|
|
661 |
by (simp add: x algebra_simps inner_commute)
|
|
662 |
then have "(1+u) * (b \<bullet> b) * (1-u) \<le> 2 * (b \<bullet> c) * (1-u)"
|
|
663 |
by (simp add: algebra_simps)
|
|
664 |
then have "(1+u) * (b \<bullet> b) \<le> 2 * (b \<bullet> c)"
|
|
665 |
using \<open>u < 1\<close> by auto
|
|
666 |
with xb have "c \<bullet> b \<ge> x \<bullet> b"
|
|
667 |
by (auto simp: x algebra_simps inner_commute)
|
|
668 |
with less show False by auto
|
|
669 |
qed
|
|
670 |
|
|
671 |
proposition dist_decreases_open_segment:
|
|
672 |
fixes a :: "'a :: euclidean_space"
|
|
673 |
assumes "x \<in> open_segment a b"
|
|
674 |
shows "dist c x < dist c a \<or> dist c x < dist c b"
|
|
675 |
proof -
|
|
676 |
have *: "x - a \<in> open_segment 0 (b - a)" using assms
|
|
677 |
by (metis diff_self open_segment_translation_eq uminus_add_conv_diff)
|
|
678 |
show ?thesis
|
|
679 |
using dist_decreases_open_segment_0 [OF *, of "c-a"] assms
|
|
680 |
by (simp add: dist_norm)
|
|
681 |
qed
|
|
682 |
|
|
683 |
corollary open_segment_furthest_le:
|
|
684 |
fixes a b x y :: "'a::euclidean_space"
|
|
685 |
assumes "x \<in> open_segment a b"
|
|
686 |
shows "norm (y - x) < norm (y - a) \<or> norm (y - x) < norm (y - b)"
|
|
687 |
by (metis assms dist_decreases_open_segment dist_norm)
|
|
688 |
|
|
689 |
corollary dist_decreases_closed_segment:
|
|
690 |
fixes a :: "'a :: euclidean_space"
|
|
691 |
assumes "x \<in> closed_segment a b"
|
|
692 |
shows "dist c x \<le> dist c a \<or> dist c x \<le> dist c b"
|
|
693 |
apply (cases "x \<in> open_segment a b")
|
|
694 |
using dist_decreases_open_segment less_eq_real_def apply blast
|
|
695 |
by (metis DiffI assms empty_iff insertE open_segment_def order_refl)
|
|
696 |
|
|
697 |
corollary segment_furthest_le:
|
|
698 |
fixes a b x y :: "'a::euclidean_space"
|
|
699 |
assumes "x \<in> closed_segment a b"
|
|
700 |
shows "norm (y - x) \<le> norm (y - a) \<or> norm (y - x) \<le> norm (y - b)"
|
|
701 |
by (metis assms dist_decreases_closed_segment dist_norm)
|
|
702 |
|
|
703 |
lemma convex_intermediate_ball:
|
|
704 |
fixes a :: "'a :: euclidean_space"
|
|
705 |
shows "\<lbrakk>ball a r \<subseteq> T; T \<subseteq> cball a r\<rbrakk> \<Longrightarrow> convex T"
|
|
706 |
apply (simp add: convex_contains_open_segment, clarify)
|
|
707 |
by (metis (no_types, hide_lams) less_le_trans mem_ball mem_cball subsetCE dist_decreases_open_segment)
|
|
708 |
|
|
709 |
lemma csegment_midpoint_subset: "closed_segment (midpoint a b) b \<subseteq> closed_segment a b"
|
|
710 |
apply (clarsimp simp: midpoint_def in_segment)
|
|
711 |
apply (rule_tac x="(1 + u) / 2" in exI)
|
|
712 |
apply (auto simp: algebra_simps add_divide_distrib diff_divide_distrib)
|
|
713 |
by (metis field_sum_of_halves scaleR_left.add)
|
|
714 |
|
|
715 |
lemma notin_segment_midpoint:
|
|
716 |
fixes a :: "'a :: euclidean_space"
|
|
717 |
shows "a \<noteq> b \<Longrightarrow> a \<notin> closed_segment (midpoint a b) b"
|
|
718 |
by (auto simp: dist_midpoint dest!: dist_in_closed_segment)
|
|
719 |
|
|
720 |
subsubsection\<open>More lemmas, especially for working with the underlying formula\<close>
|
|
721 |
|
|
722 |
lemma segment_eq_compose:
|
|
723 |
fixes a :: "'a :: real_vector"
|
|
724 |
shows "(\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) = (\<lambda>x. a + x) o (\<lambda>u. u *\<^sub>R (b - a))"
|
|
725 |
by (simp add: o_def algebra_simps)
|
|
726 |
|
|
727 |
lemma segment_degen_1:
|
|
728 |
fixes a :: "'a :: real_vector"
|
|
729 |
shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = b \<longleftrightarrow> a=b \<or> u=1"
|
|
730 |
proof -
|
|
731 |
{ assume "(1 - u) *\<^sub>R a + u *\<^sub>R b = b"
|
|
732 |
then have "(1 - u) *\<^sub>R a = (1 - u) *\<^sub>R b"
|
|
733 |
by (simp add: algebra_simps)
|
|
734 |
then have "a=b \<or> u=1"
|
|
735 |
by simp
|
|
736 |
} then show ?thesis
|
|
737 |
by (auto simp: algebra_simps)
|
|
738 |
qed
|
|
739 |
|
|
740 |
lemma segment_degen_0:
|
|
741 |
fixes a :: "'a :: real_vector"
|
|
742 |
shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = a \<longleftrightarrow> a=b \<or> u=0"
|
|
743 |
using segment_degen_1 [of "1-u" b a]
|
|
744 |
by (auto simp: algebra_simps)
|
|
745 |
|
|
746 |
lemma add_scaleR_degen:
|
|
747 |
fixes a b ::"'a::real_vector"
|
|
748 |
assumes "(u *\<^sub>R b + v *\<^sub>R a) = (u *\<^sub>R a + v *\<^sub>R b)" "u \<noteq> v"
|
|
749 |
shows "a=b"
|
|
750 |
by (metis (no_types, hide_lams) add.commute add_diff_eq diff_add_cancel real_vector.scale_cancel_left real_vector.scale_left_diff_distrib assms)
|
|
751 |
|
|
752 |
lemma closed_segment_image_interval:
|
|
753 |
"closed_segment a b = (\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0..1}"
|
|
754 |
by (auto simp: set_eq_iff image_iff closed_segment_def)
|
|
755 |
|
|
756 |
lemma open_segment_image_interval:
|
|
757 |
"open_segment a b = (if a=b then {} else (\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1})"
|
|
758 |
by (auto simp: open_segment_def closed_segment_def segment_degen_0 segment_degen_1)
|
|
759 |
|
|
760 |
lemmas segment_image_interval = closed_segment_image_interval open_segment_image_interval
|
|
761 |
|
|
762 |
lemma open_segment_bound1:
|
|
763 |
assumes "x \<in> open_segment a b"
|
|
764 |
shows "norm (x - a) < norm (b - a)"
|
|
765 |
proof -
|
|
766 |
obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 < u" "u < 1" "a \<noteq> b"
|
|
767 |
using assms by (auto simp add: open_segment_image_interval split: if_split_asm)
|
|
768 |
then show "norm (x - a) < norm (b - a)"
|
|
769 |
apply clarify
|
|
770 |
apply (auto simp: algebra_simps)
|
|
771 |
apply (simp add: scaleR_diff_right [symmetric])
|
|
772 |
done
|
|
773 |
qed
|
|
774 |
|
|
775 |
lemma compact_segment [simp]:
|
|
776 |
fixes a :: "'a::real_normed_vector"
|
|
777 |
shows "compact (closed_segment a b)"
|
|
778 |
by (auto simp: segment_image_interval intro!: compact_continuous_image continuous_intros)
|
|
779 |
|
|
780 |
lemma closed_segment [simp]:
|
|
781 |
fixes a :: "'a::real_normed_vector"
|
|
782 |
shows "closed (closed_segment a b)"
|
|
783 |
by (simp add: compact_imp_closed)
|
|
784 |
|
|
785 |
lemma closure_closed_segment [simp]:
|
|
786 |
fixes a :: "'a::real_normed_vector"
|
|
787 |
shows "closure(closed_segment a b) = closed_segment a b"
|
|
788 |
by simp
|
|
789 |
|
|
790 |
lemma open_segment_bound:
|
|
791 |
assumes "x \<in> open_segment a b"
|
|
792 |
shows "norm (x - a) < norm (b - a)" "norm (x - b) < norm (b - a)"
|
|
793 |
apply (simp add: assms open_segment_bound1)
|
|
794 |
by (metis assms norm_minus_commute open_segment_bound1 open_segment_commute)
|
|
795 |
|
|
796 |
lemma closure_open_segment [simp]:
|
|
797 |
"closure (open_segment a b) = (if a = b then {} else closed_segment a b)"
|
|
798 |
for a :: "'a::euclidean_space"
|
|
799 |
proof (cases "a = b")
|
|
800 |
case True
|
|
801 |
then show ?thesis
|
|
802 |
by simp
|
|
803 |
next
|
|
804 |
case False
|
|
805 |
have "closure ((\<lambda>u. u *\<^sub>R (b - a)) ` {0<..<1}) = (\<lambda>u. u *\<^sub>R (b - a)) ` closure {0<..<1}"
|
|
806 |
apply (rule closure_injective_linear_image [symmetric])
|
|
807 |
apply (use False in \<open>auto intro!: injI\<close>)
|
|
808 |
done
|
|
809 |
then have "closure
|
|
810 |
((\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1}) =
|
|
811 |
(\<lambda>x. (1 - x) *\<^sub>R a + x *\<^sub>R b) ` closure {0<..<1}"
|
|
812 |
using closure_translation [of a "((\<lambda>x. x *\<^sub>R b - x *\<^sub>R a) ` {0<..<1})"]
|
|
813 |
by (simp add: segment_eq_compose field_simps scaleR_diff_left scaleR_diff_right image_image)
|
|
814 |
then show ?thesis
|
|
815 |
by (simp add: segment_image_interval closure_greaterThanLessThan [symmetric] del: closure_greaterThanLessThan)
|
|
816 |
qed
|
|
817 |
|
|
818 |
lemma closed_open_segment_iff [simp]:
|
|
819 |
fixes a :: "'a::euclidean_space" shows "closed(open_segment a b) \<longleftrightarrow> a = b"
|
|
820 |
by (metis open_segment_def DiffE closure_eq closure_open_segment ends_in_segment(1) insert_iff segment_image_interval(2))
|
|
821 |
|
|
822 |
lemma compact_open_segment_iff [simp]:
|
|
823 |
fixes a :: "'a::euclidean_space" shows "compact(open_segment a b) \<longleftrightarrow> a = b"
|
|
824 |
by (simp add: bounded_open_segment compact_eq_bounded_closed)
|
|
825 |
|
|
826 |
lemma convex_closed_segment [iff]: "convex (closed_segment a b)"
|
|
827 |
unfolding segment_convex_hull by(rule convex_convex_hull)
|
|
828 |
|
|
829 |
lemma convex_open_segment [iff]: "convex (open_segment a b)"
|
|
830 |
proof -
|
|
831 |
have "convex ((\<lambda>u. u *\<^sub>R (b - a)) ` {0<..<1})"
|
|
832 |
by (rule convex_linear_image) auto
|
|
833 |
then have "convex ((+) a ` (\<lambda>u. u *\<^sub>R (b - a)) ` {0<..<1})"
|
|
834 |
by (rule convex_translation)
|
|
835 |
then show ?thesis
|
|
836 |
by (simp add: image_image open_segment_image_interval segment_eq_compose field_simps scaleR_diff_left scaleR_diff_right)
|
|
837 |
qed
|
|
838 |
|
|
839 |
lemmas convex_segment = convex_closed_segment convex_open_segment
|
|
840 |
|
|
841 |
lemma connected_segment [iff]:
|
|
842 |
fixes x :: "'a :: real_normed_vector"
|
|
843 |
shows "connected (closed_segment x y)"
|
|
844 |
by (simp add: convex_connected)
|
|
845 |
|
|
846 |
lemma is_interval_closed_segment_1[intro, simp]: "is_interval (closed_segment a b)" for a b::real
|
|
847 |
unfolding closed_segment_eq_real_ivl
|
|
848 |
by (auto simp: is_interval_def)
|
|
849 |
|
|
850 |
lemma IVT'_closed_segment_real:
|
|
851 |
fixes f :: "real \<Rightarrow> real"
|
|
852 |
assumes "y \<in> closed_segment (f a) (f b)"
|
|
853 |
assumes "continuous_on (closed_segment a b) f"
|
|
854 |
shows "\<exists>x \<in> closed_segment a b. f x = y"
|
|
855 |
using IVT'[of f a y b]
|
|
856 |
IVT'[of "-f" a "-y" b]
|
|
857 |
IVT'[of f b y a]
|
|
858 |
IVT'[of "-f" b "-y" a] assms
|
|
859 |
by (cases "a \<le> b"; cases "f b \<ge> f a") (auto simp: closed_segment_eq_real_ivl continuous_on_minus)
|
|
860 |
|
|
861 |
subsection \<open>Betweenness\<close>
|
|
862 |
|
|
863 |
definition\<^marker>\<open>tag important\<close> "between = (\<lambda>(a,b) x. x \<in> closed_segment a b)"
|
|
864 |
|
|
865 |
lemma betweenI:
|
|
866 |
assumes "0 \<le> u" "u \<le> 1" "x = (1 - u) *\<^sub>R a + u *\<^sub>R b"
|
|
867 |
shows "between (a, b) x"
|
|
868 |
using assms unfolding between_def closed_segment_def by auto
|
|
869 |
|
|
870 |
lemma betweenE:
|
|
871 |
assumes "between (a, b) x"
|
|
872 |
obtains u where "0 \<le> u" "u \<le> 1" "x = (1 - u) *\<^sub>R a + u *\<^sub>R b"
|
|
873 |
using assms unfolding between_def closed_segment_def by auto
|
|
874 |
|
|
875 |
lemma between_implies_scaled_diff:
|
|
876 |
assumes "between (S, T) X" "between (S, T) Y" "S \<noteq> Y"
|
|
877 |
obtains c where "(X - Y) = c *\<^sub>R (S - Y)"
|
|
878 |
proof -
|
|
879 |
from \<open>between (S, T) X\<close> obtain u\<^sub>X where X: "X = u\<^sub>X *\<^sub>R S + (1 - u\<^sub>X) *\<^sub>R T"
|
|
880 |
by (metis add.commute betweenE eq_diff_eq)
|
|
881 |
from \<open>between (S, T) Y\<close> obtain u\<^sub>Y where Y: "Y = u\<^sub>Y *\<^sub>R S + (1 - u\<^sub>Y) *\<^sub>R T"
|
|
882 |
by (metis add.commute betweenE eq_diff_eq)
|
|
883 |
have "X - Y = (u\<^sub>X - u\<^sub>Y) *\<^sub>R (S - T)"
|
|
884 |
proof -
|
|
885 |
from X Y have "X - Y = u\<^sub>X *\<^sub>R S - u\<^sub>Y *\<^sub>R S + ((1 - u\<^sub>X) *\<^sub>R T - (1 - u\<^sub>Y) *\<^sub>R T)" by simp
|
|
886 |
also have "\<dots> = (u\<^sub>X - u\<^sub>Y) *\<^sub>R S - (u\<^sub>X - u\<^sub>Y) *\<^sub>R T" by (simp add: scaleR_left.diff)
|
|
887 |
finally show ?thesis by (simp add: real_vector.scale_right_diff_distrib)
|
|
888 |
qed
|
|
889 |
moreover from Y have "S - Y = (1 - u\<^sub>Y) *\<^sub>R (S - T)"
|
|
890 |
by (simp add: real_vector.scale_left_diff_distrib real_vector.scale_right_diff_distrib)
|
|
891 |
moreover note \<open>S \<noteq> Y\<close>
|
|
892 |
ultimately have "(X - Y) = ((u\<^sub>X - u\<^sub>Y) / (1 - u\<^sub>Y)) *\<^sub>R (S - Y)" by auto
|
|
893 |
from this that show thesis by blast
|
|
894 |
qed
|
|
895 |
|
|
896 |
lemma between_mem_segment: "between (a,b) x \<longleftrightarrow> x \<in> closed_segment a b"
|
|
897 |
unfolding between_def by auto
|
|
898 |
|
|
899 |
lemma between: "between (a, b) (x::'a::euclidean_space) \<longleftrightarrow> dist a b = (dist a x) + (dist x b)"
|
|
900 |
proof (cases "a = b")
|
|
901 |
case True
|
|
902 |
then show ?thesis
|
|
903 |
by (auto simp add: between_def dist_commute)
|
|
904 |
next
|
|
905 |
case False
|
|
906 |
then have Fal: "norm (a - b) \<noteq> 0" and Fal2: "norm (a - b) > 0"
|
|
907 |
by auto
|
|
908 |
have *: "\<And>u. a - ((1 - u) *\<^sub>R a + u *\<^sub>R b) = u *\<^sub>R (a - b)"
|
|
909 |
by (auto simp add: algebra_simps)
|
|
910 |
have "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)" if "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1" for u
|
|
911 |
proof -
|
|
912 |
have *: "a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)"
|
|
913 |
unfolding that(1) by (auto simp add:algebra_simps)
|
|
914 |
show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)"
|
|
915 |
unfolding norm_minus_commute[of x a] * using \<open>0 \<le> u\<close> \<open>u \<le> 1\<close>
|
|
916 |
by simp
|
|
917 |
qed
|
|
918 |
moreover have "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1" if "dist a b = dist a x + dist x b"
|
|
919 |
proof -
|
|
920 |
let ?\<beta> = "norm (a - x) / norm (a - b)"
|
|
921 |
show "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1"
|
|
922 |
proof (intro exI conjI)
|
|
923 |
show "?\<beta> \<le> 1"
|
|
924 |
using Fal2 unfolding that[unfolded dist_norm] norm_ge_zero by auto
|
|
925 |
show "x = (1 - ?\<beta>) *\<^sub>R a + (?\<beta>) *\<^sub>R b"
|
|
926 |
proof (subst euclidean_eq_iff; intro ballI)
|
|
927 |
fix i :: 'a
|
|
928 |
assume i: "i \<in> Basis"
|
|
929 |
have "((1 - ?\<beta>) *\<^sub>R a + (?\<beta>) *\<^sub>R b) \<bullet> i
|
|
930 |
= ((norm (a - b) - norm (a - x)) * (a \<bullet> i) + norm (a - x) * (b \<bullet> i)) / norm (a - b)"
|
|
931 |
using Fal by (auto simp add: field_simps inner_simps)
|
|
932 |
also have "\<dots> = x\<bullet>i"
|
|
933 |
apply (rule divide_eq_imp[OF Fal])
|
|
934 |
unfolding that[unfolded dist_norm]
|
|
935 |
using that[unfolded dist_triangle_eq] i
|
|
936 |
apply (subst (asm) euclidean_eq_iff)
|
|
937 |
apply (auto simp add: field_simps inner_simps)
|
|
938 |
done
|
|
939 |
finally show "x \<bullet> i = ((1 - ?\<beta>) *\<^sub>R a + (?\<beta>) *\<^sub>R b) \<bullet> i"
|
|
940 |
by auto
|
|
941 |
qed
|
|
942 |
qed (use Fal2 in auto)
|
|
943 |
qed
|
|
944 |
ultimately show ?thesis
|
|
945 |
by (force simp add: between_def closed_segment_def dist_triangle_eq)
|
|
946 |
qed
|
|
947 |
|
|
948 |
lemma between_midpoint:
|
|
949 |
fixes a :: "'a::euclidean_space"
|
|
950 |
shows "between (a,b) (midpoint a b)" (is ?t1)
|
|
951 |
and "between (b,a) (midpoint a b)" (is ?t2)
|
|
952 |
proof -
|
|
953 |
have *: "\<And>x y z. x = (1/2::real) *\<^sub>R z \<Longrightarrow> y = (1/2) *\<^sub>R z \<Longrightarrow> norm z = norm x + norm y"
|
|
954 |
by auto
|
|
955 |
show ?t1 ?t2
|
|
956 |
unfolding between midpoint_def dist_norm
|
|
957 |
by (auto simp add: field_simps inner_simps euclidean_eq_iff[where 'a='a] intro!: *)
|
|
958 |
qed
|
|
959 |
|
|
960 |
lemma between_mem_convex_hull:
|
|
961 |
"between (a,b) x \<longleftrightarrow> x \<in> convex hull {a,b}"
|
|
962 |
unfolding between_mem_segment segment_convex_hull ..
|
|
963 |
|
|
964 |
lemma between_triv_iff [simp]: "between (a,a) b \<longleftrightarrow> a=b"
|
|
965 |
by (auto simp: between_def)
|
|
966 |
|
|
967 |
lemma between_triv1 [simp]: "between (a,b) a"
|
|
968 |
by (auto simp: between_def)
|
|
969 |
|
|
970 |
lemma between_triv2 [simp]: "between (a,b) b"
|
|
971 |
by (auto simp: between_def)
|
|
972 |
|
|
973 |
lemma between_commute:
|
|
974 |
"between (a,b) = between (b,a)"
|
|
975 |
by (auto simp: between_def closed_segment_commute)
|
|
976 |
|
|
977 |
lemma between_antisym:
|
|
978 |
fixes a :: "'a :: euclidean_space"
|
|
979 |
shows "\<lbrakk>between (b,c) a; between (a,c) b\<rbrakk> \<Longrightarrow> a = b"
|
|
980 |
by (auto simp: between dist_commute)
|
|
981 |
|
|
982 |
lemma between_trans:
|
|
983 |
fixes a :: "'a :: euclidean_space"
|
|
984 |
shows "\<lbrakk>between (b,c) a; between (a,c) d\<rbrakk> \<Longrightarrow> between (b,c) d"
|
|
985 |
using dist_triangle2 [of b c d] dist_triangle3 [of b d a]
|
|
986 |
by (auto simp: between dist_commute)
|
|
987 |
|
|
988 |
lemma between_norm:
|
|
989 |
fixes a :: "'a :: euclidean_space"
|
|
990 |
shows "between (a,b) x \<longleftrightarrow> norm(x - a) *\<^sub>R (b - x) = norm(b - x) *\<^sub>R (x - a)"
|
|
991 |
by (auto simp: between dist_triangle_eq norm_minus_commute algebra_simps)
|
|
992 |
|
|
993 |
lemma between_swap:
|
|
994 |
fixes A B X Y :: "'a::euclidean_space"
|
|
995 |
assumes "between (A, B) X"
|
|
996 |
assumes "between (A, B) Y"
|
|
997 |
shows "between (X, B) Y \<longleftrightarrow> between (A, Y) X"
|
|
998 |
using assms by (auto simp add: between)
|
|
999 |
|
|
1000 |
lemma between_translation [simp]: "between (a + y,a + z) (a + x) \<longleftrightarrow> between (y,z) x"
|
|
1001 |
by (auto simp: between_def)
|
|
1002 |
|
|
1003 |
lemma between_trans_2:
|
|
1004 |
fixes a :: "'a :: euclidean_space"
|
|
1005 |
shows "\<lbrakk>between (b,c) a; between (a,b) d\<rbrakk> \<Longrightarrow> between (c,d) a"
|
|
1006 |
by (metis between_commute between_swap between_trans)
|
|
1007 |
|
|
1008 |
lemma between_scaleR_lift [simp]:
|
|
1009 |
fixes v :: "'a::euclidean_space"
|
|
1010 |
shows "between (a *\<^sub>R v, b *\<^sub>R v) (c *\<^sub>R v) \<longleftrightarrow> v = 0 \<or> between (a, b) c"
|
|
1011 |
by (simp add: between dist_norm scaleR_left_diff_distrib [symmetric] distrib_right [symmetric])
|
|
1012 |
|
|
1013 |
lemma between_1:
|
|
1014 |
fixes x::real
|
|
1015 |
shows "between (a,b) x \<longleftrightarrow> (a \<le> x \<and> x \<le> b) \<or> (b \<le> x \<and> x \<le> a)"
|
|
1016 |
by (auto simp: between_mem_segment closed_segment_eq_real_ivl)
|
|
1017 |
|
|
1018 |
end |