author | nipkow |
Tue, 17 Jun 2025 14:11:40 +0200 | |
changeset 82733 | 8b537e1af2ec |
parent 79582 | 7822b55b26ce |
permissions | -rw-r--r-- |
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1 |
(* Title: HOL/Analysis/Convex_Euclidean_Space.thy |
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
2 |
Author: L C Paulson, University of Cambridge |
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
3 |
Author: Robert Himmelmann, TU Muenchen |
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
4 |
Author: Bogdan Grechuk, University of Edinburgh |
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
5 |
Author: Armin Heller, TU Muenchen |
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
6 |
Author: Johannes Hoelzl, TU Muenchen |
33175 | 7 |
*) |
8 |
||
69619
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents:
69618
diff
changeset
|
9 |
section \<open>Convex Sets and Functions on (Normed) Euclidean Spaces\<close> |
33175 | 10 |
|
11 |
theory Convex_Euclidean_Space |
|
44132 | 12 |
imports |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
13 |
Convex Topology_Euclidean_Space Line_Segment |
33175 | 14 |
begin |
15 |
||
70136 | 16 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Topological Properties of Convex Sets and Functions\<close> |
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
17 |
|
40377 | 18 |
lemma aff_dim_cball: |
53347 | 19 |
fixes a :: "'n::euclidean_space" |
20 |
assumes "e > 0" |
|
21 |
shows "aff_dim (cball a e) = int (DIM('n))" |
|
22 |
proof - |
|
23 |
have "(\<lambda>x. a + x) ` (cball 0 e) \<subseteq> cball a e" |
|
24 |
unfolding cball_def dist_norm by auto |
|
25 |
then have "aff_dim (cball (0 :: 'n::euclidean_space) e) \<le> aff_dim (cball a e)" |
|
26 |
using aff_dim_translation_eq[of a "cball 0 e"] |
|
67399 | 27 |
aff_dim_subset[of "(+) a ` cball 0 e" "cball a e"] |
53347 | 28 |
by auto |
29 |
moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))" |
|
30 |
using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"] |
|
31 |
centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms |
|
32 |
by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"]) |
|
33 |
ultimately show ?thesis |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
34 |
using aff_dim_le_DIM[of "cball a e"] by auto |
40377 | 35 |
qed |
36 |
||
37 |
lemma aff_dim_open: |
|
53347 | 38 |
fixes S :: "'n::euclidean_space set" |
39 |
assumes "open S" |
|
40 |
and "S \<noteq> {}" |
|
41 |
shows "aff_dim S = int (DIM('n))" |
|
42 |
proof - |
|
43 |
obtain x where "x \<in> S" |
|
44 |
using assms by auto |
|
45 |
then obtain e where e: "e > 0" "cball x e \<subseteq> S" |
|
46 |
using open_contains_cball[of S] assms by auto |
|
47 |
then have "aff_dim (cball x e) \<le> aff_dim S" |
|
48 |
using aff_dim_subset by auto |
|
49 |
with e show ?thesis |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
50 |
using aff_dim_cball[of e x] aff_dim_le_DIM[of S] by auto |
40377 | 51 |
qed |
52 |
||
53 |
lemma low_dim_interior: |
|
53347 | 54 |
fixes S :: "'n::euclidean_space set" |
55 |
assumes "\<not> aff_dim S = int (DIM('n))" |
|
56 |
shows "interior S = {}" |
|
57 |
proof - |
|
58 |
have "aff_dim(interior S) \<le> aff_dim S" |
|
59 |
using interior_subset aff_dim_subset[of "interior S" S] by auto |
|
60 |
then show ?thesis |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
61 |
using aff_dim_open[of "interior S"] aff_dim_le_DIM[of S] assms by auto |
40377 | 62 |
qed |
63 |
||
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
64 |
corollary empty_interior_lowdim: |
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
65 |
fixes S :: "'n::euclidean_space set" |
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
66 |
shows "dim S < DIM ('n) \<Longrightarrow> interior S = {}" |
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
67 |
by (metis low_dim_interior affine_hull_UNIV dim_affine_hull less_not_refl dim_UNIV) |
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
68 |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
69 |
corollary aff_dim_nonempty_interior: |
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
70 |
fixes S :: "'a::euclidean_space set" |
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
71 |
shows "interior S \<noteq> {} \<Longrightarrow> aff_dim S = DIM('a)" |
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
72 |
by (metis low_dim_interior) |
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
73 |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
74 |
|
60420 | 75 |
subsection \<open>Relative interior of a set\<close> |
40377 | 76 |
|
70136 | 77 |
definition\<^marker>\<open>tag important\<close> "rel_interior S = |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69710
diff
changeset
|
78 |
{x. \<exists>T. openin (top_of_set (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}" |
53347 | 79 |
|
64287 | 80 |
lemma rel_interior_mono: |
81 |
"\<lbrakk>S \<subseteq> T; affine hull S = affine hull T\<rbrakk> |
|
82 |
\<Longrightarrow> (rel_interior S) \<subseteq> (rel_interior T)" |
|
83 |
by (auto simp: rel_interior_def) |
|
84 |
||
85 |
lemma rel_interior_maximal: |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69710
diff
changeset
|
86 |
"\<lbrakk>T \<subseteq> S; openin(top_of_set (affine hull S)) T\<rbrakk> \<Longrightarrow> T \<subseteq> (rel_interior S)" |
64287 | 87 |
by (auto simp: rel_interior_def) |
88 |
||
72356 | 89 |
lemma rel_interior: "rel_interior S = {x \<in> S. \<exists>T. open T \<and> x \<in> T \<and> T \<inter> affine hull S \<subseteq> S}" |
90 |
(is "?lhs = ?rhs") |
|
91 |
proof |
|
92 |
show "?lhs \<subseteq> ?rhs" |
|
93 |
by (force simp add: rel_interior_def openin_open) |
|
94 |
{ fix x T |
|
95 |
assume *: "x \<in> S" "open T" "x \<in> T" "T \<inter> affine hull S \<subseteq> S" |
|
96 |
then have **: "x \<in> T \<inter> affine hull S" |
|
97 |
using hull_inc by auto |
|
98 |
with * have "\<exists>Tb. (\<exists>Ta. open Ta \<and> Tb = affine hull S \<inter> Ta) \<and> x \<in> Tb \<and> Tb \<subseteq> S" |
|
99 |
by (rule_tac x = "T \<inter> (affine hull S)" in exI) auto |
|
100 |
} |
|
101 |
then show "?rhs \<subseteq> ?lhs" |
|
102 |
by (force simp add: rel_interior_def openin_open) |
|
53347 | 103 |
qed |
104 |
||
105 |
lemma mem_rel_interior: "x \<in> rel_interior S \<longleftrightarrow> (\<exists>T. open T \<and> x \<in> T \<inter> S \<and> T \<inter> affine hull S \<subseteq> S)" |
|
68031 | 106 |
by (auto simp: rel_interior) |
53347 | 107 |
|
108 |
lemma mem_rel_interior_ball: |
|
109 |
"x \<in> rel_interior S \<longleftrightarrow> x \<in> S \<and> (\<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S)" |
|
72356 | 110 |
(is "?lhs = ?rhs") |
111 |
proof |
|
112 |
assume ?rhs then show ?lhs |
|
113 |
by (simp add: rel_interior) (meson Elementary_Metric_Spaces.open_ball centre_in_ball) |
|
114 |
qed (force simp: rel_interior open_contains_ball) |
|
40377 | 115 |
|
49531 | 116 |
lemma rel_interior_ball: |
53347 | 117 |
"rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S}" |
118 |
using mem_rel_interior_ball [of _ S] by auto |
|
119 |
||
120 |
lemma mem_rel_interior_cball: |
|
121 |
"x \<in> rel_interior S \<longleftrightarrow> x \<in> S \<and> (\<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S)" |
|
72356 | 122 |
(is "?lhs = ?rhs") |
123 |
proof |
|
124 |
assume ?rhs then obtain e where "x \<in> S" "e > 0" "cball x e \<inter> affine hull S \<subseteq> S" |
|
125 |
by (auto simp: rel_interior) |
|
126 |
then have "ball x e \<inter> affine hull S \<subseteq> S" |
|
127 |
by auto |
|
128 |
then show ?lhs |
|
129 |
using \<open>0 < e\<close> \<open>x \<in> S\<close> rel_interior_ball by auto |
|
130 |
qed (force simp: rel_interior open_contains_cball) |
|
40377 | 131 |
|
53347 | 132 |
lemma rel_interior_cball: |
133 |
"rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S}" |
|
134 |
using mem_rel_interior_cball [of _ S] by auto |
|
40377 | 135 |
|
60303 | 136 |
lemma rel_interior_empty [simp]: "rel_interior {} = {}" |
68031 | 137 |
by (auto simp: rel_interior_def) |
40377 | 138 |
|
60303 | 139 |
lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}" |
53347 | 140 |
by (metis affine_hull_eq affine_sing) |
40377 | 141 |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
142 |
lemma rel_interior_sing [simp]: |
72356 | 143 |
fixes a :: "'n::euclidean_space" shows "rel_interior {a} = {a}" |
144 |
proof - |
|
145 |
have "\<exists>x::real. 0 < x" |
|
146 |
using zero_less_one by blast |
|
147 |
then show ?thesis |
|
148 |
by (auto simp: rel_interior_ball) |
|
149 |
qed |
|
40377 | 150 |
|
151 |
lemma subset_rel_interior: |
|
53347 | 152 |
fixes S T :: "'n::euclidean_space set" |
153 |
assumes "S \<subseteq> T" |
|
154 |
and "affine hull S = affine hull T" |
|
155 |
shows "rel_interior S \<subseteq> rel_interior T" |
|
68031 | 156 |
using assms by (auto simp: rel_interior_def) |
49531 | 157 |
|
53347 | 158 |
lemma rel_interior_subset: "rel_interior S \<subseteq> S" |
68031 | 159 |
by (auto simp: rel_interior_def) |
53347 | 160 |
|
161 |
lemma rel_interior_subset_closure: "rel_interior S \<subseteq> closure S" |
|
68031 | 162 |
using rel_interior_subset by (auto simp: closure_def) |
53347 | 163 |
|
164 |
lemma interior_subset_rel_interior: "interior S \<subseteq> rel_interior S" |
|
68031 | 165 |
by (auto simp: rel_interior interior_def) |
40377 | 166 |
|
167 |
lemma interior_rel_interior: |
|
53347 | 168 |
fixes S :: "'n::euclidean_space set" |
169 |
assumes "aff_dim S = int(DIM('n))" |
|
170 |
shows "rel_interior S = interior S" |
|
40377 | 171 |
proof - |
53347 | 172 |
have "affine hull S = UNIV" |
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
173 |
using assms affine_hull_UNIV[of S] by auto |
53347 | 174 |
then show ?thesis |
175 |
unfolding rel_interior interior_def by auto |
|
40377 | 176 |
qed |
177 |
||
60303 | 178 |
lemma rel_interior_interior: |
179 |
fixes S :: "'n::euclidean_space set" |
|
180 |
assumes "affine hull S = UNIV" |
|
181 |
shows "rel_interior S = interior S" |
|
182 |
using assms unfolding rel_interior interior_def by auto |
|
183 |
||
40377 | 184 |
lemma rel_interior_open: |
53347 | 185 |
fixes S :: "'n::euclidean_space set" |
186 |
assumes "open S" |
|
187 |
shows "rel_interior S = S" |
|
188 |
by (metis assms interior_eq interior_subset_rel_interior rel_interior_subset set_eq_subset) |
|
40377 | 189 |
|
190 |
lemma interior_rel_interior_gen: |
|
53347 | 191 |
fixes S :: "'n::euclidean_space set" |
192 |
shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})" |
|
193 |
by (metis interior_rel_interior low_dim_interior) |
|
40377 | 194 |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
195 |
lemma rel_interior_nonempty_interior: |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
196 |
fixes S :: "'n::euclidean_space set" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
197 |
shows "interior S \<noteq> {} \<Longrightarrow> rel_interior S = interior S" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
198 |
by (metis interior_rel_interior_gen) |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
199 |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
200 |
lemma affine_hull_nonempty_interior: |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
201 |
fixes S :: "'n::euclidean_space set" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
202 |
shows "interior S \<noteq> {} \<Longrightarrow> affine hull S = UNIV" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
203 |
by (metis affine_hull_UNIV interior_rel_interior_gen) |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
204 |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
205 |
lemma rel_interior_affine_hull [simp]: |
53347 | 206 |
fixes S :: "'n::euclidean_space set" |
207 |
shows "rel_interior (affine hull S) = affine hull S" |
|
208 |
proof - |
|
209 |
have *: "rel_interior (affine hull S) \<subseteq> affine hull S" |
|
210 |
using rel_interior_subset by auto |
|
211 |
{ |
|
212 |
fix x |
|
213 |
assume x: "x \<in> affine hull S" |
|
63040 | 214 |
define e :: real where "e = 1" |
53347 | 215 |
then have "e > 0" "ball x e \<inter> affine hull (affine hull S) \<subseteq> affine hull S" |
216 |
using hull_hull[of _ S] by auto |
|
217 |
then have "x \<in> rel_interior (affine hull S)" |
|
218 |
using x rel_interior_ball[of "affine hull S"] by auto |
|
219 |
} |
|
220 |
then show ?thesis using * by auto |
|
40377 | 221 |
qed |
222 |
||
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
223 |
lemma rel_interior_UNIV [simp]: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV" |
53347 | 224 |
by (metis open_UNIV rel_interior_open) |
40377 | 225 |
|
226 |
lemma rel_interior_convex_shrink: |
|
53347 | 227 |
fixes S :: "'a::euclidean_space set" |
228 |
assumes "convex S" |
|
229 |
and "c \<in> rel_interior S" |
|
230 |
and "x \<in> S" |
|
231 |
and "0 < e" |
|
232 |
and "e \<le> 1" |
|
233 |
shows "x - e *\<^sub>R (x - c) \<in> rel_interior S" |
|
234 |
proof - |
|
54465 | 235 |
obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S" |
53347 | 236 |
using assms(2) unfolding mem_rel_interior_ball by auto |
237 |
{ |
|
238 |
fix y |
|
239 |
assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \<in> affine hull S" |
|
240 |
have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" |
|
68031 | 241 |
using \<open>e > 0\<close> by (auto simp: scaleR_left_diff_distrib scaleR_right_diff_distrib) |
53347 | 242 |
have "x \<in> affine hull S" |
243 |
using assms hull_subset[of S] by auto |
|
49531 | 244 |
moreover have "1 / e + - ((1 - e) / e) = 1" |
60420 | 245 |
using \<open>e > 0\<close> left_diff_distrib[of "1" "(1-e)" "1/e"] by auto |
53347 | 246 |
ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \<in> affine hull S" |
247 |
using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"] |
|
248 |
by (simp add: algebra_simps) |
|
72356 | 249 |
have "c - ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = (1 / e) *\<^sub>R (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" |
60420 | 250 |
using \<open>e > 0\<close> |
72356 | 251 |
by (auto simp: euclidean_eq_iff[where 'a='a] field_simps inner_simps) |
252 |
then have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \<bar>1/e\<bar> * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" |
|
253 |
unfolding dist_norm norm_scaleR[symmetric] by auto |
|
61945 | 254 |
also have "\<dots> = \<bar>1/e\<bar> * norm (x - e *\<^sub>R (x - c) - y)" |
53347 | 255 |
by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) |
256 |
also have "\<dots> < d" |
|
60420 | 257 |
using as[unfolded dist_norm] and \<open>e > 0\<close> |
68031 | 258 |
by (auto simp:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute) |
72356 | 259 |
finally have "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \<in> S" |
260 |
using "**" d by auto |
|
261 |
then have "y \<in> S" |
|
262 |
using * convexD [OF \<open>convex S\<close>] assms(3-5) |
|
263 |
by (metis diff_add_cancel diff_ge_0_iff_ge le_add_same_cancel1 less_eq_real_def) |
|
53347 | 264 |
} |
265 |
then have "ball (x - e *\<^sub>R (x - c)) (e*d) \<inter> affine hull S \<subseteq> S" |
|
266 |
by auto |
|
267 |
moreover have "e * d > 0" |
|
60420 | 268 |
using \<open>e > 0\<close> \<open>d > 0\<close> by simp |
53347 | 269 |
moreover have c: "c \<in> S" |
270 |
using assms rel_interior_subset by auto |
|
271 |
moreover from c have "x - e *\<^sub>R (x - c) \<in> S" |
|
72356 | 272 |
using convexD_alt[of S x c e] assms |
273 |
by (metis diff_add_eq diff_diff_eq2 less_eq_real_def scaleR_diff_left scaleR_one scale_right_diff_distrib) |
|
53347 | 274 |
ultimately show ?thesis |
60420 | 275 |
using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] \<open>e > 0\<close> by auto |
40377 | 276 |
qed |
277 |
||
69710
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
278 |
lemma interior_real_atLeast [simp]: |
53347 | 279 |
fixes a :: real |
280 |
shows "interior {a..} = {a<..}" |
|
281 |
proof - |
|
282 |
{ |
|
283 |
fix y |
|
72356 | 284 |
have "ball y (y - a) \<subseteq> {a..}" |
285 |
by (auto simp: dist_norm) |
|
286 |
moreover assume "a < y" |
|
287 |
ultimately have "y \<in> interior {a..}" |
|
288 |
by (force simp add: mem_interior) |
|
53347 | 289 |
} |
290 |
moreover |
|
291 |
{ |
|
292 |
fix y |
|
293 |
assume "y \<in> interior {a..}" |
|
294 |
then obtain e where e: "e > 0" "cball y e \<subseteq> {a..}" |
|
295 |
using mem_interior_cball[of y "{a..}"] by auto |
|
296 |
moreover from e have "y - e \<in> cball y e" |
|
68031 | 297 |
by (auto simp: cball_def dist_norm) |
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
298 |
ultimately have "a \<le> y - e" by blast |
53347 | 299 |
then have "a < y" using e by auto |
300 |
} |
|
301 |
ultimately show ?thesis by auto |
|
40377 | 302 |
qed |
303 |
||
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
304 |
lemma continuous_ge_on_Ioo: |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
305 |
assumes "continuous_on {c..d} g" "\<And>x. x \<in> {c<..<d} \<Longrightarrow> g x \<ge> a" "c < d" "x \<in> {c..d}" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
306 |
shows "g (x::real) \<ge> (a::real)" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
307 |
proof- |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
308 |
from assms(3) have "{c..d} = closure {c<..<d}" by (rule closure_greaterThanLessThan[symmetric]) |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
309 |
also from assms(2) have "{c<..<d} \<subseteq> (g -` {a..} \<inter> {c..d})" by auto |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
310 |
hence "closure {c<..<d} \<subseteq> closure (g -` {a..} \<inter> {c..d})" by (rule closure_mono) |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
311 |
also from assms(1) have "closed (g -` {a..} \<inter> {c..d})" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
312 |
by (auto simp: continuous_on_closed_vimage) |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
313 |
hence "closure (g -` {a..} \<inter> {c..d}) = g -` {a..} \<inter> {c..d}" by simp |
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61952
diff
changeset
|
314 |
finally show ?thesis using \<open>x \<in> {c..d}\<close> by auto |
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61952
diff
changeset
|
315 |
qed |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
316 |
|
69710
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
317 |
lemma interior_real_atMost [simp]: |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
318 |
fixes a :: real |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
319 |
shows "interior {..a} = {..<a}" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
320 |
proof - |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
321 |
{ |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
322 |
fix y |
72356 | 323 |
have "ball y (a - y) \<subseteq> {..a}" |
324 |
by (auto simp: dist_norm) |
|
325 |
moreover assume "a > y" |
|
326 |
ultimately have "y \<in> interior {..a}" |
|
327 |
by (force simp add: mem_interior) |
|
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
328 |
} |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
329 |
moreover |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
330 |
{ |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
331 |
fix y |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
332 |
assume "y \<in> interior {..a}" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
333 |
then obtain e where e: "e > 0" "cball y e \<subseteq> {..a}" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
334 |
using mem_interior_cball[of y "{..a}"] by auto |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
335 |
moreover from e have "y + e \<in> cball y e" |
68031 | 336 |
by (auto simp: cball_def dist_norm) |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
337 |
ultimately have "a \<ge> y + e" by auto |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
338 |
then have "a > y" using e by auto |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
339 |
} |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
340 |
ultimately show ?thesis by auto |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
341 |
qed |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
342 |
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
343 |
lemma interior_atLeastAtMost_real [simp]: "interior {a..b} = {a<..<b :: real}" |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
344 |
proof- |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
345 |
have "{a..b} = {a..} \<inter> {..b}" by auto |
68031 | 346 |
also have "interior \<dots> = {a<..} \<inter> {..<b}" |
71633 | 347 |
by (simp) |
68031 | 348 |
also have "\<dots> = {a<..<b}" by auto |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
349 |
finally show ?thesis . |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
350 |
qed |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
351 |
|
66793
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
352 |
lemma interior_atLeastLessThan [simp]: |
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
353 |
fixes a::real shows "interior {a..<b} = {a<..<b}" |
69710
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
354 |
by (metis atLeastLessThan_def greaterThanLessThan_def interior_atLeastAtMost_real interior_Int interior_interior interior_real_atLeast) |
66793
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
355 |
|
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
356 |
lemma interior_lessThanAtMost [simp]: |
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
357 |
fixes a::real shows "interior {a<..b} = {a<..<b}" |
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
358 |
by (metis atLeastAtMost_def greaterThanAtMost_def interior_atLeastAtMost_real interior_Int |
69710
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
359 |
interior_interior interior_real_atLeast) |
66793
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
360 |
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
361 |
lemma interior_greaterThanLessThan_real [simp]: "interior {a<..<b} = {a<..<b :: real}" |
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
362 |
by (metis interior_atLeastAtMost_real interior_interior) |
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
363 |
|
69710
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
364 |
lemma frontier_real_atMost [simp]: |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
365 |
fixes a :: real |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
366 |
shows "frontier {..a} = {a}" |
72356 | 367 |
unfolding frontier_def by auto |
69710
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
368 |
|
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
369 |
lemma frontier_real_atLeast [simp]: "frontier {a..} = {a::real}" |
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
370 |
by (auto simp: frontier_def) |
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
371 |
|
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
372 |
lemma frontier_real_greaterThan [simp]: "frontier {a<..} = {a::real}" |
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
373 |
by (auto simp: interior_open frontier_def) |
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
374 |
|
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
375 |
lemma frontier_real_lessThan [simp]: "frontier {..<a} = {a::real}" |
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
376 |
by (auto simp: interior_open frontier_def) |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
377 |
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
378 |
lemma rel_interior_real_box [simp]: |
53347 | 379 |
fixes a b :: real |
380 |
assumes "a < b" |
|
56188 | 381 |
shows "rel_interior {a .. b} = {a <..< b}" |
53347 | 382 |
proof - |
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54465
diff
changeset
|
383 |
have "box a b \<noteq> {}" |
53347 | 384 |
using assms |
385 |
unfolding set_eq_iff |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
386 |
by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def) |
40377 | 387 |
then show ?thesis |
56188 | 388 |
using interior_rel_interior_gen[of "cbox a b", symmetric] |
72356 | 389 |
by (simp split: if_split_asm del: box_real add: box_real[symmetric]) |
40377 | 390 |
qed |
391 |
||
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
392 |
lemma rel_interior_real_semiline [simp]: |
53347 | 393 |
fixes a :: real |
394 |
shows "rel_interior {a..} = {a<..}" |
|
395 |
proof - |
|
396 |
have *: "{a<..} \<noteq> {}" |
|
397 |
unfolding set_eq_iff by (auto intro!: exI[of _ "a + 1"]) |
|
69710
61372780515b
some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents:
69619
diff
changeset
|
398 |
then show ?thesis using interior_real_atLeast interior_rel_interior_gen[of "{a..}"] |
62390 | 399 |
by (auto split: if_split_asm) |
40377 | 400 |
qed |
401 |
||
60420 | 402 |
subsubsection \<open>Relative open sets\<close> |
40377 | 403 |
|
70136 | 404 |
definition\<^marker>\<open>tag important\<close> "rel_open S \<longleftrightarrow> rel_interior S = S" |
53347 | 405 |
|
72356 | 406 |
lemma rel_open: "rel_open S \<longleftrightarrow> openin (top_of_set (affine hull S)) S" (is "?lhs = ?rhs") |
407 |
proof |
|
408 |
assume ?lhs |
|
409 |
then show ?rhs |
|
410 |
unfolding rel_open_def rel_interior_def |
|
411 |
using openin_subopen[of "top_of_set (affine hull S)" S] by auto |
|
412 |
qed (auto simp: rel_open_def rel_interior_def) |
|
53347 | 413 |
|
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69710
diff
changeset
|
414 |
lemma openin_rel_interior: "openin (top_of_set (affine hull S)) (rel_interior S)" |
72356 | 415 |
using openin_subopen by (fastforce simp add: rel_interior_def) |
40377 | 416 |
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
417 |
lemma openin_set_rel_interior: |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69710
diff
changeset
|
418 |
"openin (top_of_set S) (rel_interior S)" |
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
419 |
by (rule openin_subset_trans [OF openin_rel_interior rel_interior_subset hull_subset]) |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
420 |
|
49531 | 421 |
lemma affine_rel_open: |
53347 | 422 |
fixes S :: "'n::euclidean_space set" |
423 |
assumes "affine S" |
|
424 |
shows "rel_open S" |
|
425 |
unfolding rel_open_def |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
426 |
using assms rel_interior_affine_hull[of S] affine_hull_eq[of S] |
53347 | 427 |
by metis |
40377 | 428 |
|
49531 | 429 |
lemma affine_closed: |
53347 | 430 |
fixes S :: "'n::euclidean_space set" |
431 |
assumes "affine S" |
|
432 |
shows "closed S" |
|
433 |
proof - |
|
434 |
{ |
|
435 |
assume "S \<noteq> {}" |
|
436 |
then obtain L where L: "subspace L" "affine_parallel S L" |
|
437 |
using assms affine_parallel_subspace[of S] by auto |
|
67399 | 438 |
then obtain a where a: "S = ((+) a ` L)" |
78516
56a408fa2440
substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents:
77935
diff
changeset
|
439 |
using affine_parallel_def[of L S] affine_parallel_commute by auto |
53347 | 440 |
from L have "closed L" using closed_subspace by auto |
441 |
then have "closed S" |
|
442 |
using closed_translation a by auto |
|
443 |
} |
|
444 |
then show ?thesis by auto |
|
40377 | 445 |
qed |
446 |
||
447 |
lemma closure_affine_hull: |
|
53347 | 448 |
fixes S :: "'n::euclidean_space set" |
449 |
shows "closure S \<subseteq> affine hull S" |
|
44524 | 450 |
by (intro closure_minimal hull_subset affine_closed affine_affine_hull) |
40377 | 451 |
|
71225 | 452 |
lemma closed_affine_hull [iff]: |
453 |
fixes S :: "'n::euclidean_space set" |
|
454 |
shows "closed (affine hull S)" |
|
455 |
by (metis affine_affine_hull affine_closed) |
|
456 |
||
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
457 |
lemma closure_same_affine_hull [simp]: |
53347 | 458 |
fixes S :: "'n::euclidean_space set" |
40377 | 459 |
shows "affine hull (closure S) = affine hull S" |
53347 | 460 |
proof - |
461 |
have "affine hull (closure S) \<subseteq> affine hull S" |
|
462 |
using hull_mono[of "closure S" "affine hull S" "affine"] |
|
463 |
closure_affine_hull[of S] hull_hull[of "affine" S] |
|
464 |
by auto |
|
465 |
moreover have "affine hull (closure S) \<supseteq> affine hull S" |
|
466 |
using hull_mono[of "S" "closure S" "affine"] closure_subset by auto |
|
467 |
ultimately show ?thesis by auto |
|
49531 | 468 |
qed |
469 |
||
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
470 |
lemma closure_aff_dim [simp]: |
53347 | 471 |
fixes S :: "'n::euclidean_space set" |
40377 | 472 |
shows "aff_dim (closure S) = aff_dim S" |
53347 | 473 |
proof - |
474 |
have "aff_dim S \<le> aff_dim (closure S)" |
|
475 |
using aff_dim_subset closure_subset by auto |
|
476 |
moreover have "aff_dim (closure S) \<le> aff_dim (affine hull S)" |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
477 |
using aff_dim_subset closure_affine_hull by blast |
53347 | 478 |
moreover have "aff_dim (affine hull S) = aff_dim S" |
479 |
using aff_dim_affine_hull by auto |
|
480 |
ultimately show ?thesis by auto |
|
40377 | 481 |
qed |
482 |
||
483 |
lemma rel_interior_closure_convex_shrink: |
|
53347 | 484 |
fixes S :: "_::euclidean_space set" |
485 |
assumes "convex S" |
|
486 |
and "c \<in> rel_interior S" |
|
487 |
and "x \<in> closure S" |
|
488 |
and "e > 0" |
|
489 |
and "e \<le> 1" |
|
490 |
shows "x - e *\<^sub>R (x - c) \<in> rel_interior S" |
|
491 |
proof - |
|
492 |
obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S" |
|
493 |
using assms(2) unfolding mem_rel_interior_ball by auto |
|
494 |
have "\<exists>y \<in> S. norm (y - x) * (1 - e) < e * d" |
|
495 |
proof (cases "x \<in> S") |
|
496 |
case True |
|
72356 | 497 |
then show ?thesis using \<open>e > 0\<close> \<open>d > 0\<close> by force |
53347 | 498 |
next |
499 |
case False |
|
500 |
then have x: "x islimpt S" |
|
501 |
using assms(3)[unfolded closure_def] by auto |
|
502 |
show ?thesis |
|
503 |
proof (cases "e = 1") |
|
504 |
case True |
|
505 |
obtain y where "y \<in> S" "y \<noteq> x" "dist y x < 1" |
|
40377 | 506 |
using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto |
53347 | 507 |
then show ?thesis |
72356 | 508 |
unfolding True using \<open>d > 0\<close> by (force simp add: ) |
53347 | 509 |
next |
510 |
case False |
|
511 |
then have "0 < e * d / (1 - e)" and *: "1 - e > 0" |
|
68031 | 512 |
using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by auto |
53347 | 513 |
then obtain y where "y \<in> S" "y \<noteq> x" "dist y x < e * d / (1 - e)" |
40377 | 514 |
using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto |
53347 | 515 |
then show ?thesis |
72356 | 516 |
unfolding dist_norm using pos_less_divide_eq[OF *] by force |
53347 | 517 |
qed |
518 |
qed |
|
519 |
then obtain y where "y \<in> S" and y: "norm (y - x) * (1 - e) < e * d" |
|
520 |
by auto |
|
63040 | 521 |
define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" |
53347 | 522 |
have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" |
60420 | 523 |
unfolding z_def using \<open>e > 0\<close> |
68031 | 524 |
by (auto simp: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) |
53347 | 525 |
have zball: "z \<in> ball c d" |
526 |
using mem_ball z_def dist_norm[of c] |
|
527 |
using y and assms(4,5) |
|
70802
160eaf566bcb
formally augmented corresponding rules for field_simps
haftmann
parents:
70136
diff
changeset
|
528 |
by (simp add: norm_minus_commute) (simp add: field_simps) |
53347 | 529 |
have "x \<in> affine hull S" |
530 |
using closure_affine_hull assms by auto |
|
531 |
moreover have "y \<in> affine hull S" |
|
60420 | 532 |
using \<open>y \<in> S\<close> hull_subset[of S] by auto |
53347 | 533 |
moreover have "c \<in> affine hull S" |
534 |
using assms rel_interior_subset hull_subset[of S] by auto |
|
535 |
ultimately have "z \<in> affine hull S" |
|
49531 | 536 |
using z_def affine_affine_hull[of S] |
53347 | 537 |
mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"] |
538 |
assms |
|
70802
160eaf566bcb
formally augmented corresponding rules for field_simps
haftmann
parents:
70136
diff
changeset
|
539 |
by simp |
53347 | 540 |
then have "z \<in> S" using d zball by auto |
541 |
obtain d1 where "d1 > 0" and d1: "ball z d1 \<le> ball c d" |
|
40377 | 542 |
using zball open_ball[of c d] openE[of "ball c d" z] by auto |
53347 | 543 |
then have "ball z d1 \<inter> affine hull S \<subseteq> ball c d \<inter> affine hull S" |
544 |
by auto |
|
545 |
then have "ball z d1 \<inter> affine hull S \<subseteq> S" |
|
546 |
using d by auto |
|
547 |
then have "z \<in> rel_interior S" |
|
60420 | 548 |
using mem_rel_interior_ball using \<open>d1 > 0\<close> \<open>z \<in> S\<close> by auto |
53347 | 549 |
then have "y - e *\<^sub>R (y - z) \<in> rel_interior S" |
60420 | 550 |
using rel_interior_convex_shrink[of S z y e] assms \<open>y \<in> S\<close> by auto |
53347 | 551 |
then show ?thesis using * by auto |
552 |
qed |
|
553 |
||
62620
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
554 |
lemma rel_interior_eq: |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69710
diff
changeset
|
555 |
"rel_interior s = s \<longleftrightarrow> openin(top_of_set (affine hull s)) s" |
62620
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
556 |
using rel_open rel_open_def by blast |
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
557 |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
558 |
lemma rel_interior_openin: |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69710
diff
changeset
|
559 |
"openin(top_of_set (affine hull s)) s \<Longrightarrow> rel_interior s = s" |
62620
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
560 |
by (simp add: rel_interior_eq) |
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
561 |
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
562 |
lemma rel_interior_affine: |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
563 |
fixes S :: "'n::euclidean_space set" |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
564 |
shows "affine S \<Longrightarrow> rel_interior S = S" |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
565 |
using affine_rel_open rel_open_def by auto |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
566 |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
567 |
lemma rel_interior_eq_closure: |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
568 |
fixes S :: "'n::euclidean_space set" |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
569 |
shows "rel_interior S = closure S \<longleftrightarrow> affine S" |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
570 |
proof (cases "S = {}") |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
571 |
case True |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
572 |
then show ?thesis |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
573 |
by auto |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
574 |
next |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
575 |
case False show ?thesis |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
576 |
proof |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
577 |
assume eq: "rel_interior S = closure S" |
72356 | 578 |
have "openin (top_of_set (affine hull S)) S" |
579 |
by (metis eq closure_subset openin_rel_interior rel_interior_subset subset_antisym) |
|
580 |
moreover have "closedin (top_of_set (affine hull S)) S" |
|
581 |
by (metis closed_subset closure_subset_eq eq hull_subset rel_interior_subset) |
|
582 |
ultimately have "S = {} \<or> S = affine hull S" |
|
583 |
using convex_connected connected_clopen convex_affine_hull by metis |
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
584 |
with False have "affine hull S = S" |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
585 |
by auto |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
586 |
then show "affine S" |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
587 |
by (metis affine_hull_eq) |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
588 |
next |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
589 |
assume "affine S" |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
590 |
then show "rel_interior S = closure S" |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
591 |
by (simp add: rel_interior_affine affine_closed) |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
592 |
qed |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
593 |
qed |
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
594 |
|
40377 | 595 |
|
70136 | 596 |
subsubsection\<^marker>\<open>tag unimportant\<close>\<open>Relative interior preserves under linear transformations\<close> |
40377 | 597 |
|
598 |
lemma rel_interior_translation_aux: |
|
53347 | 599 |
fixes a :: "'n::euclidean_space" |
600 |
shows "((\<lambda>x. a + x) ` rel_interior S) \<subseteq> rel_interior ((\<lambda>x. a + x) ` S)" |
|
601 |
proof - |
|
602 |
{ |
|
603 |
fix x |
|
604 |
assume x: "x \<in> rel_interior S" |
|
605 |
then obtain T where "open T" "x \<in> T \<inter> S" "T \<inter> affine hull S \<subseteq> S" |
|
606 |
using mem_rel_interior[of x S] by auto |
|
607 |
then have "open ((\<lambda>x. a + x) ` T)" |
|
608 |
and "a + x \<in> ((\<lambda>x. a + x) ` T) \<inter> ((\<lambda>x. a + x) ` S)" |
|
609 |
and "((\<lambda>x. a + x) ` T) \<inter> affine hull ((\<lambda>x. a + x) ` S) \<subseteq> (\<lambda>x. a + x) ` S" |
|
610 |
using affine_hull_translation[of a S] open_translation[of T a] x by auto |
|
611 |
then have "a + x \<in> rel_interior ((\<lambda>x. a + x) ` S)" |
|
612 |
using mem_rel_interior[of "a+x" "((\<lambda>x. a + x) ` S)"] by auto |
|
613 |
} |
|
614 |
then show ?thesis by auto |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
615 |
qed |
40377 | 616 |
|
617 |
lemma rel_interior_translation: |
|
53347 | 618 |
fixes a :: "'n::euclidean_space" |
619 |
shows "rel_interior ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` rel_interior S" |
|
620 |
proof - |
|
621 |
have "(\<lambda>x. (-a) + x) ` rel_interior ((\<lambda>x. a + x) ` S) \<subseteq> rel_interior S" |
|
622 |
using rel_interior_translation_aux[of "-a" "(\<lambda>x. a + x) ` S"] |
|
623 |
translation_assoc[of "-a" "a"] |
|
624 |
by auto |
|
625 |
then have "((\<lambda>x. a + x) ` rel_interior S) \<supseteq> rel_interior ((\<lambda>x. a + x) ` S)" |
|
67399 | 626 |
using translation_inverse_subset[of a "rel_interior ((+) a ` S)" "rel_interior S"] |
53347 | 627 |
by auto |
628 |
then show ?thesis |
|
629 |
using rel_interior_translation_aux[of a S] by auto |
|
40377 | 630 |
qed |
631 |
||
632 |
||
633 |
lemma affine_hull_linear_image: |
|
53347 | 634 |
assumes "bounded_linear f" |
635 |
shows "f ` (affine hull s) = affine hull f ` s" |
|
636 |
proof - |
|
40377 | 637 |
interpret f: bounded_linear f by fact |
68058 | 638 |
have "affine {x. f x \<in> affine hull f ` s}" |
53347 | 639 |
unfolding affine_def |
68031 | 640 |
by (auto simp: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format]) |
68058 | 641 |
moreover have "affine {x. x \<in> f ` (affine hull s)}" |
53347 | 642 |
using affine_affine_hull[unfolded affine_def, of s] |
68031 | 643 |
unfolding affine_def by (auto simp: f.scaleR [symmetric] f.add [symmetric]) |
68058 | 644 |
ultimately show ?thesis |
645 |
by (auto simp: hull_inc elim!: hull_induct) |
|
646 |
qed |
|
40377 | 647 |
|
648 |
||
649 |
lemma rel_interior_injective_on_span_linear_image: |
|
53347 | 650 |
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" |
651 |
and S :: "'m::euclidean_space set" |
|
652 |
assumes "bounded_linear f" |
|
653 |
and "inj_on f (span S)" |
|
654 |
shows "rel_interior (f ` S) = f ` (rel_interior S)" |
|
655 |
proof - |
|
656 |
{ |
|
657 |
fix z |
|
658 |
assume z: "z \<in> rel_interior (f ` S)" |
|
659 |
then have "z \<in> f ` S" |
|
660 |
using rel_interior_subset[of "f ` S"] by auto |
|
661 |
then obtain x where x: "x \<in> S" "f x = z" by auto |
|
662 |
obtain e2 where e2: "e2 > 0" "cball z e2 \<inter> affine hull (f ` S) \<subseteq> (f ` S)" |
|
663 |
using z rel_interior_cball[of "f ` S"] by auto |
|
664 |
obtain K where K: "K > 0" "\<And>x. norm (f x) \<le> norm x * K" |
|
665 |
using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto |
|
63040 | 666 |
define e1 where "e1 = 1 / K" |
53347 | 667 |
then have e1: "e1 > 0" "\<And>x. e1 * norm (f x) \<le> norm x" |
668 |
using K pos_le_divide_eq[of e1] by auto |
|
63040 | 669 |
define e where "e = e1 * e2" |
56544 | 670 |
then have "e > 0" using e1 e2 by auto |
53347 | 671 |
{ |
672 |
fix y |
|
673 |
assume y: "y \<in> cball x e \<inter> affine hull S" |
|
674 |
then have h1: "f y \<in> affine hull (f ` S)" |
|
675 |
using affine_hull_linear_image[of f S] assms by auto |
|
676 |
from y have "norm (x-y) \<le> e1 * e2" |
|
677 |
using cball_def[of x e] dist_norm[of x y] e_def by auto |
|
678 |
moreover have "f x - f y = f (x - y)" |
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
679 |
using assms linear_diff[of f x y] linear_conv_bounded_linear[of f] by auto |
53347 | 680 |
moreover have "e1 * norm (f (x-y)) \<le> norm (x - y)" |
681 |
using e1 by auto |
|
682 |
ultimately have "e1 * norm ((f x)-(f y)) \<le> e1 * e2" |
|
683 |
by auto |
|
684 |
then have "f y \<in> cball z e2" |
|
685 |
using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1 x by auto |
|
686 |
then have "f y \<in> f ` S" |
|
687 |
using y e2 h1 by auto |
|
688 |
then have "y \<in> S" |
|
689 |
using assms y hull_subset[of S] affine_hull_subset_span |
|
61520
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
690 |
inj_on_image_mem_iff [OF \<open>inj_on f (span S)\<close>] |
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67982
diff
changeset
|
691 |
by (metis Int_iff span_superset subsetCE) |
53347 | 692 |
} |
693 |
then have "z \<in> f ` (rel_interior S)" |
|
60420 | 694 |
using mem_rel_interior_cball[of x S] \<open>e > 0\<close> x by auto |
49531 | 695 |
} |
53347 | 696 |
moreover |
697 |
{ |
|
698 |
fix x |
|
699 |
assume x: "x \<in> rel_interior S" |
|
54465 | 700 |
then obtain e2 where e2: "e2 > 0" "cball x e2 \<inter> affine hull S \<subseteq> S" |
53347 | 701 |
using rel_interior_cball[of S] by auto |
702 |
have "x \<in> S" using x rel_interior_subset by auto |
|
703 |
then have *: "f x \<in> f ` S" by auto |
|
704 |
have "\<forall>x\<in>span S. f x = 0 \<longrightarrow> x = 0" |
|
705 |
using assms subspace_span linear_conv_bounded_linear[of f] |
|
706 |
linear_injective_on_subspace_0[of f "span S"] |
|
707 |
by auto |
|
708 |
then obtain e1 where e1: "e1 > 0" "\<forall>x \<in> span S. e1 * norm x \<le> norm (f x)" |
|
709 |
using assms injective_imp_isometric[of "span S" f] |
|
710 |
subspace_span[of S] closed_subspace[of "span S"] |
|
711 |
by auto |
|
63040 | 712 |
define e where "e = e1 * e2" |
56544 | 713 |
hence "e > 0" using e1 e2 by auto |
53347 | 714 |
{ |
715 |
fix y |
|
716 |
assume y: "y \<in> cball (f x) e \<inter> affine hull (f ` S)" |
|
717 |
then have "y \<in> f ` (affine hull S)" |
|
718 |
using affine_hull_linear_image[of f S] assms by auto |
|
719 |
then obtain xy where xy: "xy \<in> affine hull S" "f xy = y" by auto |
|
720 |
with y have "norm (f x - f xy) \<le> e1 * e2" |
|
721 |
using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto |
|
722 |
moreover have "f x - f xy = f (x - xy)" |
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
723 |
using assms linear_diff[of f x xy] linear_conv_bounded_linear[of f] by auto |
53347 | 724 |
moreover have *: "x - xy \<in> span S" |
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
725 |
using subspace_diff[of "span S" x xy] subspace_span \<open>x \<in> S\<close> xy |
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67982
diff
changeset
|
726 |
affine_hull_subset_span[of S] span_superset |
53347 | 727 |
by auto |
728 |
moreover from * have "e1 * norm (x - xy) \<le> norm (f (x - xy))" |
|
729 |
using e1 by auto |
|
730 |
ultimately have "e1 * norm (x - xy) \<le> e1 * e2" |
|
731 |
by auto |
|
732 |
then have "xy \<in> cball x e2" |
|
733 |
using cball_def[of x e2] dist_norm[of x xy] e1 by auto |
|
734 |
then have "y \<in> f ` S" |
|
735 |
using xy e2 by auto |
|
736 |
} |
|
737 |
then have "f x \<in> rel_interior (f ` S)" |
|
60420 | 738 |
using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * \<open>e > 0\<close> by auto |
49531 | 739 |
} |
53347 | 740 |
ultimately show ?thesis by auto |
40377 | 741 |
qed |
742 |
||
743 |
lemma rel_interior_injective_linear_image: |
|
53347 | 744 |
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" |
745 |
assumes "bounded_linear f" |
|
746 |
and "inj f" |
|
747 |
shows "rel_interior (f ` S) = f ` (rel_interior S)" |
|
748 |
using assms rel_interior_injective_on_span_linear_image[of f S] |
|
749 |
subset_inj_on[of f "UNIV" "span S"] |
|
750 |
by auto |
|
751 |
||
40377 | 752 |
|
70136 | 753 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Openness and compactness are preserved by convex hull operation\<close> |
33175 | 754 |
|
34964 | 755 |
lemma open_convex_hull[intro]: |
68052 | 756 |
fixes S :: "'a::real_normed_vector set" |
757 |
assumes "open S" |
|
758 |
shows "open (convex hull S)" |
|
759 |
proof (clarsimp simp: open_contains_cball convex_hull_explicit) |
|
760 |
fix T and u :: "'a\<Rightarrow>real" |
|
761 |
assume obt: "finite T" "T\<subseteq>S" "\<forall>x\<in>T. 0 \<le> u x" "sum u T = 1" |
|
53347 | 762 |
|
763 |
from assms[unfolded open_contains_cball] obtain b |
|
68052 | 764 |
where b: "\<And>x. x\<in>S \<Longrightarrow> 0 < b x \<and> cball x (b x) \<subseteq> S" by metis |
765 |
have "b ` T \<noteq> {}" |
|
56889
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
hoelzl
parents:
56571
diff
changeset
|
766 |
using obt by auto |
68052 | 767 |
define i where "i = b ` T" |
768 |
let ?\<Phi> = "\<lambda>y. \<exists>F. finite F \<and> F \<subseteq> S \<and> (\<exists>u. (\<forall>x\<in>F. 0 \<le> u x) \<and> sum u F = 1 \<and> (\<Sum>v\<in>F. u v *\<^sub>R v) = y)" |
|
769 |
let ?a = "\<Sum>v\<in>T. u v *\<^sub>R v" |
|
770 |
show "\<exists>e > 0. cball ?a e \<subseteq> {y. ?\<Phi> y}" |
|
771 |
proof (intro exI subsetI conjI) |
|
53347 | 772 |
show "0 < Min i" |
68052 | 773 |
unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \<open>b ` T\<noteq>{}\<close>] |
774 |
using b \<open>T\<subseteq>S\<close> by auto |
|
53347 | 775 |
next |
776 |
fix y |
|
68052 | 777 |
assume "y \<in> cball ?a (Min i)" |
778 |
then have y: "norm (?a - y) \<le> Min i" |
|
53347 | 779 |
unfolding dist_norm[symmetric] by auto |
68052 | 780 |
{ fix x |
781 |
assume "x \<in> T" |
|
53347 | 782 |
then have "Min i \<le> b x" |
68052 | 783 |
by (simp add: i_def obt(1)) |
784 |
then have "x + (y - ?a) \<in> cball x (b x)" |
|
53347 | 785 |
using y unfolding mem_cball dist_norm by auto |
68052 | 786 |
moreover have "x \<in> S" |
787 |
using \<open>x\<in>T\<close> \<open>T\<subseteq>S\<close> by auto |
|
788 |
ultimately have "x + (y - ?a) \<in> S" |
|
789 |
using y b by blast |
|
53347 | 790 |
} |
33175 | 791 |
moreover |
68052 | 792 |
have *: "inj_on (\<lambda>v. v + (y - ?a)) T" |
53347 | 793 |
unfolding inj_on_def by auto |
68052 | 794 |
have "(\<Sum>v\<in>(\<lambda>v. v + (y - ?a)) ` T. u (v - (y - ?a)) *\<^sub>R v) = y" |
795 |
unfolding sum.reindex[OF *] o_def using obt(4) |
|
64267 | 796 |
by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib) |
68052 | 797 |
ultimately show "y \<in> {y. ?\<Phi> y}" |
798 |
proof (intro CollectI exI conjI) |
|
799 |
show "finite ((\<lambda>v. v + (y - ?a)) ` T)" |
|
800 |
by (simp add: obt(1)) |
|
801 |
show "sum (\<lambda>v. u (v - (y - ?a))) ((\<lambda>v. v + (y - ?a)) ` T) = 1" |
|
802 |
unfolding sum.reindex[OF *] o_def using obt(4) by auto |
|
803 |
qed (use obt(1, 3) in auto) |
|
33175 | 804 |
qed |
805 |
qed |
|
806 |
||
807 |
lemma compact_convex_combinations: |
|
68052 | 808 |
fixes S T :: "'a::real_normed_vector set" |
809 |
assumes "compact S" "compact T" |
|
810 |
shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> S \<and> y \<in> T}" |
|
53347 | 811 |
proof - |
68052 | 812 |
let ?X = "{0..1} \<times> S \<times> T" |
33175 | 813 |
let ?h = "(\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" |
68052 | 814 |
have *: "{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \<le> u \<and> u \<le> 1 \<and> x \<in> S \<and> y \<in> T} = ?h ` ?X" |
815 |
by force |
|
56188 | 816 |
have "continuous_on ?X (\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" |
33175 | 817 |
unfolding continuous_on by (rule ballI) (intro tendsto_intros) |
68052 | 818 |
with assms show ?thesis |
819 |
by (simp add: * compact_Times compact_continuous_image) |
|
33175 | 820 |
qed |
821 |
||
44525
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
822 |
lemma finite_imp_compact_convex_hull: |
68052 | 823 |
fixes S :: "'a::real_normed_vector set" |
824 |
assumes "finite S" |
|
825 |
shows "compact (convex hull S)" |
|
826 |
proof (cases "S = {}") |
|
53347 | 827 |
case True |
828 |
then show ?thesis by simp |
|
44525
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
829 |
next |
53347 | 830 |
case False |
831 |
with assms show ?thesis |
|
44525
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
832 |
proof (induct rule: finite_ne_induct) |
53347 | 833 |
case (singleton x) |
834 |
show ?case by simp |
|
44525
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
835 |
next |
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
836 |
case (insert x A) |
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
837 |
let ?f = "\<lambda>(u, y::'a). u *\<^sub>R x + (1 - u) *\<^sub>R y" |
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
838 |
let ?T = "{0..1::real} \<times> (convex hull A)" |
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
839 |
have "continuous_on ?T ?f" |
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
840 |
unfolding split_def continuous_on by (intro ballI tendsto_intros) |
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
841 |
moreover have "compact ?T" |
56188 | 842 |
by (intro compact_Times compact_Icc insert) |
44525
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
843 |
ultimately have "compact (?f ` ?T)" |
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
844 |
by (rule compact_continuous_image) |
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
845 |
also have "?f ` ?T = convex hull (insert x A)" |
60420 | 846 |
unfolding convex_hull_insert [OF \<open>A \<noteq> {}\<close>] |
44525
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
847 |
apply safe |
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
848 |
apply (rule_tac x=a in exI, simp) |
68031 | 849 |
apply (rule_tac x="1 - a" in exI, simp, fast) |
44525
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
850 |
apply (rule_tac x="(u, b)" in image_eqI, simp_all) |
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
851 |
done |
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
852 |
finally show "compact (convex hull (insert x A))" . |
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
853 |
qed |
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
854 |
qed |
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
855 |
|
53347 | 856 |
lemma compact_convex_hull: |
68052 | 857 |
fixes S :: "'a::euclidean_space set" |
858 |
assumes "compact S" |
|
859 |
shows "compact (convex hull S)" |
|
860 |
proof (cases "S = {}") |
|
53347 | 861 |
case True |
862 |
then show ?thesis using compact_empty by simp |
|
33175 | 863 |
next |
53347 | 864 |
case False |
68052 | 865 |
then obtain w where "w \<in> S" by auto |
53347 | 866 |
show ?thesis |
68052 | 867 |
unfolding caratheodory[of S] |
53347 | 868 |
proof (induct ("DIM('a) + 1")) |
869 |
case 0 |
|
68052 | 870 |
have *: "{x.\<exists>sa. finite sa \<and> sa \<subseteq> S \<and> card sa \<le> 0 \<and> x \<in> convex hull sa} = {}" |
36362
06475a1547cb
fix lots of looping simp calls and other warnings
huffman
parents:
36341
diff
changeset
|
871 |
using compact_empty by auto |
53347 | 872 |
from 0 show ?case unfolding * by simp |
33175 | 873 |
next |
874 |
case (Suc n) |
|
53347 | 875 |
show ?case |
876 |
proof (cases "n = 0") |
|
877 |
case True |
|
68052 | 878 |
have "{x. \<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T} = S" |
53347 | 879 |
unfolding set_eq_iff and mem_Collect_eq |
880 |
proof (rule, rule) |
|
881 |
fix x |
|
68052 | 882 |
assume "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T" |
883 |
then obtain T where T: "finite T" "T \<subseteq> S" "card T \<le> Suc n" "x \<in> convex hull T" |
|
53347 | 884 |
by auto |
68052 | 885 |
show "x \<in> S" |
886 |
proof (cases "card T = 0") |
|
53347 | 887 |
case True |
888 |
then show ?thesis |
|
68052 | 889 |
using T(4) unfolding card_0_eq[OF T(1)] by simp |
33175 | 890 |
next |
53347 | 891 |
case False |
68052 | 892 |
then have "card T = Suc 0" using T(3) \<open>n=0\<close> by auto |
893 |
then obtain a where "T = {a}" unfolding card_Suc_eq by auto |
|
894 |
then show ?thesis using T(2,4) by simp |
|
33175 | 895 |
qed |
896 |
next |
|
68052 | 897 |
fix x assume "x\<in>S" |
898 |
then show "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T" |
|
72356 | 899 |
by (rule_tac x="{x}" in exI) (use convex_hull_singleton in auto) |
53347 | 900 |
qed |
901 |
then show ?thesis using assms by simp |
|
33175 | 902 |
next |
53347 | 903 |
case False |
68052 | 904 |
have "{x. \<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T} = |
53347 | 905 |
{(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. |
68052 | 906 |
0 \<le> u \<and> u \<le> 1 \<and> x \<in> S \<and> y \<in> {x. \<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> n \<and> x \<in> convex hull T}}" |
53347 | 907 |
unfolding set_eq_iff and mem_Collect_eq |
908 |
proof (rule, rule) |
|
909 |
fix x |
|
910 |
assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and> |
|
68052 | 911 |
0 \<le> c \<and> c \<le> 1 \<and> u \<in> S \<and> (\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> n \<and> v \<in> convex hull T)" |
912 |
then obtain u v c T where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v" |
|
913 |
"0 \<le> c \<and> c \<le> 1" "u \<in> S" "finite T" "T \<subseteq> S" "card T \<le> n" "v \<in> convex hull T" |
|
53347 | 914 |
by auto |
68052 | 915 |
moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \<in> convex hull insert u T" |
72356 | 916 |
by (meson convexD_alt convex_convex_hull hull_inc hull_mono in_mono insertCI obt(2) obt(7) subset_insertI) |
68052 | 917 |
ultimately show "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T" |
72356 | 918 |
by (rule_tac x="insert u T" in exI) (auto simp: card_insert_if) |
33175 | 919 |
next |
53347 | 920 |
fix x |
68052 | 921 |
assume "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T" |
922 |
then obtain T where T: "finite T" "T \<subseteq> S" "card T \<le> Suc n" "x \<in> convex hull T" |
|
53347 | 923 |
by auto |
924 |
show "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and> |
|
68052 | 925 |
0 \<le> c \<and> c \<le> 1 \<and> u \<in> S \<and> (\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> n \<and> v \<in> convex hull T)" |
926 |
proof (cases "card T = Suc n") |
|
53347 | 927 |
case False |
68052 | 928 |
then have "card T \<le> n" using T(3) by auto |
53347 | 929 |
then show ?thesis |
68052 | 930 |
using \<open>w\<in>S\<close> and T |
72356 | 931 |
by (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) auto |
33175 | 932 |
next |
53347 | 933 |
case True |
68052 | 934 |
then obtain a u where au: "T = insert a u" "a\<notin>u" |
72356 | 935 |
by (metis card_le_Suc_iff order_refl) |
53347 | 936 |
show ?thesis |
937 |
proof (cases "u = {}") |
|
938 |
case True |
|
68052 | 939 |
then have "x = a" using T(4)[unfolded au] by auto |
60420 | 940 |
show ?thesis unfolding \<open>x = a\<close> |
72356 | 941 |
using T \<open>n \<noteq> 0\<close> unfolding au |
942 |
by (rule_tac x=a in exI, rule_tac x=a in exI, rule_tac x=1 in exI) force |
|
33175 | 943 |
next |
53347 | 944 |
case False |
945 |
obtain ux vx b where obt: "ux\<ge>0" "vx\<ge>0" "ux + vx = 1" |
|
946 |
"b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b" |
|
68052 | 947 |
using T(4)[unfolded au convex_hull_insert[OF False]] |
53347 | 948 |
by auto |
949 |
have *: "1 - vx = ux" using obt(3) by auto |
|
950 |
show ?thesis |
|
72356 | 951 |
using obt T(1-3) card_insert_disjoint[OF _ au(2)] unfolding au * |
952 |
by (rule_tac x=a in exI, rule_tac x=b in exI, rule_tac x=vx in exI) force |
|
33175 | 953 |
qed |
954 |
qed |
|
955 |
qed |
|
53347 | 956 |
then show ?thesis |
957 |
using compact_convex_combinations[OF assms Suc] by simp |
|
33175 | 958 |
qed |
36362
06475a1547cb
fix lots of looping simp calls and other warnings
huffman
parents:
36341
diff
changeset
|
959 |
qed |
33175 | 960 |
qed |
961 |
||
53347 | 962 |
|
70136 | 963 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Extremal points of a simplex are some vertices\<close> |
33175 | 964 |
|
965 |
lemma dist_increases_online: |
|
966 |
fixes a b d :: "'a::real_inner" |
|
967 |
assumes "d \<noteq> 0" |
|
968 |
shows "dist a (b + d) > dist a b \<or> dist a (b - d) > dist a b" |
|
53347 | 969 |
proof (cases "inner a d - inner b d > 0") |
970 |
case True |
|
971 |
then have "0 < inner d d + (inner a d * 2 - inner b d * 2)" |
|
972 |
using assms |
|
72356 | 973 |
by (intro add_pos_pos) auto |
53347 | 974 |
then show ?thesis |
975 |
unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff |
|
72356 | 976 |
by (simp add: algebra_simps inner_commute) |
33175 | 977 |
next |
53347 | 978 |
case False |
979 |
then have "0 < inner d d + (inner b d * 2 - inner a d * 2)" |
|
980 |
using assms |
|
72356 | 981 |
by (intro add_pos_nonneg) auto |
53347 | 982 |
then show ?thesis |
983 |
unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff |
|
72356 | 984 |
by (simp add: algebra_simps inner_commute) |
33175 | 985 |
qed |
986 |
||
987 |
lemma norm_increases_online: |
|
988 |
fixes d :: "'a::real_inner" |
|
53347 | 989 |
shows "d \<noteq> 0 \<Longrightarrow> norm (a + d) > norm a \<or> norm(a - d) > norm a" |
33175 | 990 |
using dist_increases_online[of d a 0] unfolding dist_norm by auto |
991 |
||
992 |
lemma simplex_furthest_lt: |
|
68052 | 993 |
fixes S :: "'a::real_inner set" |
994 |
assumes "finite S" |
|
995 |
shows "\<forall>x \<in> convex hull S. x \<notin> S \<longrightarrow> (\<exists>y \<in> convex hull S. norm (x - a) < norm(y - a))" |
|
53347 | 996 |
using assms |
997 |
proof induct |
|
68052 | 998 |
fix x S |
999 |
assume as: "finite S" "x\<notin>S" "\<forall>x\<in>convex hull S. x \<notin> S \<longrightarrow> (\<exists>y\<in>convex hull S. norm (x - a) < norm (y - a))" |
|
1000 |
show "\<forall>xa\<in>convex hull insert x S. xa \<notin> insert x S \<longrightarrow> |
|
1001 |
(\<exists>y\<in>convex hull insert x S. norm (xa - a) < norm (y - a))" |
|
1002 |
proof (intro impI ballI, cases "S = {}") |
|
53347 | 1003 |
case False |
1004 |
fix y |
|
68052 | 1005 |
assume y: "y \<in> convex hull insert x S" "y \<notin> insert x S" |
1006 |
obtain u v b where obt: "u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull S" "y = u *\<^sub>R x + v *\<^sub>R b" |
|
33175 | 1007 |
using y(1)[unfolded convex_hull_insert[OF False]] by auto |
68052 | 1008 |
show "\<exists>z\<in>convex hull insert x S. norm (y - a) < norm (z - a)" |
1009 |
proof (cases "y \<in> convex hull S") |
|
53347 | 1010 |
case True |
68052 | 1011 |
then obtain z where "z \<in> convex hull S" "norm (y - a) < norm (z - a)" |
33175 | 1012 |
using as(3)[THEN bspec[where x=y]] and y(2) by auto |
53347 | 1013 |
then show ?thesis |
72356 | 1014 |
by (meson hull_mono subsetD subset_insertI) |
33175 | 1015 |
next |
53347 | 1016 |
case False |
1017 |
show ?thesis |
|
72356 | 1018 |
proof (cases "u = 0 \<or> v = 0") |
1019 |
case True |
|
1020 |
with False show ?thesis |
|
1021 |
using obt y by auto |
|
33175 | 1022 |
next |
72356 | 1023 |
case False |
53347 | 1024 |
then obtain w where w: "w>0" "w<u" "w<v" |
68527
2f4e2aab190a
Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents:
68074
diff
changeset
|
1025 |
using field_lbound_gt_zero[of u v] and obt(1,2) by auto |
53347 | 1026 |
have "x \<noteq> b" |
1027 |
proof |
|
1028 |
assume "x = b" |
|
1029 |
then have "y = b" unfolding obt(5) |
|
68031 | 1030 |
using obt(3) by (auto simp: scaleR_left_distrib[symmetric]) |
72356 | 1031 |
then show False using obt(4) and False |
1032 |
using \<open>x = b\<close> y(2) by blast |
|
53347 | 1033 |
qed |
1034 |
then have *: "w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto |
|
1035 |
show ?thesis |
|
1036 |
using dist_increases_online[OF *, of a y] |
|
1037 |
proof (elim disjE) |
|
33175 | 1038 |
assume "dist a y < dist a (y + w *\<^sub>R (x - b))" |
53347 | 1039 |
then have "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)" |
1040 |
unfolding dist_commute[of a] |
|
1041 |
unfolding dist_norm obt(5) |
|
1042 |
by (simp add: algebra_simps) |
|
68052 | 1043 |
moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x S" |
1044 |
unfolding convex_hull_insert[OF \<open>S\<noteq>{}\<close>] |
|
1045 |
proof (intro CollectI conjI exI) |
|
1046 |
show "u + w \<ge> 0" "v - w \<ge> 0" |
|
1047 |
using obt(1) w by auto |
|
1048 |
qed (use obt in auto) |
|
33175 | 1049 |
ultimately show ?thesis by auto |
1050 |
next |
|
1051 |
assume "dist a y < dist a (y - w *\<^sub>R (x - b))" |
|
53347 | 1052 |
then have "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)" |
1053 |
unfolding dist_commute[of a] |
|
1054 |
unfolding dist_norm obt(5) |
|
1055 |
by (simp add: algebra_simps) |
|
68052 | 1056 |
moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x S" |
1057 |
unfolding convex_hull_insert[OF \<open>S\<noteq>{}\<close>] |
|
1058 |
proof (intro CollectI conjI exI) |
|
1059 |
show "u - w \<ge> 0" "v + w \<ge> 0" |
|
1060 |
using obt(1) w by auto |
|
1061 |
qed (use obt in auto) |
|
33175 | 1062 |
ultimately show ?thesis by auto |
1063 |
qed |
|
72356 | 1064 |
qed |
33175 | 1065 |
qed |
1066 |
qed auto |
|
68031 | 1067 |
qed (auto simp: assms) |
33175 | 1068 |
|
1069 |
lemma simplex_furthest_le: |
|
68052 | 1070 |
fixes S :: "'a::real_inner set" |
1071 |
assumes "finite S" |
|
1072 |
and "S \<noteq> {}" |
|
1073 |
shows "\<exists>y\<in>S. \<forall>x\<in> convex hull S. norm (x - a) \<le> norm (y - a)" |
|
53347 | 1074 |
proof - |
68052 | 1075 |
have "convex hull S \<noteq> {}" |
1076 |
using hull_subset[of S convex] and assms(2) by auto |
|
1077 |
then obtain x where x: "x \<in> convex hull S" "\<forall>y\<in>convex hull S. norm (y - a) \<le> norm (x - a)" |
|
1078 |
using distance_attains_sup[OF finite_imp_compact_convex_hull[OF \<open>finite S\<close>], of a] |
|
53347 | 1079 |
unfolding dist_commute[of a] |
1080 |
unfolding dist_norm |
|
1081 |
by auto |
|
1082 |
show ?thesis |
|
68052 | 1083 |
proof (cases "x \<in> S") |
53347 | 1084 |
case False |
68052 | 1085 |
then obtain y where "y \<in> convex hull S" "norm (x - a) < norm (y - a)" |
53347 | 1086 |
using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) |
1087 |
by auto |
|
1088 |
then show ?thesis |
|
1089 |
using x(2)[THEN bspec[where x=y]] by auto |
|
1090 |
next |
|
1091 |
case True |
|
1092 |
with x show ?thesis by auto |
|
1093 |
qed |
|
33175 | 1094 |
qed |
1095 |
||
1096 |
lemma simplex_furthest_le_exists: |
|
68052 | 1097 |
fixes S :: "('a::real_inner) set" |
1098 |
shows "finite S \<Longrightarrow> \<forall>x\<in>(convex hull S). \<exists>y\<in>S. norm (x - a) \<le> norm (y - a)" |
|
1099 |
using simplex_furthest_le[of S] by (cases "S = {}") auto |
|
33175 | 1100 |
|
1101 |
lemma simplex_extremal_le: |
|
68052 | 1102 |
fixes S :: "'a::real_inner set" |
1103 |
assumes "finite S" |
|
1104 |
and "S \<noteq> {}" |
|
1105 |
shows "\<exists>u\<in>S. \<exists>v\<in>S. \<forall>x\<in>convex hull S. \<forall>y \<in> convex hull S. norm (x - y) \<le> norm (u - v)" |
|
53347 | 1106 |
proof - |
68052 | 1107 |
have "convex hull S \<noteq> {}" |
1108 |
using hull_subset[of S convex] and assms(2) by auto |
|
1109 |
then obtain u v where obt: "u \<in> convex hull S" "v \<in> convex hull S" |
|
1110 |
"\<forall>x\<in>convex hull S. \<forall>y\<in>convex hull S. norm (x - y) \<le> norm (u - v)" |
|
53347 | 1111 |
using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] |
1112 |
by (auto simp: dist_norm) |
|
1113 |
then show ?thesis |
|
68052 | 1114 |
proof (cases "u\<notin>S \<or> v\<notin>S", elim disjE) |
1115 |
assume "u \<notin> S" |
|
1116 |
then obtain y where "y \<in> convex hull S" "norm (u - v) < norm (y - v)" |
|
53347 | 1117 |
using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) |
1118 |
by auto |
|
1119 |
then show ?thesis |
|
1120 |
using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) |
|
1121 |
by auto |
|
33175 | 1122 |
next |
68052 | 1123 |
assume "v \<notin> S" |
1124 |
then obtain y where "y \<in> convex hull S" "norm (v - u) < norm (y - u)" |
|
53347 | 1125 |
using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) |
1126 |
by auto |
|
1127 |
then show ?thesis |
|
1128 |
using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1) |
|
68031 | 1129 |
by (auto simp: norm_minus_commute) |
33175 | 1130 |
qed auto |
49531 | 1131 |
qed |
33175 | 1132 |
|
1133 |
lemma simplex_extremal_le_exists: |
|
68052 | 1134 |
fixes S :: "'a::real_inner set" |
1135 |
shows "finite S \<Longrightarrow> x \<in> convex hull S \<Longrightarrow> y \<in> convex hull S \<Longrightarrow> |
|
1136 |
\<exists>u\<in>S. \<exists>v\<in>S. norm (x - y) \<le> norm (u - v)" |
|
1137 |
using convex_hull_empty simplex_extremal_le[of S] |
|
1138 |
by(cases "S = {}") auto |
|
53347 | 1139 |
|
33175 | 1140 |
|
67968 | 1141 |
subsection \<open>Closest point of a convex set is unique, with a continuous projection\<close> |
33175 | 1142 |
|
70136 | 1143 |
definition\<^marker>\<open>tag important\<close> closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a" |
68052 | 1144 |
where "closest_point S a = (SOME x. x \<in> S \<and> (\<forall>y\<in>S. dist a x \<le> dist a y))" |
33175 | 1145 |
|
1146 |
lemma closest_point_exists: |
|
68052 | 1147 |
assumes "closed S" |
1148 |
and "S \<noteq> {}" |
|
71244 | 1149 |
shows closest_point_in_set: "closest_point S a \<in> S" |
68052 | 1150 |
and "\<forall>y\<in>S. dist a (closest_point S a) \<le> dist a y" |
53347 | 1151 |
unfolding closest_point_def |
72356 | 1152 |
by (rule_tac someI2_ex, auto intro: distance_attains_inf[OF assms(1,2), of a])+ |
53347 | 1153 |
|
68052 | 1154 |
lemma closest_point_le: "closed S \<Longrightarrow> x \<in> S \<Longrightarrow> dist a (closest_point S a) \<le> dist a x" |
1155 |
using closest_point_exists[of S] by auto |
|
33175 | 1156 |
|
1157 |
lemma closest_point_self: |
|
68052 | 1158 |
assumes "x \<in> S" |
1159 |
shows "closest_point S x = x" |
|
53347 | 1160 |
unfolding closest_point_def |
72356 | 1161 |
by (rule some1_equality, rule ex1I[of _ x]) (use assms in auto) |
53347 | 1162 |
|
68052 | 1163 |
lemma closest_point_refl: "closed S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> closest_point S x = x \<longleftrightarrow> x \<in> S" |
1164 |
using closest_point_in_set[of S x] closest_point_self[of x S] |
|
53347 | 1165 |
by auto |
33175 | 1166 |
|
36337 | 1167 |
lemma closer_points_lemma: |
33175 | 1168 |
assumes "inner y z > 0" |
1169 |
shows "\<exists>u>0. \<forall>v>0. v \<le> u \<longrightarrow> norm(v *\<^sub>R z - y) < norm y" |
|
53347 | 1170 |
proof - |
1171 |
have z: "inner z z > 0" |
|
1172 |
unfolding inner_gt_zero_iff using assms by auto |
|
68031 | 1173 |
have "norm (v *\<^sub>R z - y) < norm y" |
1174 |
if "0 < v" and "v \<le> inner y z / inner z z" for v |
|
1175 |
unfolding norm_lt using z assms that |
|
1176 |
by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \<open>0<v\<close>]) |
|
53347 | 1177 |
then show ?thesis |
68031 | 1178 |
using assms z |
1179 |
by (rule_tac x = "inner y z / inner z z" in exI) auto |
|
53347 | 1180 |
qed |
33175 | 1181 |
|
1182 |
lemma closer_point_lemma: |
|
1183 |
assumes "inner (y - x) (z - x) > 0" |
|
1184 |
shows "\<exists>u>0. u \<le> 1 \<and> dist (x + u *\<^sub>R (z - x)) y < dist x y" |
|
53347 | 1185 |
proof - |
1186 |
obtain u where "u > 0" |
|
72356 | 1187 |
and u: "\<And>v. \<lbrakk>0<v; v \<le> u\<rbrakk> \<Longrightarrow> norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)" |
33175 | 1188 |
using closer_points_lemma[OF assms] by auto |
53347 | 1189 |
show ?thesis |
72356 | 1190 |
using u[of "min u 1"] and \<open>u > 0\<close> |
1191 |
by (metis diff_diff_add dist_commute dist_norm less_eq_real_def not_less u zero_less_one) |
|
53347 | 1192 |
qed |
33175 | 1193 |
|
1194 |
lemma any_closest_point_dot: |
|
68052 | 1195 |
assumes "convex S" "closed S" "x \<in> S" "y \<in> S" "\<forall>z\<in>S. dist a x \<le> dist a z" |
33175 | 1196 |
shows "inner (a - x) (y - x) \<le> 0" |
53347 | 1197 |
proof (rule ccontr) |
1198 |
assume "\<not> ?thesis" |
|
1199 |
then obtain u where u: "u>0" "u\<le>1" "dist (x + u *\<^sub>R (y - x)) a < dist x a" |
|
1200 |
using closer_point_lemma[of a x y] by auto |
|
1201 |
let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y" |
|
68052 | 1202 |
have "?z \<in> S" |
61426
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61222
diff
changeset
|
1203 |
using convexD_alt[OF assms(1,3,4), of u] using u by auto |
53347 | 1204 |
then show False |
1205 |
using assms(5)[THEN bspec[where x="?z"]] and u(3) |
|
68031 | 1206 |
by (auto simp: dist_commute algebra_simps) |
53347 | 1207 |
qed |
33175 | 1208 |
|
1209 |
lemma any_closest_point_unique: |
|
36337 | 1210 |
fixes x :: "'a::real_inner" |
68052 | 1211 |
assumes "convex S" "closed S" "x \<in> S" "y \<in> S" |
1212 |
"\<forall>z\<in>S. dist a x \<le> dist a z" "\<forall>z\<in>S. dist a y \<le> dist a z" |
|
53347 | 1213 |
shows "x = y" |
1214 |
using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)] |
|
33175 | 1215 |
unfolding norm_pths(1) and norm_le_square |
68031 | 1216 |
by (auto simp: algebra_simps) |
33175 | 1217 |
|
1218 |
lemma closest_point_unique: |
|
68052 | 1219 |
assumes "convex S" "closed S" "x \<in> S" "\<forall>z\<in>S. dist a x \<le> dist a z" |
1220 |
shows "x = closest_point S a" |
|
1221 |
using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point S a"] |
|
33175 | 1222 |
using closest_point_exists[OF assms(2)] and assms(3) by auto |
1223 |
||
1224 |
lemma closest_point_dot: |
|
68052 | 1225 |
assumes "convex S" "closed S" "x \<in> S" |
1226 |
shows "inner (a - closest_point S a) (x - closest_point S a) \<le> 0" |
|
72356 | 1227 |
using any_closest_point_dot[OF assms(1,2) _ assms(3)] |
1228 |
by (metis assms(2) assms(3) closest_point_in_set closest_point_le empty_iff) |
|
33175 | 1229 |
|
1230 |
lemma closest_point_lt: |
|
68052 | 1231 |
assumes "convex S" "closed S" "x \<in> S" "x \<noteq> closest_point S a" |
1232 |
shows "dist a (closest_point S a) < dist a x" |
|
72356 | 1233 |
using closest_point_unique[where a=a] closest_point_le[where a=a] assms by fastforce |
33175 | 1234 |
|
69618 | 1235 |
lemma setdist_closest_point: |
1236 |
"\<lbrakk>closed S; S \<noteq> {}\<rbrakk> \<Longrightarrow> setdist {a} S = dist a (closest_point S a)" |
|
72356 | 1237 |
by (metis closest_point_exists(2) closest_point_in_set emptyE insert_iff setdist_unique) |
69618 | 1238 |
|
33175 | 1239 |
lemma closest_point_lipschitz: |
68052 | 1240 |
assumes "convex S" |
1241 |
and "closed S" "S \<noteq> {}" |
|
1242 |
shows "dist (closest_point S x) (closest_point S y) \<le> dist x y" |
|
53347 | 1243 |
proof - |
68052 | 1244 |
have "inner (x - closest_point S x) (closest_point S y - closest_point S x) \<le> 0" |
1245 |
and "inner (y - closest_point S y) (closest_point S x - closest_point S y) \<le> 0" |
|
72356 | 1246 |
by (simp_all add: assms closest_point_dot closest_point_in_set) |
53347 | 1247 |
then show ?thesis unfolding dist_norm and norm_le |
68052 | 1248 |
using inner_ge_zero[of "(x - closest_point S x) - (y - closest_point S y)"] |
53347 | 1249 |
by (simp add: inner_add inner_diff inner_commute) |
1250 |
qed |
|
33175 | 1251 |
|
1252 |
lemma continuous_at_closest_point: |
|
68052 | 1253 |
assumes "convex S" |
1254 |
and "closed S" |
|
1255 |
and "S \<noteq> {}" |
|
1256 |
shows "continuous (at x) (closest_point S)" |
|
49531 | 1257 |
unfolding continuous_at_eps_delta |
33175 | 1258 |
using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto |
1259 |
||
1260 |
lemma continuous_on_closest_point: |
|
68052 | 1261 |
assumes "convex S" |
1262 |
and "closed S" |
|
1263 |
and "S \<noteq> {}" |
|
1264 |
shows "continuous_on t (closest_point S)" |
|
53347 | 1265 |
by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms]) |
1266 |
||
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1267 |
proposition closest_point_in_rel_interior: |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1268 |
assumes "closed S" "S \<noteq> {}" and x: "x \<in> affine hull S" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1269 |
shows "closest_point S x \<in> rel_interior S \<longleftrightarrow> x \<in> rel_interior S" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1270 |
proof (cases "x \<in> S") |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1271 |
case True |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1272 |
then show ?thesis |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1273 |
by (simp add: closest_point_self) |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1274 |
next |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1275 |
case False |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1276 |
then have "False" if asm: "closest_point S x \<in> rel_interior S" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1277 |
proof - |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1278 |
obtain e where "e > 0" and clox: "closest_point S x \<in> S" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1279 |
and e: "cball (closest_point S x) e \<inter> affine hull S \<subseteq> S" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1280 |
using asm mem_rel_interior_cball by blast |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1281 |
then have clo_notx: "closest_point S x \<noteq> x" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1282 |
using \<open>x \<notin> S\<close> by auto |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1283 |
define y where "y \<equiv> closest_point S x - |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1284 |
(min 1 (e / norm(closest_point S x - x))) *\<^sub>R (closest_point S x - x)" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1285 |
have "x - y = (1 - min 1 (e / norm (closest_point S x - x))) *\<^sub>R (x - closest_point S x)" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1286 |
by (simp add: y_def algebra_simps) |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1287 |
then have "norm (x - y) = abs ((1 - min 1 (e / norm (closest_point S x - x)))) * norm(x - closest_point S x)" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1288 |
by simp |
68031 | 1289 |
also have "\<dots> < norm(x - closest_point S x)" |
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1290 |
using clo_notx \<open>e > 0\<close> |
70817
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents:
70802
diff
changeset
|
1291 |
by (auto simp: mult_less_cancel_right2 field_split_simps) |
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1292 |
finally have no_less: "norm (x - y) < norm (x - closest_point S x)" . |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1293 |
have "y \<in> affine hull S" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1294 |
unfolding y_def |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1295 |
by (meson affine_affine_hull clox hull_subset mem_affine_3_minus2 subsetD x) |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1296 |
moreover have "dist (closest_point S x) y \<le> e" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1297 |
using \<open>e > 0\<close> by (auto simp: y_def min_mult_distrib_right) |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1298 |
ultimately have "y \<in> S" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1299 |
using subsetD [OF e] by simp |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1300 |
then have "dist x (closest_point S x) \<le> dist x y" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1301 |
by (simp add: closest_point_le \<open>closed S\<close>) |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1302 |
with no_less show False |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1303 |
by (simp add: dist_norm) |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1304 |
qed |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1305 |
moreover have "x \<notin> rel_interior S" |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1306 |
using rel_interior_subset False by blast |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1307 |
ultimately show ?thesis by blast |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1308 |
qed |
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1309 |
|
33175 | 1310 |
|
70136 | 1311 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Various point-to-set separating/supporting hyperplane theorems\<close> |
33175 | 1312 |
|
1313 |
lemma supporting_hyperplane_closed_point: |
|
36337 | 1314 |
fixes z :: "'a::{real_inner,heine_borel}" |
68052 | 1315 |
assumes "convex S" |
1316 |
and "closed S" |
|
1317 |
and "S \<noteq> {}" |
|
1318 |
and "z \<notin> S" |
|
1319 |
shows "\<exists>a b. \<exists>y\<in>S. inner a z < b \<and> inner a y = b \<and> (\<forall>x\<in>S. inner a x \<ge> b)" |
|
53347 | 1320 |
proof - |
68052 | 1321 |
obtain y where "y \<in> S" and y: "\<forall>x\<in>S. dist z y \<le> dist z x" |
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
1322 |
by (metis distance_attains_inf[OF assms(2-3)]) |
53347 | 1323 |
show ?thesis |
68052 | 1324 |
proof (intro exI bexI conjI ballI) |
1325 |
show "(y - z) \<bullet> z < (y - z) \<bullet> y" |
|
1326 |
by (metis \<open>y \<in> S\<close> assms(4) diff_gt_0_iff_gt inner_commute inner_diff_left inner_gt_zero_iff right_minus_eq) |
|
1327 |
show "(y - z) \<bullet> y \<le> (y - z) \<bullet> x" if "x \<in> S" for x |
|
1328 |
proof (rule ccontr) |
|
1329 |
have *: "\<And>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> dist z y \<le> dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)" |
|
1330 |
using assms(1)[unfolded convex_alt] and y and \<open>x\<in>S\<close> and \<open>y\<in>S\<close> by auto |
|
1331 |
assume "\<not> (y - z) \<bullet> y \<le> (y - z) \<bullet> x" |
|
1332 |
then obtain v where "v > 0" "v \<le> 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" |
|
1333 |
using closer_point_lemma[of z y x] by (auto simp: inner_diff) |
|
1334 |
then show False |
|
1335 |
using *[of v] by (auto simp: dist_commute algebra_simps) |
|
1336 |
qed |
|
1337 |
qed (use \<open>y \<in> S\<close> in auto) |
|
33175 | 1338 |
qed |
1339 |
||
1340 |
lemma separating_hyperplane_closed_point: |
|
36337 | 1341 |
fixes z :: "'a::{real_inner,heine_borel}" |
68052 | 1342 |
assumes "convex S" |
1343 |
and "closed S" |
|
1344 |
and "z \<notin> S" |
|
1345 |
shows "\<exists>a b. inner a z < b \<and> (\<forall>x\<in>S. inner a x > b)" |
|
1346 |
proof (cases "S = {}") |
|
53347 | 1347 |
case True |
1348 |
then show ?thesis |
|
68052 | 1349 |
by (simp add: gt_ex) |
33175 | 1350 |
next |
53347 | 1351 |
case False |
68052 | 1352 |
obtain y where "y \<in> S" and y: "\<And>x. x \<in> S \<Longrightarrow> dist z y \<le> dist z x" |
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1353 |
by (metis distance_attains_inf[OF assms(2) False]) |
53347 | 1354 |
show ?thesis |
68052 | 1355 |
proof (intro exI conjI ballI) |
1356 |
show "(y - z) \<bullet> z < inner (y - z) z + (norm (y - z))\<^sup>2 / 2" |
|
1357 |
using \<open>y\<in>S\<close> \<open>z\<notin>S\<close> by auto |
|
1358 |
next |
|
53347 | 1359 |
fix x |
68052 | 1360 |
assume "x \<in> S" |
1361 |
have "False" if *: "0 < inner (z - y) (x - y)" |
|
53347 | 1362 |
proof - |
68052 | 1363 |
obtain u where "u > 0" "u \<le> 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" |
1364 |
using * closer_point_lemma by blast |
|
1365 |
then show False using y[of "y + u *\<^sub>R (x - y)"] convexD_alt [OF \<open>convex S\<close>] |
|
1366 |
using \<open>x\<in>S\<close> \<open>y\<in>S\<close> by (auto simp: dist_commute algebra_simps) |
|
53347 | 1367 |
qed |
1368 |
moreover have "0 < (norm (y - z))\<^sup>2" |
|
68052 | 1369 |
using \<open>y\<in>S\<close> \<open>z\<notin>S\<close> by auto |
53347 | 1370 |
then have "0 < inner (y - z) (y - z)" |
1371 |
unfolding power2_norm_eq_inner by simp |
|
68052 | 1372 |
ultimately show "(y - z) \<bullet> z + (norm (y - z))\<^sup>2 / 2 < (y - z) \<bullet> x" |
1373 |
by (force simp: field_simps power2_norm_eq_inner inner_commute inner_diff) |
|
1374 |
qed |
|
33175 | 1375 |
qed |
1376 |
||
1377 |
lemma separating_hyperplane_closed_0: |
|
68052 | 1378 |
assumes "convex (S::('a::euclidean_space) set)" |
1379 |
and "closed S" |
|
1380 |
and "0 \<notin> S" |
|
1381 |
shows "\<exists>a b. a \<noteq> 0 \<and> 0 < b \<and> (\<forall>x\<in>S. inner a x > b)" |
|
1382 |
proof (cases "S = {}") |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
1383 |
case True |
68052 | 1384 |
have "(SOME i. i\<in>Basis) \<noteq> (0::'a)" |
1385 |
by (metis Basis_zero SOME_Basis) |
|
53347 | 1386 |
then show ?thesis |
68052 | 1387 |
using True zero_less_one by blast |
53347 | 1388 |
next |
1389 |
case False |
|
1390 |
then show ?thesis |
|
1391 |
using False using separating_hyperplane_closed_point[OF assms] |
|
68052 | 1392 |
by (metis all_not_in_conv inner_zero_left inner_zero_right less_eq_real_def not_le) |
53347 | 1393 |
qed |
1394 |
||
33175 | 1395 |
|
70136 | 1396 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Now set-to-set for closed/compact sets\<close> |
33175 | 1397 |
|
1398 |
lemma separating_hyperplane_closed_compact: |
|
65038
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1399 |
fixes S :: "'a::euclidean_space set" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1400 |
assumes "convex S" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1401 |
and "closed S" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1402 |
and "convex T" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1403 |
and "compact T" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1404 |
and "T \<noteq> {}" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1405 |
and "S \<inter> T = {}" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1406 |
shows "\<exists>a b. (\<forall>x\<in>S. inner a x < b) \<and> (\<forall>x\<in>T. inner a x > b)" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1407 |
proof (cases "S = {}") |
33175 | 1408 |
case True |
65038
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1409 |
obtain b where b: "b > 0" "\<forall>x\<in>T. norm x \<le> b" |
53347 | 1410 |
using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto |
1411 |
obtain z :: 'a where z: "norm z = b + 1" |
|
1412 |
using vector_choose_size[of "b + 1"] and b(1) by auto |
|
65038
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1413 |
then have "z \<notin> T" using b(2)[THEN bspec[where x=z]] by auto |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1414 |
then obtain a b where ab: "inner a z < b" "\<forall>x\<in>T. b < inner a x" |
53347 | 1415 |
using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z] |
1416 |
by auto |
|
1417 |
then show ?thesis |
|
1418 |
using True by auto |
|
33175 | 1419 |
next |
53347 | 1420 |
case False |
65038
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1421 |
then obtain y where "y \<in> S" by auto |
72356 | 1422 |
obtain a b where "0 < b" and \<section>: "\<And>x. x \<in> (\<Union>x\<in> S. \<Union>y \<in> T. {x - y}) \<Longrightarrow> b < inner a x" |
33175 | 1423 |
using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0] |
72356 | 1424 |
using closed_compact_differences assms by fastforce |
1425 |
have ab: "b + inner a y < inner a x" if "x\<in>S" "y\<in>T" for x y |
|
1426 |
using \<section> [of "x-y"] that by (auto simp add: inner_diff_right less_diff_eq) |
|
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69064
diff
changeset
|
1427 |
define k where "k = (SUP x\<in>T. a \<bullet> x)" |
72356 | 1428 |
have "k + b / 2 < a \<bullet> x" if "x \<in> S" for x |
53347 | 1429 |
proof - |
72356 | 1430 |
have "k \<le> inner a x - b" |
1431 |
unfolding k_def |
|
1432 |
using \<open>T \<noteq> {}\<close> ab that by (fastforce intro: cSUP_least) |
|
1433 |
then show ?thesis |
|
1434 |
using \<open>0 < b\<close> by auto |
|
1435 |
qed |
|
1436 |
moreover |
|
1437 |
have "- (k + b / 2) < - a \<bullet> x" if "x \<in> T" for x |
|
1438 |
proof - |
|
1439 |
have "inner a x - b / 2 < k" |
|
53347 | 1440 |
unfolding k_def |
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
1441 |
proof (subst less_cSUP_iff) |
65038
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1442 |
show "T \<noteq> {}" by fact |
67399 | 1443 |
show "bdd_above ((\<bullet>) a ` T)" |
65038
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1444 |
using ab[rule_format, of y] \<open>y \<in> S\<close> |
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
1445 |
by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le) |
72356 | 1446 |
show "\<exists>y\<in>T. a \<bullet> x - b / 2 < a \<bullet> y" |
1447 |
using \<open>0 < b\<close> that by force |
|
1448 |
qed |
|
1449 |
then show ?thesis |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
1450 |
by auto |
33175 | 1451 |
qed |
72356 | 1452 |
ultimately show ?thesis |
1453 |
by (metis inner_minus_left neg_less_iff_less) |
|
33175 | 1454 |
qed |
1455 |
||
1456 |
lemma separating_hyperplane_compact_closed: |
|
65038
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1457 |
fixes S :: "'a::euclidean_space set" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1458 |
assumes "convex S" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1459 |
and "compact S" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1460 |
and "S \<noteq> {}" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1461 |
and "convex T" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1462 |
and "closed T" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1463 |
and "S \<inter> T = {}" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1464 |
shows "\<exists>a b. (\<forall>x\<in>S. inner a x < b) \<and> (\<forall>x\<in>T. inner a x > b)" |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1465 |
proof - |
9391ea7daa17
new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
1466 |
obtain a b where "(\<forall>x\<in>T. inner a x < b) \<and> (\<forall>x\<in>S. b < inner a x)" |
72356 | 1467 |
by (metis disjoint_iff_not_equal separating_hyperplane_closed_compact assms) |
53347 | 1468 |
then show ?thesis |
72356 | 1469 |
by (metis inner_minus_left neg_less_iff_less) |
53347 | 1470 |
qed |
1471 |
||
33175 | 1472 |
|
70136 | 1473 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>General case without assuming closure and getting non-strict separation\<close> |
33175 | 1474 |
|
1475 |
lemma separating_hyperplane_set_0: |
|
68031 | 1476 |
assumes "convex S" "(0::'a::euclidean_space) \<notin> S" |
1477 |
shows "\<exists>a. a \<noteq> 0 \<and> (\<forall>x\<in>S. 0 \<le> inner a x)" |
|
53347 | 1478 |
proof - |
1479 |
let ?k = "\<lambda>c. {x::'a. 0 \<le> inner c x}" |
|
68031 | 1480 |
have *: "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}" if as: "f \<subseteq> ?k ` S" "finite f" for f |
53347 | 1481 |
proof - |
68031 | 1482 |
obtain c where c: "f = ?k ` c" "c \<subseteq> S" "finite c" |
53347 | 1483 |
using finite_subset_image[OF as(2,1)] by auto |
1484 |
then obtain a b where ab: "a \<noteq> 0" "0 < b" "\<forall>x\<in>convex hull c. b < inner a x" |
|
33175 | 1485 |
using separating_hyperplane_closed_0[OF convex_convex_hull, of c] |
1486 |
using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2) |
|
53347 | 1487 |
using subset_hull[of convex, OF assms(1), symmetric, of c] |
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:
61531
diff
changeset
|
1488 |
by force |
72356 | 1489 |
have "norm (a /\<^sub>R norm a) = 1" |
1490 |
by (simp add: ab(1)) |
|
1491 |
moreover have "(\<forall>y\<in>c. 0 \<le> y \<bullet> (a /\<^sub>R norm a))" |
|
1492 |
using hull_subset[of c convex] ab by (force simp: inner_commute) |
|
1493 |
ultimately have "\<exists>x. norm x = 1 \<and> (\<forall>y\<in>c. 0 \<le> inner y x)" |
|
1494 |
by blast |
|
53347 | 1495 |
then show "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}" |
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1496 |
unfolding c(1) frontier_cball sphere_def dist_norm by auto |
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1497 |
qed |
68031 | 1498 |
have "frontier (cball 0 1) \<inter> (\<Inter>(?k ` S)) \<noteq> {}" |
72356 | 1499 |
by (rule compact_imp_fip) (use * closed_halfspace_ge in auto) |
68031 | 1500 |
then obtain x where "norm x = 1" "\<forall>y\<in>S. x\<in>?k y" |
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1501 |
unfolding frontier_cball dist_norm sphere_def by auto |
53347 | 1502 |
then show ?thesis |
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1503 |
by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one) |
53347 | 1504 |
qed |
33175 | 1505 |
|
1506 |
lemma separating_hyperplane_sets: |
|
68031 | 1507 |
fixes S T :: "'a::euclidean_space set" |
1508 |
assumes "convex S" |
|
1509 |
and "convex T" |
|
1510 |
and "S \<noteq> {}" |
|
1511 |
and "T \<noteq> {}" |
|
1512 |
and "S \<inter> T = {}" |
|
1513 |
shows "\<exists>a b. a \<noteq> 0 \<and> (\<forall>x\<in>S. inner a x \<le> b) \<and> (\<forall>x\<in>T. inner a x \<ge> b)" |
|
53347 | 1514 |
proof - |
1515 |
from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] |
|
68031 | 1516 |
obtain a where "a \<noteq> 0" "\<forall>x\<in>{x - y |x y. x \<in> T \<and> y \<in> S}. 0 \<le> inner a x" |
1517 |
using assms(3-5) by force |
|
1518 |
then have *: "\<And>x y. x \<in> T \<Longrightarrow> y \<in> S \<Longrightarrow> inner a y \<le> inner a x" |
|
1519 |
by (force simp: inner_diff) |
|
1520 |
then have bdd: "bdd_above (((\<bullet>) a)`S)" |
|
1521 |
using \<open>T \<noteq> {}\<close> by (auto intro: bdd_aboveI2[OF *]) |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
1522 |
show ?thesis |
60420 | 1523 |
using \<open>a\<noteq>0\<close> |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69064
diff
changeset
|
1524 |
by (intro exI[of _ a] exI[of _ "SUP x\<in>S. a \<bullet> x"]) |
68031 | 1525 |
(auto intro!: cSUP_upper bdd cSUP_least \<open>a \<noteq> 0\<close> \<open>S \<noteq> {}\<close> *) |
60420 | 1526 |
qed |
1527 |
||
1528 |
||
70136 | 1529 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>More convexity generalities\<close> |
33175 | 1530 |
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
1531 |
lemma convex_closure [intro,simp]: |
68031 | 1532 |
fixes S :: "'a::real_normed_vector set" |
1533 |
assumes "convex S" |
|
1534 |
shows "convex (closure S)" |
|
53676 | 1535 |
apply (rule convexI) |
72356 | 1536 |
unfolding closure_sequential |
1537 |
apply (elim exE) |
|
1538 |
subgoal for x y u v f g |
|
1539 |
by (rule_tac x="\<lambda>n. u *\<^sub>R f n + v *\<^sub>R g n" in exI) (force intro: tendsto_intros dest: convexD [OF assms]) |
|
53347 | 1540 |
done |
33175 | 1541 |
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
1542 |
lemma convex_interior [intro,simp]: |
68031 | 1543 |
fixes S :: "'a::real_normed_vector set" |
1544 |
assumes "convex S" |
|
1545 |
shows "convex (interior S)" |
|
53347 | 1546 |
unfolding convex_alt Ball_def mem_interior |
68052 | 1547 |
proof clarify |
53347 | 1548 |
fix x y u |
1549 |
assume u: "0 \<le> u" "u \<le> (1::real)" |
|
1550 |
fix e d |
|
68031 | 1551 |
assume ed: "ball x e \<subseteq> S" "ball y d \<subseteq> S" "0<d" "0<e" |
1552 |
show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> S" |
|
68052 | 1553 |
proof (intro exI conjI subsetI) |
53347 | 1554 |
fix z |
72356 | 1555 |
assume z: "z \<in> ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)" |
1556 |
have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> S" |
|
1557 |
proof (rule_tac assms[unfolded convex_alt, rule_format]) |
|
1558 |
show "z - u *\<^sub>R (y - x) \<in> S" "z + (1 - u) *\<^sub>R (y - x) \<in> S" |
|
1559 |
using ed z u by (auto simp add: algebra_simps dist_norm) |
|
1560 |
qed (use u in auto) |
|
68031 | 1561 |
then show "z \<in> S" |
1562 |
using u by (auto simp: algebra_simps) |
|
72356 | 1563 |
qed(use u ed in auto) |
53347 | 1564 |
qed |
33175 | 1565 |
|
68031 | 1566 |
lemma convex_hull_eq_empty[simp]: "convex hull S = {} \<longleftrightarrow> S = {}" |
1567 |
using hull_subset[of S convex] convex_hull_empty by auto |
|
33175 | 1568 |
|
53347 | 1569 |
|
70136 | 1570 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Convex set as intersection of halfspaces\<close> |
33175 | 1571 |
|
1572 |
lemma convex_halfspace_intersection: |
|
72356 | 1573 |
fixes S :: "('a::euclidean_space) set" |
1574 |
assumes "closed S" "convex S" |
|
1575 |
shows "S = \<Inter>{h. S \<subseteq> h \<and> (\<exists>a b. h = {x. inner a x \<le> b})}" |
|
53347 | 1576 |
proof - |
72356 | 1577 |
{ fix z |
1578 |
assume "\<forall>T. S \<subseteq> T \<and> (\<exists>a b. T = {x. inner a x \<le> b}) \<longrightarrow> z \<in> T" "z \<notin> S" |
|
1579 |
then have \<section>: "\<And>a b. S \<subseteq> {x. inner a x \<le> b} \<Longrightarrow> z \<in> {x. inner a x \<le> b}" |
|
1580 |
by blast |
|
1581 |
obtain a b where "inner a z < b" "(\<forall>x\<in>S. inner a x > b)" |
|
1582 |
using \<open>z \<notin> S\<close> assms separating_hyperplane_closed_point by blast |
|
1583 |
then have False |
|
1584 |
using \<section> [of "-a" "-b"] by fastforce |
|
1585 |
} |
|
1586 |
then show ?thesis |
|
1587 |
by force |
|
1588 |
qed |
|
33175 | 1589 |
|
53347 | 1590 |
|
70136 | 1591 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Convexity of general and special intervals\<close> |
33175 | 1592 |
|
1593 |
lemma is_interval_convex: |
|
68052 | 1594 |
fixes S :: "'a::euclidean_space set" |
1595 |
assumes "is_interval S" |
|
1596 |
shows "convex S" |
|
37732
6432bf0d7191
generalize type of is_interval to class euclidean_space
huffman
parents:
37673
diff
changeset
|
1597 |
proof (rule convexI) |
53348 | 1598 |
fix x y and u v :: real |
72356 | 1599 |
assume "x \<in> S" "y \<in> S" and uv: "0 \<le> u" "0 \<le> v" "u + v = 1" |
53348 | 1600 |
then have *: "u = 1 - v" "1 - v \<ge> 0" and **: "v = 1 - u" "1 - u \<ge> 0" |
1601 |
by auto |
|
1602 |
{ |
|
1603 |
fix a b |
|
1604 |
assume "\<not> b \<le> u * a + v * b" |
|
1605 |
then have "u * a < (1 - v) * b" |
|
72356 | 1606 |
unfolding not_le using \<open>0 \<le> v\<close>by (auto simp: field_simps) |
53348 | 1607 |
then have "a < b" |
72356 | 1608 |
using "*"(1) less_eq_real_def uv(1) by auto |
53348 | 1609 |
then have "a \<le> u * a + v * b" |
72356 | 1610 |
unfolding * using \<open>0 \<le> v\<close> |
68031 | 1611 |
by (auto simp: field_simps intro!:mult_right_mono) |
53348 | 1612 |
} |
1613 |
moreover |
|
1614 |
{ |
|
1615 |
fix a b |
|
1616 |
assume "\<not> u * a + v * b \<le> a" |
|
1617 |
then have "v * b > (1 - u) * a" |
|
72356 | 1618 |
unfolding not_le using \<open>0 \<le> v\<close> by (auto simp: field_simps) |
53348 | 1619 |
then have "a < b" |
72356 | 1620 |
unfolding * using \<open>0 \<le> v\<close> |
1621 |
by (rule_tac mult_left_less_imp_less) (auto simp: field_simps) |
|
53348 | 1622 |
then have "u * a + v * b \<le> b" |
1623 |
unfolding ** |
|
72356 | 1624 |
using **(2) \<open>0 \<le> u\<close> by (auto simp: algebra_simps mult_right_mono) |
53348 | 1625 |
} |
68052 | 1626 |
ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> S" |
72356 | 1627 |
using DIM_positive[where 'a='a] |
1628 |
by (intro mem_is_intervalI [OF assms \<open>x \<in> S\<close> \<open>y \<in> S\<close>]) (auto simp: inner_simps) |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
1629 |
qed |
33175 | 1630 |
|
1631 |
lemma is_interval_connected: |
|
68052 | 1632 |
fixes S :: "'a::euclidean_space set" |
1633 |
shows "is_interval S \<Longrightarrow> connected S" |
|
33175 | 1634 |
using is_interval_convex convex_connected by auto |
1635 |
||
62618
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1636 |
lemma convex_box [simp]: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))" |
72356 | 1637 |
by (auto simp add: is_interval_convex) |
33175 | 1638 |
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1639 |
text\<open>A non-singleton connected set is perfect (i.e. has no isolated points). \<close> |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1640 |
lemma connected_imp_perfect: |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1641 |
fixes a :: "'a::metric_space" |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1642 |
assumes "connected S" "a \<in> S" and S: "\<And>x. S \<noteq> {x}" |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1643 |
shows "a islimpt S" |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1644 |
proof - |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1645 |
have False if "a \<in> T" "open T" "\<And>y. \<lbrakk>y \<in> S; y \<in> T\<rbrakk> \<Longrightarrow> y = a" for T |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1646 |
proof - |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1647 |
obtain e where "e > 0" and e: "cball a e \<subseteq> T" |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1648 |
using \<open>open T\<close> \<open>a \<in> T\<close> by (auto simp: open_contains_cball) |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69710
diff
changeset
|
1649 |
have "openin (top_of_set S) {a}" |
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1650 |
unfolding openin_open using that \<open>a \<in> S\<close> by blast |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69710
diff
changeset
|
1651 |
moreover have "closedin (top_of_set S) {a}" |
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1652 |
by (simp add: assms) |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1653 |
ultimately show "False" |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1654 |
using \<open>connected S\<close> connected_clopen S by blast |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1655 |
qed |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1656 |
then show ?thesis |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1657 |
unfolding islimpt_def by blast |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1658 |
qed |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1659 |
|
73795 | 1660 |
lemma islimpt_Ioc [simp]: |
1661 |
fixes a :: real |
|
1662 |
assumes "a<b" |
|
1663 |
shows "x islimpt {a<..b} \<longleftrightarrow> x \<in> {a..b}" (is "?lhs = ?rhs") |
|
1664 |
proof |
|
1665 |
show "?lhs \<Longrightarrow> ?rhs" |
|
1666 |
by (metis assms closed_atLeastAtMost closed_limpt closure_greaterThanAtMost closure_subset islimpt_subset) |
|
1667 |
next |
|
1668 |
assume ?rhs |
|
1669 |
then have "x \<in> closure {a<..<b}" |
|
1670 |
using assms closure_greaterThanLessThan by blast |
|
1671 |
then show ?lhs |
|
1672 |
by (metis (no_types) Diff_empty Diff_insert0 interior_lessThanAtMost interior_limit_point interior_subset islimpt_in_closure islimpt_subset) |
|
1673 |
qed |
|
1674 |
||
1675 |
lemma islimpt_Ico [simp]: |
|
1676 |
fixes a :: real |
|
1677 |
assumes "a<b" shows "x islimpt {a..<b} \<longleftrightarrow> x \<in> {a..b}" |
|
1678 |
by (metis assms closure_atLeastLessThan closure_greaterThanAtMost islimpt_Ioc limpt_of_closure) |
|
1679 |
||
1680 |
lemma islimpt_Icc [simp]: |
|
1681 |
fixes a :: real |
|
1682 |
assumes "a<b" shows "x islimpt {a..b} \<longleftrightarrow> x \<in> {a..b}" |
|
1683 |
by (metis assms closure_atLeastLessThan islimpt_Ico limpt_of_closure) |
|
1684 |
||
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1685 |
lemma connected_imp_perfect_aff_dim: |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1686 |
"\<lbrakk>connected S; aff_dim S \<noteq> 0; a \<in> S\<rbrakk> \<Longrightarrow> a islimpt S" |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1687 |
using aff_dim_sing connected_imp_perfect by blast |
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1688 |
|
70136 | 1689 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>On \<open>real\<close>, \<open>is_interval\<close>, \<open>convex\<close> and \<open>connected\<close> are all equivalent\<close> |
33175 | 1690 |
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1691 |
lemma mem_is_interval_1_I: |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1692 |
fixes a b c::real |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1693 |
assumes "is_interval S" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1694 |
assumes "a \<in> S" "c \<in> S" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1695 |
assumes "a \<le> b" "b \<le> c" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1696 |
shows "b \<in> S" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1697 |
using assms is_interval_1 by blast |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1698 |
|
54465 | 1699 |
lemma is_interval_connected_1: |
72356 | 1700 |
fixes S :: "real set" |
1701 |
shows "is_interval S \<longleftrightarrow> connected S" |
|
1702 |
by (meson connected_iff_interval is_interval_1) |
|
33175 | 1703 |
|
1704 |
lemma is_interval_convex_1: |
|
72356 | 1705 |
fixes S :: "real set" |
1706 |
shows "is_interval S \<longleftrightarrow> convex S" |
|
54465 | 1707 |
by (metis is_interval_convex convex_connected is_interval_connected_1) |
33175 | 1708 |
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
1709 |
lemma connected_compact_interval_1: |
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
1710 |
"connected S \<and> compact S \<longleftrightarrow> (\<exists>a b. S = {a..b::real})" |
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
1711 |
by (auto simp: is_interval_connected_1 [symmetric] is_interval_compact) |
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
1712 |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1713 |
lemma connected_convex_1: |
72356 | 1714 |
fixes S :: "real set" |
1715 |
shows "connected S \<longleftrightarrow> convex S" |
|
54465 | 1716 |
by (metis is_interval_convex convex_connected is_interval_connected_1) |
53348 | 1717 |
|
77935
7f240b0dabd9
More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents:
73795
diff
changeset
|
1718 |
lemma connected_space_iff_is_interval_1 [iff]: |
7f240b0dabd9
More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents:
73795
diff
changeset
|
1719 |
fixes S :: "real set" |
7f240b0dabd9
More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents:
73795
diff
changeset
|
1720 |
shows "connected_space (top_of_set S) \<longleftrightarrow> is_interval S" |
7f240b0dabd9
More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents:
73795
diff
changeset
|
1721 |
using connectedin_topspace is_interval_connected_1 by force |
7f240b0dabd9
More new theorems, and a necessary correction
paulson <lp15@cam.ac.uk>
parents:
73795
diff
changeset
|
1722 |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1723 |
lemma connected_convex_1_gen: |
72356 | 1724 |
fixes S :: "'a :: euclidean_space set" |
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1725 |
assumes "DIM('a) = 1" |
72356 | 1726 |
shows "connected S \<longleftrightarrow> convex S" |
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1727 |
proof - |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1728 |
obtain f:: "'a \<Rightarrow> real" where linf: "linear f" and "inj f" |
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67982
diff
changeset
|
1729 |
using subspace_isomorphism[OF subspace_UNIV subspace_UNIV, |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67982
diff
changeset
|
1730 |
where 'a='a and 'b=real] |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67982
diff
changeset
|
1731 |
unfolding Euclidean_Space.dim_UNIV |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67982
diff
changeset
|
1732 |
by (auto simp: assms) |
72356 | 1733 |
then have "f -` (f ` S) = S" |
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1734 |
by (simp add: inj_vimage_image_eq) |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1735 |
then show ?thesis |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1736 |
by (metis connected_convex_1 convex_linear_vimage linf convex_connected connected_linear_image) |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1737 |
qed |
53348 | 1738 |
|
70097
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1739 |
lemma [simp]: |
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1740 |
fixes r s::real |
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1741 |
shows is_interval_io: "is_interval {..<r}" |
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1742 |
and is_interval_oi: "is_interval {r<..}" |
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1743 |
and is_interval_oo: "is_interval {r<..<s}" |
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1744 |
and is_interval_oc: "is_interval {r<..s}" |
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1745 |
and is_interval_co: "is_interval {r..<s}" |
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
69922
diff
changeset
|
1746 |
by (simp_all add: is_interval_convex_1) |
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1747 |
|
70136 | 1748 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Another intermediate value theorem formulation\<close> |
33175 | 1749 |
|
53348 | 1750 |
lemma ivt_increasing_component_on_1: |
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:
61531
diff
changeset
|
1751 |
fixes f :: "real \<Rightarrow> 'a::euclidean_space" |
53348 | 1752 |
assumes "a \<le> b" |
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1753 |
and "continuous_on {a..b} f" |
53348 | 1754 |
and "(f a)\<bullet>k \<le> y" "y \<le> (f b)\<bullet>k" |
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1755 |
shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y" |
56188 | 1756 |
proof - |
1757 |
have "f a \<in> f ` cbox a b" "f b \<in> f ` cbox a b" |
|
72356 | 1758 |
using \<open>a \<le> b\<close> by auto |
53348 | 1759 |
then show ?thesis |
56188 | 1760 |
using connected_ivt_component[of "f ` cbox a b" "f a" "f b" k 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:
66793
diff
changeset
|
1761 |
by (simp add: connected_continuous_image assms) |
53348 | 1762 |
qed |
1763 |
||
1764 |
lemma ivt_increasing_component_1: |
|
1765 |
fixes f :: "real \<Rightarrow> 'a::euclidean_space" |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1766 |
shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a..b}. continuous (at x) f \<Longrightarrow> |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1767 |
f a\<bullet>k \<le> y \<Longrightarrow> y \<le> f b\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>k = y" |
68031 | 1768 |
by (rule ivt_increasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) |
53348 | 1769 |
|
1770 |
lemma ivt_decreasing_component_on_1: |
|
1771 |
fixes f :: "real \<Rightarrow> 'a::euclidean_space" |
|
1772 |
assumes "a \<le> b" |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1773 |
and "continuous_on {a..b} f" |
53348 | 1774 |
and "(f b)\<bullet>k \<le> y" |
1775 |
and "y \<le> (f a)\<bullet>k" |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1776 |
shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y" |
72356 | 1777 |
using ivt_increasing_component_on_1[of a b "\<lambda>x. - f x" k "- y"] neg_equal_iff_equal |
1778 |
using assms continuous_on_minus by force |
|
53348 | 1779 |
|
1780 |
lemma ivt_decreasing_component_1: |
|
1781 |
fixes f :: "real \<Rightarrow> 'a::euclidean_space" |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1782 |
shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a..b}. continuous (at x) f \<Longrightarrow> |
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1783 |
f b\<bullet>k \<le> y \<Longrightarrow> y \<le> f a\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>k = y" |
53348 | 1784 |
by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) |
1785 |
||
33175 | 1786 |
|
70136 | 1787 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>A bound within an interval\<close> |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1788 |
|
56188 | 1789 |
lemma convex_hull_eq_real_cbox: |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1790 |
fixes x y :: real assumes "x \<le> y" |
56188 | 1791 |
shows "convex hull {x, y} = cbox x y" |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1792 |
proof (rule hull_unique) |
60420 | 1793 |
show "{x, y} \<subseteq> cbox x y" using \<open>x \<le> y\<close> by auto |
56188 | 1794 |
show "convex (cbox x y)" |
1795 |
by (rule convex_box) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1796 |
next |
68058 | 1797 |
fix S assume "{x, y} \<subseteq> S" and "convex S" |
1798 |
then show "cbox x y \<subseteq> S" |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1799 |
unfolding is_interval_convex_1 [symmetric] is_interval_def Basis_real_def |
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1800 |
by - (clarify, simp (no_asm_use), fast) |
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1801 |
qed |
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
1802 |
|
33175 | 1803 |
lemma unit_interval_convex_hull: |
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1804 |
"cbox (0::'a::euclidean_space) One = convex hull {x. \<forall>i\<in>Basis. (x\<bullet>i = 0) \<or> (x\<bullet>i = 1)}" |
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
1805 |
(is "?int = convex hull ?points") |
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
1806 |
proof - |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
1807 |
have One[simp]: "\<And>i. i \<in> Basis \<Longrightarrow> One \<bullet> i = 1" |
64267 | 1808 |
by (simp add: inner_sum_left sum.If_cases inner_Basis) |
56188 | 1809 |
have "?int = {x. \<forall>i\<in>Basis. x \<bullet> i \<in> cbox 0 1}" |
1810 |
by (auto simp: cbox_def) |
|
1811 |
also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` cbox 0 1)" |
|
64267 | 1812 |
by (simp only: box_eq_set_sum_Basis) |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1813 |
also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` (convex hull {0, 1}))" |
56188 | 1814 |
by (simp only: convex_hull_eq_real_cbox zero_le_one) |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1815 |
also have "\<dots> = (\<Sum>i\<in>Basis. convex hull ((\<lambda>x. x *\<^sub>R i) ` {0, 1}))" |
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67982
diff
changeset
|
1816 |
by (simp add: convex_hull_linear_image) |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1817 |
also have "\<dots> = convex hull (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` {0, 1})" |
64267 | 1818 |
by (simp only: convex_hull_set_sum) |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1819 |
also have "\<dots> = convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}}" |
64267 | 1820 |
by (simp only: box_eq_set_sum_Basis) |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1821 |
also have "convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}} = convex hull ?points" |
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1822 |
by simp |
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
1823 |
finally show ?thesis . |
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
1824 |
qed |
33175 | 1825 |
|
60420 | 1826 |
text \<open>And this is a finite set of vertices.\<close> |
33175 | 1827 |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
1828 |
lemma unit_cube_convex_hull: |
68058 | 1829 |
obtains S :: "'a::euclidean_space set" |
1830 |
where "finite S" and "cbox 0 (\<Sum>Basis) = convex hull S" |
|
1831 |
proof |
|
1832 |
show "finite {x::'a. \<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1}" |
|
1833 |
proof (rule finite_subset, clarify) |
|
1834 |
show "finite ((\<lambda>S. \<Sum>i\<in>Basis. (if i \<in> S then 1 else 0) *\<^sub>R i) ` Pow Basis)" |
|
1835 |
using finite_Basis by blast |
|
1836 |
fix x :: 'a |
|
72356 | 1837 |
assume x: "\<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1" |
68058 | 1838 |
show "x \<in> (\<lambda>S. \<Sum>i\<in>Basis. (if i\<in>S then 1 else 0) *\<^sub>R i) ` Pow Basis" |
72356 | 1839 |
apply (rule image_eqI[where x="{i. i \<in> Basis \<and> x\<bullet>i = 1}"]) |
1840 |
using x |
|
1841 |
by (subst euclidean_eq_iff, auto) |
|
68058 | 1842 |
qed |
1843 |
show "cbox 0 One = convex hull {x. \<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1}" |
|
1844 |
using unit_interval_convex_hull by blast |
|
1845 |
qed |
|
33175 | 1846 |
|
60420 | 1847 |
text \<open>Hence any cube (could do any nonempty interval).\<close> |
33175 | 1848 |
|
1849 |
lemma cube_convex_hull: |
|
53348 | 1850 |
assumes "d > 0" |
68058 | 1851 |
obtains S :: "'a::euclidean_space set" where |
1852 |
"finite S" and "cbox (x - (\<Sum>i\<in>Basis. d*\<^sub>Ri)) (x + (\<Sum>i\<in>Basis. d*\<^sub>Ri)) = convex hull S" |
|
53348 | 1853 |
proof - |
68058 | 1854 |
let ?d = "(\<Sum>i\<in>Basis. d *\<^sub>R i)::'a" |
56188 | 1855 |
have *: "cbox (x - ?d) (x + ?d) = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\<Sum>Basis)" |
68058 | 1856 |
proof (intro set_eqI iffI) |
53348 | 1857 |
fix y |
68058 | 1858 |
assume "y \<in> cbox (x - ?d) (x + ?d)" |
56188 | 1859 |
then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> cbox 0 (\<Sum>Basis)" |
70802
160eaf566bcb
formally augmented corresponding rules for field_simps
haftmann
parents:
70136
diff
changeset
|
1860 |
using assms by (simp add: mem_box inner_simps) (simp add: field_simps) |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68607
diff
changeset
|
1861 |
with \<open>0 < d\<close> show "y \<in> (\<lambda>y. x - sum ((*\<^sub>R) d) Basis + (2 * d) *\<^sub>R y) ` cbox 0 One" |
68058 | 1862 |
by (auto intro: image_eqI[where x= "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) |
33175 | 1863 |
next |
68058 | 1864 |
fix y |
1865 |
assume "y \<in> (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 One" |
|
1866 |
then obtain z where z: "z \<in> cbox 0 One" "y = x - ?d + (2*d) *\<^sub>R z" |
|
68031 | 1867 |
by auto |
56188 | 1868 |
then show "y \<in> cbox (x - ?d) (x + ?d)" |
68058 | 1869 |
using z assms by (auto simp: mem_box inner_simps) |
53348 | 1870 |
qed |
68058 | 1871 |
obtain S where "finite S" "cbox 0 (\<Sum>Basis::'a) = convex hull S" |
53348 | 1872 |
using unit_cube_convex_hull by auto |
1873 |
then show ?thesis |
|
68058 | 1874 |
by (rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` S"]) (auto simp: convex_hull_affinity *) |
53348 | 1875 |
qed |
1876 |
||
70136 | 1877 |
subsection\<^marker>\<open>tag unimportant\<close>\<open>Representation of any interval as a finite convex hull\<close> |
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1878 |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1879 |
lemma image_stretch_interval: |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1880 |
"(\<lambda>x. \<Sum>k\<in>Basis. (m k * (x\<bullet>k)) *\<^sub>R k) ` cbox a (b::'a::euclidean_space) = |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1881 |
(if (cbox a b) = {} then {} else |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1882 |
cbox (\<Sum>k\<in>Basis. (min (m k * (a\<bullet>k)) (m k * (b\<bullet>k))) *\<^sub>R k::'a) |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1883 |
(\<Sum>k\<in>Basis. (max (m k * (a\<bullet>k)) (m k * (b\<bullet>k))) *\<^sub>R k))" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1884 |
proof cases |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1885 |
assume *: "cbox a b \<noteq> {}" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1886 |
show ?thesis |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1887 |
unfolding box_ne_empty if_not_P[OF *] |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1888 |
apply (simp add: cbox_def image_Collect set_eq_iff euclidean_eq_iff[where 'a='a] ball_conj_distrib[symmetric]) |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1889 |
apply (subst choice_Basis_iff[symmetric]) |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1890 |
proof (intro allI ball_cong refl) |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1891 |
fix x i :: 'a assume "i \<in> Basis" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1892 |
with * have a_le_b: "a \<bullet> i \<le> b \<bullet> i" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1893 |
unfolding box_ne_empty by auto |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1894 |
show "(\<exists>xa. x \<bullet> i = m i * xa \<and> a \<bullet> i \<le> xa \<and> xa \<le> b \<bullet> i) \<longleftrightarrow> |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1895 |
min (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) \<le> x \<bullet> i \<and> x \<bullet> i \<le> max (m i * (a \<bullet> i)) (m i * (b \<bullet> i))" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1896 |
proof (cases "m i = 0") |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1897 |
case True |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1898 |
with a_le_b show ?thesis by auto |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1899 |
next |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1900 |
case False |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1901 |
then have *: "\<And>a b. a = m i * b \<longleftrightarrow> b = a / m i" |
68031 | 1902 |
by (auto simp: field_simps) |
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1903 |
from False have |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1904 |
"min (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) = (if 0 < m i then m i * (a \<bullet> i) else m i * (b \<bullet> i))" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1905 |
"max (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) = (if 0 < m i then m i * (b \<bullet> i) else m i * (a \<bullet> i))" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1906 |
using a_le_b by (auto simp: min_def max_def mult_le_cancel_left) |
71938 | 1907 |
with False show ?thesis using a_le_b * |
1908 |
by (simp add: le_divide_eq divide_le_eq) (simp add: ac_simps) |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1909 |
qed |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1910 |
qed |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1911 |
qed simp |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1912 |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1913 |
lemma interval_image_stretch_interval: |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1914 |
"\<exists>u v. (\<lambda>x. \<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k) ` cbox a (b::'a::euclidean_space) = cbox u (v::'a::euclidean_space)" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1915 |
unfolding image_stretch_interval by auto |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1916 |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1917 |
lemma cbox_translation: "cbox (c + a) (c + b) = image (\<lambda>x. c + x) (cbox a b)" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1918 |
using image_affinity_cbox [of 1 c a b] |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1919 |
using box_ne_empty [of "a+c" "b+c"] box_ne_empty [of a b] |
68031 | 1920 |
by (auto simp: inner_left_distrib add.commute) |
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1921 |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1922 |
lemma cbox_image_unit_interval: |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1923 |
fixes a :: "'a::euclidean_space" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1924 |
assumes "cbox a b \<noteq> {}" |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1925 |
shows "cbox a b = |
67399 | 1926 |
(+) a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` cbox 0 One" |
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1927 |
using assms |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1928 |
apply (simp add: box_ne_empty image_stretch_interval cbox_translation [symmetric]) |
64267 | 1929 |
apply (simp add: min_def max_def algebra_simps sum_subtractf euclidean_representation) |
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1930 |
done |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1931 |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1932 |
lemma closed_interval_as_convex_hull: |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1933 |
fixes a :: "'a::euclidean_space" |
68058 | 1934 |
obtains S where "finite S" "cbox a b = convex hull S" |
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1935 |
proof (cases "cbox a b = {}") |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1936 |
case True with convex_hull_empty that show ?thesis |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1937 |
by blast |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1938 |
next |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1939 |
case False |
68058 | 1940 |
obtain S::"'a set" where "finite S" and eq: "cbox 0 One = convex hull S" |
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1941 |
by (blast intro: unit_cube_convex_hull) |
72356 | 1942 |
let ?S = "((+) a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` S)" |
1943 |
show thesis |
|
1944 |
proof |
|
1945 |
show "finite ?S" |
|
1946 |
by (simp add: \<open>finite S\<close>) |
|
1947 |
have lin: "linear (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k)" |
|
1948 |
by (rule linear_compose_sum) (auto simp: algebra_simps linearI) |
|
1949 |
show "cbox a b = convex hull ?S" |
|
1950 |
using convex_hull_linear_image [OF lin] |
|
1951 |
by (simp add: convex_hull_translation eq cbox_image_unit_interval [OF False]) |
|
1952 |
qed |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1953 |
qed |
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1954 |
|
33175 | 1955 |
|
70136 | 1956 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Bounded convex function on open set is continuous\<close> |
33175 | 1957 |
|
1958 |
lemma convex_on_bounded_continuous: |
|
68058 | 1959 |
fixes S :: "('a::real_normed_vector) set" |
1960 |
assumes "open S" |
|
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
1961 |
and f: "convex_on S f" |
68058 | 1962 |
and "\<forall>x\<in>S. \<bar>f x\<bar> \<le> b" |
1963 |
shows "continuous_on S f" |
|
72356 | 1964 |
proof - |
1965 |
have "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e" if "x \<in> S" "e > 0" for x and e :: real |
|
1966 |
proof - |
|
1967 |
define B where "B = \<bar>b\<bar> + 1" |
|
1968 |
then have B: "0 < B""\<And>x. x\<in>S \<Longrightarrow> \<bar>f x\<bar> \<le> B" |
|
1969 |
using assms(3) by auto |
|
1970 |
obtain k where "k > 0" and k: "cball x k \<subseteq> S" |
|
1971 |
using \<open>x \<in> S\<close> assms(1) open_contains_cball_eq by blast |
|
1972 |
show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e" |
|
1973 |
proof (intro exI conjI allI impI) |
|
1974 |
fix y |
|
1975 |
assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)" |
|
1976 |
show "\<bar>f y - f x\<bar> < e" |
|
1977 |
proof (cases "y = x") |
|
1978 |
case False |
|
1979 |
define t where "t = k / norm (y - x)" |
|
1980 |
have "2 < t" "0<t" |
|
1981 |
unfolding t_def using as False and \<open>k>0\<close> |
|
68031 | 1982 |
by (auto simp:field_simps) |
72356 | 1983 |
have "y \<in> S" |
1984 |
apply (rule k[THEN subsetD]) |
|
1985 |
unfolding mem_cball dist_norm |
|
1986 |
apply (rule order_trans[of _ "2 * norm (x - y)"]) |
|
1987 |
using as |
|
1988 |
by (auto simp: field_simps norm_minus_commute) |
|
1989 |
{ |
|
1990 |
define w where "w = x + t *\<^sub>R (y - x)" |
|
1991 |
have "w \<in> S" |
|
1992 |
using \<open>k>0\<close> by (auto simp: dist_norm t_def w_def k[THEN subsetD]) |
|
1993 |
have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" |
|
1994 |
by (auto simp: algebra_simps) |
|
1995 |
also have "\<dots> = 0" |
|
1996 |
using \<open>t > 0\<close> by (auto simp:field_simps) |
|
1997 |
finally have w: "(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" |
|
1998 |
unfolding w_def using False and \<open>t > 0\<close> |
|
1999 |
by (auto simp: algebra_simps) |
|
2000 |
have 2: "2 * B < e * t" |
|
2001 |
unfolding t_def using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False |
|
2002 |
by (auto simp:field_simps) |
|
2003 |
have "f y - f x \<le> (f w - f x) / t" |
|
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2004 |
using convex_onD [OF f, of "(t - 1)/t" w x] \<open>0 < t\<close> \<open>2 < t\<close> \<open>x \<in> S\<close> \<open>w \<in> S\<close> |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2005 |
by (simp add: w field_simps) |
72356 | 2006 |
also have "... < e" |
2007 |
using B(2)[OF \<open>w\<in>S\<close>] and B(2)[OF \<open>x\<in>S\<close>] 2 \<open>t > 0\<close> by (auto simp: field_simps) |
|
2008 |
finally have th1: "f y - f x < e" . |
|
2009 |
} |
|
2010 |
moreover |
|
2011 |
{ |
|
2012 |
define w where "w = x - t *\<^sub>R (y - x)" |
|
2013 |
have "w \<in> S" |
|
2014 |
using \<open>k > 0\<close> by (auto simp: dist_norm t_def w_def k[THEN subsetD]) |
|
2015 |
have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" |
|
2016 |
by (auto simp: algebra_simps) |
|
2017 |
also have "\<dots> = x" |
|
2018 |
using \<open>t > 0\<close> by (auto simp:field_simps) |
|
2019 |
finally have w: "(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" |
|
2020 |
unfolding w_def using False and \<open>t > 0\<close> |
|
2021 |
by (auto simp: algebra_simps) |
|
2022 |
have "2 * B < e * t" |
|
2023 |
unfolding t_def |
|
2024 |
using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False |
|
2025 |
by (auto simp:field_simps) |
|
2026 |
then have *: "(f w - f y) / t < e" |
|
2027 |
using B(2)[OF \<open>w\<in>S\<close>] and B(2)[OF \<open>y\<in>S\<close>] |
|
2028 |
using \<open>t > 0\<close> |
|
2029 |
by (auto simp:field_simps) |
|
2030 |
have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y" |
|
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2031 |
using convex_onD [OF f, of "t / (1+t)" w y] \<open>0 < t\<close> \<open>2 < t\<close> \<open>y \<in> S\<close> \<open>w \<in> S\<close> |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2032 |
by (simp add: w field_simps) |
72356 | 2033 |
also have "\<dots> = (f w + t * f y) / (1 + t)" |
2034 |
using \<open>t > 0\<close> by (simp add: add_divide_distrib) |
|
2035 |
also have "\<dots> < e + f y" |
|
2036 |
using \<open>t > 0\<close> * \<open>e > 0\<close> by (auto simp: field_simps) |
|
2037 |
finally have "f x - f y < e" by auto |
|
2038 |
} |
|
2039 |
ultimately show ?thesis by auto |
|
2040 |
qed (use \<open>0<e\<close> in auto) |
|
2041 |
qed (use \<open>0<e\<close> \<open>0<k\<close> \<open>0<B\<close> in \<open>auto simp: field_simps\<close>) |
|
2042 |
qed |
|
2043 |
then show ?thesis |
|
2044 |
by (metis continuous_on_iff dist_norm real_norm_def) |
|
60420 | 2045 |
qed |
2046 |
||
70136 | 2047 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Upper bound on a ball implies upper and lower bounds\<close> |
33175 | 2048 |
|
2049 |
lemma convex_bounds_lemma: |
|
36338 | 2050 |
fixes x :: "'a::real_normed_vector" |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2051 |
assumes f: "convex_on (cball x e) f" |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2052 |
and b: "\<And>y. y \<in> cball x e \<Longrightarrow> f y \<le> b" and y: "y \<in> cball x e" |
72356 | 2053 |
shows "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" |
53348 | 2054 |
proof (cases "0 \<le> e") |
2055 |
case True |
|
63040 | 2056 |
define z where "z = 2 *\<^sub>R x - y" |
53348 | 2057 |
have *: "x - (2 *\<^sub>R x - y) = y - x" |
2058 |
by (simp add: scaleR_2) |
|
2059 |
have z: "z \<in> cball x e" |
|
68031 | 2060 |
using y unfolding z_def mem_cball dist_norm * by (auto simp: norm_minus_commute) |
53348 | 2061 |
have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" |
68031 | 2062 |
unfolding z_def by (auto simp: algebra_simps) |
53348 | 2063 |
then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2064 |
using convex_onD [OF f, of "1/2" y z] b[OF y] b y z |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2065 |
by (fastforce simp add: field_simps) |
53348 | 2066 |
next |
2067 |
case False |
|
72356 | 2068 |
have "dist x y < 0" |
2069 |
using False y unfolding mem_cball not_le by (auto simp del: dist_not_less_zero) |
|
53348 | 2070 |
then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" |
2071 |
using zero_le_dist[of x y] by auto |
|
2072 |
qed |
|
2073 |
||
33175 | 2074 |
|
70136 | 2075 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Hence a convex function on an open set is continuous\<close> |
33175 | 2076 |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
2077 |
lemma real_of_nat_ge_one_iff: "1 \<le> real (n::nat) \<longleftrightarrow> 1 \<le> n" |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
2078 |
by auto |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
2079 |
|
33175 | 2080 |
lemma convex_on_continuous: |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2081 |
fixes S :: "'a::euclidean_space set" |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2082 |
assumes "open S" "convex_on S f" |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2083 |
shows "continuous_on S f" |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2084 |
unfolding continuous_on_eq_continuous_at[OF \<open>open S\<close>] |
53348 | 2085 |
proof |
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
2086 |
note dimge1 = DIM_positive[where 'a='a] |
53348 | 2087 |
fix x |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2088 |
assume "x \<in> S" |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2089 |
then obtain e where e: "cball x e \<subseteq> S" "e > 0" |
53348 | 2090 |
using assms(1) unfolding open_contains_cball by auto |
63040 | 2091 |
define d where "d = e / real DIM('a)" |
53348 | 2092 |
have "0 < d" |
60420 | 2093 |
unfolding d_def using \<open>e > 0\<close> dimge1 by auto |
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
2094 |
let ?d = "(\<Sum>i\<in>Basis. d *\<^sub>R i)::'a" |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2095 |
obtain c |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2096 |
where c: "finite c" and c1: "convex hull c \<subseteq> cball x e" and c2: "cball x d \<subseteq> convex hull c" |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2097 |
proof |
63040 | 2098 |
define c where "c = (\<Sum>i\<in>Basis. (\<lambda>a. a *\<^sub>R i) ` {x\<bullet>i - d, x\<bullet>i + d})" |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2099 |
show "finite c" |
64267 | 2100 |
unfolding c_def by (simp add: finite_set_sum) |
72356 | 2101 |
have "\<And>i. i \<in> Basis \<Longrightarrow> convex hull {x \<bullet> i - d, x \<bullet> i + d} = cbox (x \<bullet> i - d) (x \<bullet> i + d)" |
2102 |
using \<open>0 < d\<close> convex_hull_eq_real_cbox by auto |
|
2103 |
then have 1: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cbox (x \<bullet> i - d) (x \<bullet> i + d)}" |
|
2104 |
unfolding box_eq_set_sum_Basis c_def convex_hull_set_sum |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2105 |
apply (subst convex_hull_linear_image [symmetric]) |
72356 | 2106 |
by (force simp add: linear_iff scaleR_add_left)+ |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2107 |
then have 2: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cball (x \<bullet> i) d}" |
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2108 |
by (simp add: dist_norm abs_le_iff algebra_simps) |
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2109 |
show "cball x d \<subseteq> convex hull c" |
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2110 |
unfolding 2 |
68058 | 2111 |
by (clarsimp simp: dist_norm) (metis inner_commute inner_diff_right norm_bound_Basis_le) |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2112 |
have e': "e = (\<Sum>(i::'a)\<in>Basis. d)" |
71172 | 2113 |
by (simp add: d_def) |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2114 |
show "convex hull c \<subseteq> cball x e" |
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2115 |
unfolding 2 |
72356 | 2116 |
proof clarsimp |
2117 |
show "dist x y \<le> e" if "\<forall>i\<in>Basis. dist (x \<bullet> i) (y \<bullet> i) \<le> d" for y |
|
2118 |
proof - |
|
2119 |
have "\<And>i. i \<in> Basis \<Longrightarrow> 0 \<le> dist (x \<bullet> i) (y \<bullet> i)" |
|
2120 |
by simp |
|
2121 |
have "(\<Sum>i\<in>Basis. dist (x \<bullet> i) (y \<bullet> i)) \<le> e" |
|
2122 |
using e' sum_mono that by fastforce |
|
2123 |
then show ?thesis |
|
2124 |
by (metis (mono_tags) euclidean_dist_l2 order_trans [OF L2_set_le_sum] zero_le_dist) |
|
2125 |
qed |
|
2126 |
qed |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2127 |
qed |
63040 | 2128 |
define k where "k = Max (f ` c)" |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2129 |
have "convex_on (convex hull c) f" |
72356 | 2130 |
using assms(2) c1 convex_on_subset e(1) by blast |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2131 |
then have k: "\<forall>y\<in>convex hull c. f y \<le> k" |
72356 | 2132 |
using c convex_on_convex_hull_bound k_def by fastforce |
68048 | 2133 |
have "e \<le> e * real DIM('a)" |
2134 |
using e(2) real_of_nat_ge_one_iff by auto |
|
2135 |
then have "d \<le> e" |
|
70817
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents:
70802
diff
changeset
|
2136 |
by (simp add: d_def field_split_simps) |
53348 | 2137 |
then have dsube: "cball x d \<subseteq> cball x e" |
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2138 |
by (rule subset_cball) |
53348 | 2139 |
have conv: "convex_on (cball x d) f" |
68031 | 2140 |
using \<open>convex_on (convex hull c) f\<close> c2 convex_on_subset by blast |
72356 | 2141 |
then have "\<And>y. y\<in>cball x d \<Longrightarrow> \<bar>f y\<bar> \<le> k + 2 * \<bar>f x\<bar>" |
68048 | 2142 |
by (rule convex_bounds_lemma) (use c2 k in blast) |
79582
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2143 |
moreover have "convex_on (ball x d) f" |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2144 |
using conv convex_on_subset by fastforce |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2145 |
ultimately |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2146 |
have "continuous_on (ball x d) f" |
7822b55b26ce
Correct the definition of a convex function, and updated the proofs
paulson <lp15@cam.ac.uk>
parents:
78516
diff
changeset
|
2147 |
by (metis convex_on_bounded_continuous Elementary_Metric_Spaces.open_ball mem_ball_imp_mem_cball) |
53348 | 2148 |
then show "continuous (at x) f" |
2149 |
unfolding continuous_on_eq_continuous_at[OF open_ball] |
|
60420 | 2150 |
using \<open>d > 0\<close> by auto |
2151 |
qed |
|
2152 |
||
33175 | 2153 |
end |