| author | wenzelm |
| Mon, 16 Jan 2017 21:20:30 +0100 | |
| changeset 64907 | 354bfbb27fbb |
| parent 64788 | 19f3d4af7a7d |
| child 65036 | ab7e11730ad8 |
| 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 |
||
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9 |
section \<open>Convex sets, functions and related things\<close> |
| 33175 | 10 |
|
11 |
theory Convex_Euclidean_Space |
|
| 44132 | 12 |
imports |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
13 |
Topology_Euclidean_Space |
| 44132 | 14 |
"~~/src/HOL/Library/Set_Algebras" |
| 33175 | 15 |
begin |
16 |
||
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
17 |
lemma swap_continuous: (*move to Topological_Spaces?*) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
18 |
assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
19 |
shows "continuous_on (cbox (c,a) (d,b)) (\<lambda>(x, y). f y x)" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
20 |
proof - |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
21 |
have "(\<lambda>(x, y). f y x) = (\<lambda>(x, y). f x y) \<circ> prod.swap" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
22 |
by auto |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
23 |
then show ?thesis |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
24 |
apply (rule ssubst) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
25 |
apply (rule continuous_on_compose) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
26 |
apply (simp add: split_def) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
27 |
apply (rule continuous_intros | simp add: assms)+ |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
28 |
done |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
29 |
qed |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
30 |
|
| 40377 | 31 |
lemma dim_image_eq: |
| 53339 | 32 |
fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" |
| 53333 | 33 |
assumes lf: "linear f" |
34 |
and fi: "inj_on f (span S)" |
|
| 53406 | 35 |
shows "dim (f ` S) = dim (S::'n::euclidean_space set)" |
36 |
proof - |
|
37 |
obtain B where B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" |
|
| 49529 | 38 |
using basis_exists[of S] by auto |
39 |
then have "span S = span B" |
|
40 |
using span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto |
|
41 |
then have "independent (f ` B)" |
|
| 63128 | 42 |
using independent_inj_on_image[of B f] B assms by auto |
| 49529 | 43 |
moreover have "card (f ` B) = card B" |
| 53406 | 44 |
using assms card_image[of f B] subset_inj_on[of f "span S" B] B span_inc by auto |
| 53333 | 45 |
moreover have "(f ` B) \<subseteq> (f ` S)" |
| 53406 | 46 |
using B by auto |
| 53302 | 47 |
ultimately have "dim (f ` S) \<ge> dim S" |
| 53406 | 48 |
using independent_card_le_dim[of "f ` B" "f ` S"] B by auto |
| 53333 | 49 |
then show ?thesis |
50 |
using dim_image_le[of f S] assms by auto |
|
| 40377 | 51 |
qed |
52 |
||
53 |
lemma linear_injective_on_subspace_0: |
|
| 53302 | 54 |
assumes lf: "linear f" |
55 |
and "subspace S" |
|
56 |
shows "inj_on f S \<longleftrightarrow> (\<forall>x \<in> S. f x = 0 \<longrightarrow> x = 0)" |
|
| 49529 | 57 |
proof - |
| 53302 | 58 |
have "inj_on f S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. f x = f y \<longrightarrow> x = y)" |
59 |
by (simp add: inj_on_def) |
|
60 |
also have "\<dots> \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. f x - f y = 0 \<longrightarrow> x - y = 0)" |
|
61 |
by simp |
|
62 |
also have "\<dots> \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. f (x - y) = 0 \<longrightarrow> x - y = 0)" |
|
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
63 |
by (simp add: linear_diff[OF lf]) |
| 53302 | 64 |
also have "\<dots> \<longleftrightarrow> (\<forall>x \<in> S. f x = 0 \<longrightarrow> x = 0)" |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
65 |
using \<open>subspace S\<close> subspace_def[of S] subspace_diff[of S] by auto |
| 40377 | 66 |
finally show ?thesis . |
67 |
qed |
|
68 |
||
| 61952 | 69 |
lemma subspace_Inter: "\<forall>s \<in> f. subspace s \<Longrightarrow> subspace (\<Inter>f)" |
| 49531 | 70 |
unfolding subspace_def by auto |
| 40377 | 71 |
|
| 53302 | 72 |
lemma span_eq[simp]: "span s = s \<longleftrightarrow> subspace s" |
73 |
unfolding span_def by (rule hull_eq) (rule subspace_Inter) |
|
| 40377 | 74 |
|
| 49529 | 75 |
lemma substdbasis_expansion_unique: |
|
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
|
76 |
assumes d: "d \<subseteq> Basis" |
| 53302 | 77 |
shows "(\<Sum>i\<in>d. f i *\<^sub>R i) = (x::'a::euclidean_space) \<longleftrightarrow> |
78 |
(\<forall>i\<in>Basis. (i \<in> d \<longrightarrow> f i = x \<bullet> i) \<and> (i \<notin> d \<longrightarrow> x \<bullet> i = 0))" |
|
| 49529 | 79 |
proof - |
| 53339 | 80 |
have *: "\<And>x a b P. x * (if P then a else b) = (if P then x * a else x * b)" |
| 53302 | 81 |
by auto |
82 |
have **: "finite d" |
|
83 |
by (auto intro: finite_subset[OF assms]) |
|
|
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
|
84 |
have ***: "\<And>i. i \<in> Basis \<Longrightarrow> (\<Sum>i\<in>d. f i *\<^sub>R i) \<bullet> i = (\<Sum>x\<in>d. if x = i then f x else 0)" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
85 |
using d |
| 64267 | 86 |
by (auto intro!: sum.cong simp: inner_Basis inner_sum_left) |
87 |
show ?thesis |
|
88 |
unfolding euclidean_eq_iff[where 'a='a] by (auto simp: sum.delta[OF **] ***) |
|
|
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
|
89 |
qed |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
90 |
|
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
91 |
lemma independent_substdbasis: "d \<subseteq> Basis \<Longrightarrow> independent d" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
92 |
by (rule independent_mono[OF independent_Basis]) |
| 40377 | 93 |
|
| 49531 | 94 |
lemma dim_cball: |
| 53302 | 95 |
assumes "e > 0" |
| 49529 | 96 |
shows "dim (cball (0 :: 'n::euclidean_space) e) = DIM('n)"
|
97 |
proof - |
|
| 53302 | 98 |
{
|
99 |
fix x :: "'n::euclidean_space" |
|
| 63040 | 100 |
define y where "y = (e / norm x) *\<^sub>R x" |
| 53339 | 101 |
then have "y \<in> cball 0 e" |
|
62397
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents:
62381
diff
changeset
|
102 |
using assms by auto |
| 53339 | 103 |
moreover have *: "x = (norm x / e) *\<^sub>R y" |
| 53302 | 104 |
using y_def assms by simp |
105 |
moreover from * have "x = (norm x/e) *\<^sub>R y" |
|
106 |
by auto |
|
| 53339 | 107 |
ultimately have "x \<in> span (cball 0 e)" |
|
62397
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents:
62381
diff
changeset
|
108 |
using span_mul[of y "cball 0 e" "norm x/e"] span_inc[of "cball 0 e"] |
|
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents:
62381
diff
changeset
|
109 |
by (simp add: span_superset) |
| 53302 | 110 |
} |
| 53339 | 111 |
then have "span (cball 0 e) = (UNIV :: 'n::euclidean_space set)" |
| 53302 | 112 |
by auto |
| 49529 | 113 |
then show ?thesis |
114 |
using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto simp add: dim_UNIV) |
|
| 40377 | 115 |
qed |
116 |
||
117 |
lemma indep_card_eq_dim_span: |
|
| 53339 | 118 |
fixes B :: "'n::euclidean_space set" |
| 49529 | 119 |
assumes "independent B" |
| 53339 | 120 |
shows "finite B \<and> card B = dim (span B)" |
| 40377 | 121 |
using assms basis_card_eq_dim[of B "span B"] span_inc by auto |
122 |
||
| 64267 | 123 |
lemma sum_not_0: "sum f A \<noteq> 0 \<Longrightarrow> \<exists>a \<in> A. f a \<noteq> 0" |
| 49529 | 124 |
by (rule ccontr) auto |
| 40377 | 125 |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
126 |
lemma subset_translation_eq [simp]: |
|
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
127 |
fixes a :: "'a::real_vector" shows "op + a ` s \<subseteq> op + a ` t \<longleftrightarrow> s \<subseteq> t" |
|
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
128 |
by auto |
|
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
129 |
|
| 49531 | 130 |
lemma translate_inj_on: |
| 53339 | 131 |
fixes A :: "'a::ab_group_add set" |
132 |
shows "inj_on (\<lambda>x. a + x) A" |
|
| 49529 | 133 |
unfolding inj_on_def by auto |
| 40377 | 134 |
|
135 |
lemma translation_assoc: |
|
136 |
fixes a b :: "'a::ab_group_add" |
|
| 53339 | 137 |
shows "(\<lambda>x. b + x) ` ((\<lambda>x. a + x) ` S) = (\<lambda>x. (a + b) + x) ` S" |
| 49529 | 138 |
by auto |
| 40377 | 139 |
|
140 |
lemma translation_invert: |
|
141 |
fixes a :: "'a::ab_group_add" |
|
| 53339 | 142 |
assumes "(\<lambda>x. a + x) ` A = (\<lambda>x. a + x) ` B" |
| 49529 | 143 |
shows "A = B" |
144 |
proof - |
|
| 53339 | 145 |
have "(\<lambda>x. -a + x) ` ((\<lambda>x. a + x) ` A) = (\<lambda>x. - a + x) ` ((\<lambda>x. a + x) ` B)" |
| 49529 | 146 |
using assms by auto |
147 |
then show ?thesis |
|
148 |
using translation_assoc[of "-a" a A] translation_assoc[of "-a" a B] by auto |
|
| 40377 | 149 |
qed |
150 |
||
151 |
lemma translation_galois: |
|
152 |
fixes a :: "'a::ab_group_add" |
|
| 53339 | 153 |
shows "T = ((\<lambda>x. a + x) ` S) \<longleftrightarrow> S = ((\<lambda>x. (- a) + x) ` T)" |
| 53333 | 154 |
using translation_assoc[of "-a" a S] |
155 |
apply auto |
|
156 |
using translation_assoc[of a "-a" T] |
|
157 |
apply auto |
|
| 49529 | 158 |
done |
| 40377 | 159 |
|
160 |
lemma translation_inverse_subset: |
|
| 53339 | 161 |
assumes "((\<lambda>x. - a + x) ` V) \<le> (S :: 'n::ab_group_add set)" |
162 |
shows "V \<le> ((\<lambda>x. a + x) ` S)" |
|
| 49529 | 163 |
proof - |
| 53333 | 164 |
{
|
165 |
fix x |
|
166 |
assume "x \<in> V" |
|
167 |
then have "x-a \<in> S" using assms by auto |
|
168 |
then have "x \<in> {a + v |v. v \<in> S}"
|
|
| 49529 | 169 |
apply auto |
170 |
apply (rule exI[of _ "x-a"]) |
|
171 |
apply simp |
|
172 |
done |
|
| 53333 | 173 |
then have "x \<in> ((\<lambda>x. a+x) ` S)" by auto |
174 |
} |
|
175 |
then show ?thesis by auto |
|
| 40377 | 176 |
qed |
177 |
||
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
178 |
subsection \<open>Convexity\<close> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
179 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
180 |
definition convex :: "'a::real_vector set \<Rightarrow> bool" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
181 |
where "convex s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u\<ge>0. \<forall>v\<ge>0. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
182 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
183 |
lemma convexI: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
184 |
assumes "\<And>x y u v. x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> 0 \<le> u \<Longrightarrow> 0 \<le> v \<Longrightarrow> u + v = 1 \<Longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
185 |
shows "convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
186 |
using assms unfolding convex_def by fast |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
187 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
188 |
lemma convexD: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
189 |
assumes "convex s" and "x \<in> s" and "y \<in> s" and "0 \<le> u" and "0 \<le> v" and "u + v = 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
190 |
shows "u *\<^sub>R x + v *\<^sub>R y \<in> s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
191 |
using assms unfolding convex_def by fast |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
192 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
193 |
lemma convex_alt: "convex s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> ((1 - u) *\<^sub>R x + u *\<^sub>R y) \<in> s)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
194 |
(is "_ \<longleftrightarrow> ?alt") |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
195 |
proof |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
196 |
show "convex s" if alt: ?alt |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
197 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
198 |
{
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
199 |
fix x y and u v :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
200 |
assume mem: "x \<in> s" "y \<in> s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
201 |
assume "0 \<le> u" "0 \<le> v" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
202 |
moreover |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
203 |
assume "u + v = 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
204 |
then have "u = 1 - v" by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
205 |
ultimately have "u *\<^sub>R x + v *\<^sub>R y \<in> s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
206 |
using alt [rule_format, OF mem] by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
207 |
} |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
208 |
then show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
209 |
unfolding convex_def by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
210 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
211 |
show ?alt if "convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
212 |
using that by (auto simp: convex_def) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
213 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
214 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
215 |
lemma convexD_alt: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
216 |
assumes "convex s" "a \<in> s" "b \<in> s" "0 \<le> u" "u \<le> 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
217 |
shows "((1 - u) *\<^sub>R a + u *\<^sub>R b) \<in> s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
218 |
using assms unfolding convex_alt by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
219 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
220 |
lemma mem_convex_alt: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
221 |
assumes "convex S" "x \<in> S" "y \<in> S" "u \<ge> 0" "v \<ge> 0" "u + v > 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
222 |
shows "((u/(u+v)) *\<^sub>R x + (v/(u+v)) *\<^sub>R y) \<in> S" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
223 |
apply (rule convexD) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
224 |
using assms |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
225 |
apply (simp_all add: zero_le_divide_iff add_divide_distrib [symmetric]) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
226 |
done |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
227 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
228 |
lemma convex_empty[intro,simp]: "convex {}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
229 |
unfolding convex_def by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
230 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
231 |
lemma convex_singleton[intro,simp]: "convex {a}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
232 |
unfolding convex_def by (auto simp: scaleR_left_distrib[symmetric]) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
233 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
234 |
lemma convex_UNIV[intro,simp]: "convex UNIV" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
235 |
unfolding convex_def by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
236 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
237 |
lemma convex_Inter: "(\<And>s. s\<in>f \<Longrightarrow> convex s) \<Longrightarrow> convex(\<Inter>f)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
238 |
unfolding convex_def by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
239 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
240 |
lemma convex_Int: "convex s \<Longrightarrow> convex t \<Longrightarrow> convex (s \<inter> t)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
241 |
unfolding convex_def by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
242 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
243 |
lemma convex_INT: "(\<And>i. i \<in> A \<Longrightarrow> convex (B i)) \<Longrightarrow> convex (\<Inter>i\<in>A. B i)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
244 |
unfolding convex_def by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
245 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
246 |
lemma convex_Times: "convex s \<Longrightarrow> convex t \<Longrightarrow> convex (s \<times> t)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
247 |
unfolding convex_def by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
248 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
249 |
lemma convex_halfspace_le: "convex {x. inner a x \<le> b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
250 |
unfolding convex_def |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
251 |
by (auto simp: inner_add intro!: convex_bound_le) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
252 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
253 |
lemma convex_halfspace_ge: "convex {x. inner a x \<ge> b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
254 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
255 |
have *: "{x. inner a x \<ge> b} = {x. inner (-a) x \<le> -b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
256 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
257 |
show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
258 |
unfolding * using convex_halfspace_le[of "-a" "-b"] by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
259 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
260 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
261 |
lemma convex_hyperplane: "convex {x. inner a x = b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
262 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
263 |
have *: "{x. inner a x = b} = {x. inner a x \<le> b} \<inter> {x. inner a x \<ge> b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
264 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
265 |
show ?thesis using convex_halfspace_le convex_halfspace_ge |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
266 |
by (auto intro!: convex_Int simp: *) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
267 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
268 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
269 |
lemma convex_halfspace_lt: "convex {x. inner a x < b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
270 |
unfolding convex_def |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
271 |
by (auto simp: convex_bound_lt inner_add) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
272 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
273 |
lemma convex_halfspace_gt: "convex {x. inner a x > b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
274 |
using convex_halfspace_lt[of "-a" "-b"] by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
275 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
276 |
lemma convex_real_interval [iff]: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
277 |
fixes a b :: "real" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
278 |
shows "convex {a..}" and "convex {..b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
279 |
and "convex {a<..}" and "convex {..<b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
280 |
and "convex {a..b}" and "convex {a<..b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
281 |
and "convex {a..<b}" and "convex {a<..<b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
282 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
283 |
have "{a..} = {x. a \<le> inner 1 x}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
284 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
285 |
then show 1: "convex {a..}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
286 |
by (simp only: convex_halfspace_ge) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
287 |
have "{..b} = {x. inner 1 x \<le> b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
288 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
289 |
then show 2: "convex {..b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
290 |
by (simp only: convex_halfspace_le) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
291 |
have "{a<..} = {x. a < inner 1 x}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
292 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
293 |
then show 3: "convex {a<..}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
294 |
by (simp only: convex_halfspace_gt) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
295 |
have "{..<b} = {x. inner 1 x < b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
296 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
297 |
then show 4: "convex {..<b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
298 |
by (simp only: convex_halfspace_lt) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
299 |
have "{a..b} = {a..} \<inter> {..b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
300 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
301 |
then show "convex {a..b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
302 |
by (simp only: convex_Int 1 2) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
303 |
have "{a<..b} = {a<..} \<inter> {..b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
304 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
305 |
then show "convex {a<..b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
306 |
by (simp only: convex_Int 3 2) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
307 |
have "{a..<b} = {a..} \<inter> {..<b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
308 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
309 |
then show "convex {a..<b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
310 |
by (simp only: convex_Int 1 4) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
311 |
have "{a<..<b} = {a<..} \<inter> {..<b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
312 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
313 |
then show "convex {a<..<b}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
314 |
by (simp only: convex_Int 3 4) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
315 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
316 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
317 |
lemma convex_Reals: "convex \<real>" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
318 |
by (simp add: convex_def scaleR_conv_of_real) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
319 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
320 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
321 |
subsection \<open>Explicit expressions for convexity in terms of arbitrary sums\<close> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
322 |
|
| 64267 | 323 |
lemma convex_sum: |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
324 |
fixes C :: "'a::real_vector set" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
325 |
assumes "finite s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
326 |
and "convex C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
327 |
and "(\<Sum> i \<in> s. a i) = 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
328 |
assumes "\<And>i. i \<in> s \<Longrightarrow> a i \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
329 |
and "\<And>i. i \<in> s \<Longrightarrow> y i \<in> C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
330 |
shows "(\<Sum> j \<in> s. a j *\<^sub>R y j) \<in> C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
331 |
using assms(1,3,4,5) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
332 |
proof (induct arbitrary: a set: finite) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
333 |
case empty |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
334 |
then show ?case by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
335 |
next |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
336 |
case (insert i s) note IH = this(3) |
| 64267 | 337 |
have "a i + sum a s = 1" |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
338 |
and "0 \<le> a i" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
339 |
and "\<forall>j\<in>s. 0 \<le> a j" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
340 |
and "y i \<in> C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
341 |
and "\<forall>j\<in>s. y j \<in> C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
342 |
using insert.hyps(1,2) insert.prems by simp_all |
| 64267 | 343 |
then have "0 \<le> sum a s" |
344 |
by (simp add: sum_nonneg) |
|
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
345 |
have "a i *\<^sub>R y i + (\<Sum>j\<in>s. a j *\<^sub>R y j) \<in> C" |
| 64267 | 346 |
proof (cases "sum a s = 0") |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
347 |
case True |
| 64267 | 348 |
with \<open>a i + sum a s = 1\<close> have "a i = 1" |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
349 |
by simp |
| 64267 | 350 |
from sum_nonneg_0 [OF \<open>finite s\<close> _ True] \<open>\<forall>j\<in>s. 0 \<le> a j\<close> have "\<forall>j\<in>s. a j = 0" |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
351 |
by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
352 |
show ?thesis using \<open>a i = 1\<close> and \<open>\<forall>j\<in>s. a j = 0\<close> and \<open>y i \<in> C\<close> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
353 |
by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
354 |
next |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
355 |
case False |
| 64267 | 356 |
with \<open>0 \<le> sum a s\<close> have "0 < sum a s" |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
357 |
by simp |
| 64267 | 358 |
then have "(\<Sum>j\<in>s. (a j / sum a s) *\<^sub>R y j) \<in> C" |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
359 |
using \<open>\<forall>j\<in>s. 0 \<le> a j\<close> and \<open>\<forall>j\<in>s. y j \<in> C\<close> |
| 64267 | 360 |
by (simp add: IH sum_divide_distrib [symmetric]) |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
361 |
from \<open>convex C\<close> and \<open>y i \<in> C\<close> and this and \<open>0 \<le> a i\<close> |
| 64267 | 362 |
and \<open>0 \<le> sum a s\<close> and \<open>a i + sum a s = 1\<close> |
363 |
have "a i *\<^sub>R y i + sum a s *\<^sub>R (\<Sum>j\<in>s. (a j / sum a s) *\<^sub>R y j) \<in> C" |
|
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
364 |
by (rule convexD) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
365 |
then show ?thesis |
| 64267 | 366 |
by (simp add: scaleR_sum_right False) |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
367 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
368 |
then show ?case using \<open>finite s\<close> and \<open>i \<notin> s\<close> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
369 |
by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
370 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
371 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
372 |
lemma convex: |
| 64267 | 373 |
"convex s \<longleftrightarrow> (\<forall>(k::nat) u x. (\<forall>i. 1\<le>i \<and> i\<le>k \<longrightarrow> 0 \<le> u i \<and> x i \<in>s) \<and> (sum u {1..k} = 1)
|
374 |
\<longrightarrow> sum (\<lambda>i. u i *\<^sub>R x i) {1..k} \<in> s)"
|
|
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
375 |
proof safe |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
376 |
fix k :: nat |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
377 |
fix u :: "nat \<Rightarrow> real" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
378 |
fix x |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
379 |
assume "convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
380 |
"\<forall>i. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> s" |
| 64267 | 381 |
"sum u {1..k} = 1"
|
382 |
with convex_sum[of "{1 .. k}" s] show "(\<Sum>j\<in>{1 .. k}. u j *\<^sub>R x j) \<in> s"
|
|
383 |
by auto |
|
384 |
next |
|
385 |
assume *: "\<forall>k u x. (\<forall> i :: nat. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> s) \<and> sum u {1..k} = 1
|
|
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
386 |
\<longrightarrow> (\<Sum>i = 1..k. u i *\<^sub>R (x i :: 'a)) \<in> s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
387 |
{
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
388 |
fix \<mu> :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
389 |
fix x y :: 'a |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
390 |
assume xy: "x \<in> s" "y \<in> s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
391 |
assume mu: "\<mu> \<ge> 0" "\<mu> \<le> 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
392 |
let ?u = "\<lambda>i. if (i :: nat) = 1 then \<mu> else 1 - \<mu>" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
393 |
let ?x = "\<lambda>i. if (i :: nat) = 1 then x else y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
394 |
have "{1 :: nat .. 2} \<inter> - {x. x = 1} = {2}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
395 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
396 |
then have card: "card ({1 :: nat .. 2} \<inter> - {x. x = 1}) = 1"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
397 |
by simp |
| 64267 | 398 |
then have "sum ?u {1 .. 2} = 1"
|
399 |
using sum.If_cases[of "{(1 :: nat) .. 2}" "\<lambda> x. x = 1" "\<lambda> x. \<mu>" "\<lambda> x. 1 - \<mu>"]
|
|
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
400 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
401 |
with *[rule_format, of "2" ?u ?x] have s: "(\<Sum>j \<in> {1..2}. ?u j *\<^sub>R ?x j) \<in> s"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
402 |
using mu xy by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
403 |
have grarr: "(\<Sum>j \<in> {Suc (Suc 0)..2}. ?u j *\<^sub>R ?x j) = (1 - \<mu>) *\<^sub>R y"
|
| 64267 | 404 |
using sum_head_Suc[of "Suc (Suc 0)" 2 "\<lambda> j. (1 - \<mu>) *\<^sub>R y"] by auto |
405 |
from sum_head_Suc[of "Suc 0" 2 "\<lambda> j. ?u j *\<^sub>R ?x j", simplified this] |
|
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
406 |
have "(\<Sum>j \<in> {1..2}. ?u j *\<^sub>R ?x j) = \<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
407 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
408 |
then have "(1 - \<mu>) *\<^sub>R y + \<mu> *\<^sub>R x \<in> s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
409 |
using s by (auto simp: add.commute) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
410 |
} |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
411 |
then show "convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
412 |
unfolding convex_alt by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
413 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
414 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
415 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
416 |
lemma convex_explicit: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
417 |
fixes s :: "'a::real_vector set" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
418 |
shows "convex s \<longleftrightarrow> |
| 64267 | 419 |
(\<forall>t u. finite t \<and> t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> sum u t = 1 \<longrightarrow> sum (\<lambda>x. u x *\<^sub>R x) t \<in> s)" |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
420 |
proof safe |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
421 |
fix t |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
422 |
fix u :: "'a \<Rightarrow> real" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
423 |
assume "convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
424 |
and "finite t" |
| 64267 | 425 |
and "t \<subseteq> s" "\<forall>x\<in>t. 0 \<le> u x" "sum u t = 1" |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
426 |
then show "(\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" |
| 64267 | 427 |
using convex_sum[of t s u "\<lambda> x. x"] by auto |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
428 |
next |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
429 |
assume *: "\<forall>t. \<forall> u. finite t \<and> t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> |
| 64267 | 430 |
sum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
431 |
show "convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
432 |
unfolding convex_alt |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
433 |
proof safe |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
434 |
fix x y |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
435 |
fix \<mu> :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
436 |
assume **: "x \<in> s" "y \<in> s" "0 \<le> \<mu>" "\<mu> \<le> 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
437 |
show "(1 - \<mu>) *\<^sub>R x + \<mu> *\<^sub>R y \<in> s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
438 |
proof (cases "x = y") |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
439 |
case False |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
440 |
then show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
441 |
using *[rule_format, of "{x, y}" "\<lambda> z. if z = x then 1 - \<mu> else \<mu>"] **
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
442 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
443 |
next |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
444 |
case True |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
445 |
then show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
446 |
using *[rule_format, of "{x, y}" "\<lambda> z. 1"] **
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
447 |
by (auto simp: field_simps real_vector.scale_left_diff_distrib) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
448 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
449 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
450 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
451 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
452 |
lemma convex_finite: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
453 |
assumes "finite s" |
| 64267 | 454 |
shows "convex s \<longleftrightarrow> (\<forall>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<longrightarrow> sum (\<lambda>x. u x *\<^sub>R x) s \<in> s)" |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
455 |
unfolding convex_explicit |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
456 |
apply safe |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
457 |
subgoal for u by (erule allE [where x=s], erule allE [where x=u]) auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
458 |
subgoal for t u |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
459 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
460 |
have if_distrib_arg: "\<And>P f g x. (if P then f else g) x = (if P then f x else g x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
461 |
by simp |
| 64267 | 462 |
assume sum: "\<forall>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> s" |
463 |
assume *: "\<forall>x\<in>t. 0 \<le> u x" "sum u t = 1" |
|
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
464 |
assume "t \<subseteq> s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
465 |
then have "s \<inter> t = t" by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
466 |
with sum[THEN spec[where x="\<lambda>x. if x\<in>t then u x else 0"]] * show "(\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" |
| 64267 | 467 |
by (auto simp: assms sum.If_cases if_distrib if_distrib_arg) |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
468 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
469 |
done |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
470 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
471 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
472 |
subsection \<open>Functions that are convex on a set\<close> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
473 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
474 |
definition convex_on :: "'a::real_vector set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
475 |
where "convex_on s f \<longleftrightarrow> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
476 |
(\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u\<ge>0. \<forall>v\<ge>0. u + v = 1 \<longrightarrow> f (u *\<^sub>R x + v *\<^sub>R y) \<le> u * f x + v * f y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
477 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
478 |
lemma convex_onI [intro?]: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
479 |
assumes "\<And>t x y. t > 0 \<Longrightarrow> t < 1 \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
480 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
481 |
shows "convex_on A f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
482 |
unfolding convex_on_def |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
483 |
proof clarify |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
484 |
fix x y |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
485 |
fix u v :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
486 |
assume A: "x \<in> A" "y \<in> A" "u \<ge> 0" "v \<ge> 0" "u + v = 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
487 |
from A(5) have [simp]: "v = 1 - u" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
488 |
by (simp add: algebra_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
489 |
from A(1-4) show "f (u *\<^sub>R x + v *\<^sub>R y) \<le> u * f x + v * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
490 |
using assms[of u y x] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
491 |
by (cases "u = 0 \<or> u = 1") (auto simp: algebra_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
492 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
493 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
494 |
lemma convex_on_linorderI [intro?]: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
495 |
fixes A :: "('a::{linorder,real_vector}) set"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
496 |
assumes "\<And>t x y. t > 0 \<Longrightarrow> t < 1 \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x < y \<Longrightarrow> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
497 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
498 |
shows "convex_on A f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
499 |
proof |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
500 |
fix x y |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
501 |
fix t :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
502 |
assume A: "x \<in> A" "y \<in> A" "t > 0" "t < 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
503 |
with assms [of t x y] assms [of "1 - t" y x] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
504 |
show "f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
505 |
by (cases x y rule: linorder_cases) (auto simp: algebra_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
506 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
507 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
508 |
lemma convex_onD: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
509 |
assumes "convex_on A f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
510 |
shows "\<And>t x y. t \<ge> 0 \<Longrightarrow> t \<le> 1 \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
511 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
512 |
using assms by (auto simp: convex_on_def) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
513 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
514 |
lemma convex_onD_Icc: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
515 |
assumes "convex_on {x..y} f" "x \<le> (y :: _ :: {real_vector,preorder})"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
516 |
shows "\<And>t. t \<ge> 0 \<Longrightarrow> t \<le> 1 \<Longrightarrow> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
517 |
f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \<le> (1 - t) * f x + t * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
518 |
using assms(2) by (intro convex_onD [OF assms(1)]) simp_all |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
519 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
520 |
lemma convex_on_subset: "convex_on t f \<Longrightarrow> s \<subseteq> t \<Longrightarrow> convex_on s f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
521 |
unfolding convex_on_def by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
522 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
523 |
lemma convex_on_add [intro]: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
524 |
assumes "convex_on s f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
525 |
and "convex_on s g" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
526 |
shows "convex_on s (\<lambda>x. f x + g x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
527 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
528 |
{
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
529 |
fix x y |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
530 |
assume "x \<in> s" "y \<in> s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
531 |
moreover |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
532 |
fix u v :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
533 |
assume "0 \<le> u" "0 \<le> v" "u + v = 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
534 |
ultimately |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
535 |
have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \<le> (u * f x + v * f y) + (u * g x + v * g y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
536 |
using assms unfolding convex_on_def by (auto simp: add_mono) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
537 |
then have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \<le> u * (f x + g x) + v * (f y + g y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
538 |
by (simp add: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
539 |
} |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
540 |
then show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
541 |
unfolding convex_on_def by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
542 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
543 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
544 |
lemma convex_on_cmul [intro]: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
545 |
fixes c :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
546 |
assumes "0 \<le> c" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
547 |
and "convex_on s f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
548 |
shows "convex_on s (\<lambda>x. c * f x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
549 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
550 |
have *: "u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
551 |
for u c fx v fy :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
552 |
by (simp add: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
553 |
show ?thesis using assms(2) and mult_left_mono [OF _ assms(1)] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
554 |
unfolding convex_on_def and * by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
555 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
556 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
557 |
lemma convex_lower: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
558 |
assumes "convex_on s f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
559 |
and "x \<in> s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
560 |
and "y \<in> s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
561 |
and "0 \<le> u" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
562 |
and "0 \<le> v" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
563 |
and "u + v = 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
564 |
shows "f (u *\<^sub>R x + v *\<^sub>R y) \<le> max (f x) (f y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
565 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
566 |
let ?m = "max (f x) (f y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
567 |
have "u * f x + v * f y \<le> u * max (f x) (f y) + v * max (f x) (f y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
568 |
using assms(4,5) by (auto simp: mult_left_mono add_mono) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
569 |
also have "\<dots> = max (f x) (f y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
570 |
using assms(6) by (simp add: distrib_right [symmetric]) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
571 |
finally show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
572 |
using assms unfolding convex_on_def by fastforce |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
573 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
574 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
575 |
lemma convex_on_dist [intro]: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
576 |
fixes s :: "'a::real_normed_vector set" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
577 |
shows "convex_on s (\<lambda>x. dist a x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
578 |
proof (auto simp: convex_on_def dist_norm) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
579 |
fix x y |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
580 |
assume "x \<in> s" "y \<in> s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
581 |
fix u v :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
582 |
assume "0 \<le> u" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
583 |
assume "0 \<le> v" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
584 |
assume "u + v = 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
585 |
have "a = u *\<^sub>R a + v *\<^sub>R a" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
586 |
unfolding scaleR_left_distrib[symmetric] and \<open>u + v = 1\<close> by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
587 |
then have *: "a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
588 |
by (auto simp: algebra_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
589 |
show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \<le> u * norm (a - x) + v * norm (a - y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
590 |
unfolding * using norm_triangle_ineq[of "u *\<^sub>R (a - x)" "v *\<^sub>R (a - y)"] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
591 |
using \<open>0 \<le> u\<close> \<open>0 \<le> v\<close> by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
592 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
593 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
594 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
595 |
subsection \<open>Arithmetic operations on sets preserve convexity\<close> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
596 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
597 |
lemma convex_linear_image: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
598 |
assumes "linear f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
599 |
and "convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
600 |
shows "convex (f ` s)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
601 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
602 |
interpret f: linear f by fact |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
603 |
from \<open>convex s\<close> show "convex (f ` s)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
604 |
by (simp add: convex_def f.scaleR [symmetric] f.add [symmetric]) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
605 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
606 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
607 |
lemma convex_linear_vimage: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
608 |
assumes "linear f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
609 |
and "convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
610 |
shows "convex (f -` s)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
611 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
612 |
interpret f: linear f by fact |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
613 |
from \<open>convex s\<close> show "convex (f -` s)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
614 |
by (simp add: convex_def f.add f.scaleR) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
615 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
616 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
617 |
lemma convex_scaling: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
618 |
assumes "convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
619 |
shows "convex ((\<lambda>x. c *\<^sub>R x) ` s)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
620 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
621 |
have "linear (\<lambda>x. c *\<^sub>R x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
622 |
by (simp add: linearI scaleR_add_right) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
623 |
then show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
624 |
using \<open>convex s\<close> by (rule convex_linear_image) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
625 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
626 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
627 |
lemma convex_scaled: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
628 |
assumes "convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
629 |
shows "convex ((\<lambda>x. x *\<^sub>R c) ` s)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
630 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
631 |
have "linear (\<lambda>x. x *\<^sub>R c)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
632 |
by (simp add: linearI scaleR_add_left) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
633 |
then show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
634 |
using \<open>convex s\<close> by (rule convex_linear_image) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
635 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
636 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
637 |
lemma convex_negations: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
638 |
assumes "convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
639 |
shows "convex ((\<lambda>x. - x) ` s)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
640 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
641 |
have "linear (\<lambda>x. - x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
642 |
by (simp add: linearI) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
643 |
then show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
644 |
using \<open>convex s\<close> by (rule convex_linear_image) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
645 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
646 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
647 |
lemma convex_sums: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
648 |
assumes "convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
649 |
and "convex t" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
650 |
shows "convex {x + y| x y. x \<in> s \<and> y \<in> t}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
651 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
652 |
have "linear (\<lambda>(x, y). x + y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
653 |
by (auto intro: linearI simp: scaleR_add_right) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
654 |
with assms have "convex ((\<lambda>(x, y). x + y) ` (s \<times> t))" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
655 |
by (intro convex_linear_image convex_Times) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
656 |
also have "((\<lambda>(x, y). x + y) ` (s \<times> t)) = {x + y| x y. x \<in> s \<and> y \<in> t}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
657 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
658 |
finally show ?thesis . |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
659 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
660 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
661 |
lemma convex_differences: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
662 |
assumes "convex s" "convex t" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
663 |
shows "convex {x - y| x y. x \<in> s \<and> y \<in> t}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
664 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
665 |
have "{x - y| x y. x \<in> s \<and> y \<in> t} = {x + y |x y. x \<in> s \<and> y \<in> uminus ` t}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
666 |
by (auto simp: diff_conv_add_uminus simp del: add_uminus_conv_diff) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
667 |
then show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
668 |
using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
669 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
670 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
671 |
lemma convex_translation: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
672 |
assumes "convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
673 |
shows "convex ((\<lambda>x. a + x) ` s)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
674 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
675 |
have "{a + y |y. y \<in> s} = (\<lambda>x. a + x) ` s"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
676 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
677 |
then show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
678 |
using convex_sums[OF convex_singleton[of a] assms] by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
679 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
680 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
681 |
lemma convex_affinity: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
682 |
assumes "convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
683 |
shows "convex ((\<lambda>x. a + c *\<^sub>R x) ` s)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
684 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
685 |
have "(\<lambda>x. a + c *\<^sub>R x) ` s = op + a ` op *\<^sub>R c ` s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
686 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
687 |
then show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
688 |
using convex_translation[OF convex_scaling[OF assms], of a c] by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
689 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
690 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
691 |
lemma pos_is_convex: "convex {0 :: real <..}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
692 |
unfolding convex_alt |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
693 |
proof safe |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
694 |
fix y x \<mu> :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
695 |
assume *: "y > 0" "x > 0" "\<mu> \<ge> 0" "\<mu> \<le> 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
696 |
{
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
697 |
assume "\<mu> = 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
698 |
then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y = y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
699 |
by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
700 |
then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y > 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
701 |
using * by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
702 |
} |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
703 |
moreover |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
704 |
{
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
705 |
assume "\<mu> = 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
706 |
then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y > 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
707 |
using * by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
708 |
} |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
709 |
moreover |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
710 |
{
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
711 |
assume "\<mu> \<noteq> 1" "\<mu> \<noteq> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
712 |
then have "\<mu> > 0" "(1 - \<mu>) > 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
713 |
using * by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
714 |
then have "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y > 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
715 |
using * by (auto simp: add_pos_pos) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
716 |
} |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
717 |
ultimately show "(1 - \<mu>) *\<^sub>R y + \<mu> *\<^sub>R x > 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
718 |
by fastforce |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
719 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
720 |
|
| 64267 | 721 |
lemma convex_on_sum: |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
722 |
fixes a :: "'a \<Rightarrow> real" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
723 |
and y :: "'a \<Rightarrow> 'b::real_vector" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
724 |
and f :: "'b \<Rightarrow> real" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
725 |
assumes "finite s" "s \<noteq> {}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
726 |
and "convex_on C f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
727 |
and "convex C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
728 |
and "(\<Sum> i \<in> s. a i) = 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
729 |
and "\<And>i. i \<in> s \<Longrightarrow> a i \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
730 |
and "\<And>i. i \<in> s \<Longrightarrow> y i \<in> C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
731 |
shows "f (\<Sum> i \<in> s. a i *\<^sub>R y i) \<le> (\<Sum> i \<in> s. a i * f (y i))" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
732 |
using assms |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
733 |
proof (induct s arbitrary: a rule: finite_ne_induct) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
734 |
case (singleton i) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
735 |
then have ai: "a i = 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
736 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
737 |
then show ?case |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
738 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
739 |
next |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
740 |
case (insert i s) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
741 |
then have "convex_on C f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
742 |
by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
743 |
from this[unfolded convex_on_def, rule_format] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
744 |
have conv: "\<And>x y \<mu>. x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> 0 \<le> \<mu> \<Longrightarrow> \<mu> \<le> 1 \<Longrightarrow> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
745 |
f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
746 |
by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
747 |
show ?case |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
748 |
proof (cases "a i = 1") |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
749 |
case True |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
750 |
then have "(\<Sum> j \<in> s. a j) = 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
751 |
using insert by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
752 |
then have "\<And>j. j \<in> s \<Longrightarrow> a j = 0" |
| 64267 | 753 |
using insert by (fastforce simp: sum_nonneg_eq_0_iff) |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
754 |
then show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
755 |
using insert by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
756 |
next |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
757 |
case False |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
758 |
from insert have yai: "y i \<in> C" "a i \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
759 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
760 |
have fis: "finite (insert i s)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
761 |
using insert by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
762 |
then have ai1: "a i \<le> 1" |
| 64267 | 763 |
using sum_nonneg_leq_bound[of "insert i s" a] insert by simp |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
764 |
then have "a i < 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
765 |
using False by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
766 |
then have i0: "1 - a i > 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
767 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
768 |
let ?a = "\<lambda>j. a j / (1 - a i)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
769 |
have a_nonneg: "?a j \<ge> 0" if "j \<in> s" for j |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
770 |
using i0 insert that by fastforce |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
771 |
have "(\<Sum> j \<in> insert i s. a j) = 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
772 |
using insert by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
773 |
then have "(\<Sum> j \<in> s. a j) = 1 - a i" |
| 64267 | 774 |
using sum.insert insert by fastforce |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
775 |
then have "(\<Sum> j \<in> s. a j) / (1 - a i) = 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
776 |
using i0 by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
777 |
then have a1: "(\<Sum> j \<in> s. ?a j) = 1" |
| 64267 | 778 |
unfolding sum_divide_distrib by simp |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
779 |
have "convex C" using insert by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
780 |
then have asum: "(\<Sum> j \<in> s. ?a j *\<^sub>R y j) \<in> C" |
| 64267 | 781 |
using insert convex_sum [OF \<open>finite s\<close> \<open>convex C\<close> a1 a_nonneg] by auto |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
782 |
have asum_le: "f (\<Sum> j \<in> s. ?a j *\<^sub>R y j) \<le> (\<Sum> j \<in> s. ?a j * f (y j))" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
783 |
using a_nonneg a1 insert by blast |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
784 |
have "f (\<Sum> j \<in> insert i s. a j *\<^sub>R y j) = f ((\<Sum> j \<in> s. a j *\<^sub>R y j) + a i *\<^sub>R y i)" |
| 64267 | 785 |
using sum.insert[of s i "\<lambda> j. a j *\<^sub>R y j", OF \<open>finite s\<close> \<open>i \<notin> s\<close>] insert |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
786 |
by (auto simp only: add.commute) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
787 |
also have "\<dots> = f (((1 - a i) * inverse (1 - a i)) *\<^sub>R (\<Sum> j \<in> s. a j *\<^sub>R y j) + a i *\<^sub>R y i)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
788 |
using i0 by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
789 |
also have "\<dots> = f ((1 - a i) *\<^sub>R (\<Sum> j \<in> s. (a j * inverse (1 - a i)) *\<^sub>R y j) + a i *\<^sub>R y i)" |
| 64267 | 790 |
using scaleR_right.sum[of "inverse (1 - a i)" "\<lambda> j. a j *\<^sub>R y j" s, symmetric] |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
791 |
by (auto simp: algebra_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
792 |
also have "\<dots> = f ((1 - a i) *\<^sub>R (\<Sum> j \<in> s. ?a j *\<^sub>R y j) + a i *\<^sub>R y i)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
793 |
by (auto simp: divide_inverse) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
794 |
also have "\<dots> \<le> (1 - a i) *\<^sub>R f ((\<Sum> j \<in> s. ?a j *\<^sub>R y j)) + a i * f (y i)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
795 |
using conv[of "y i" "(\<Sum> j \<in> s. ?a j *\<^sub>R y j)" "a i", OF yai(1) asum yai(2) ai1] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
796 |
by (auto simp: add.commute) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
797 |
also have "\<dots> \<le> (1 - a i) * (\<Sum> j \<in> s. ?a j * f (y j)) + a i * f (y i)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
798 |
using add_right_mono [OF mult_left_mono [of _ _ "1 - a i", |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
799 |
OF asum_le less_imp_le[OF i0]], of "a i * f (y i)"] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
800 |
by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
801 |
also have "\<dots> = (\<Sum> j \<in> s. (1 - a i) * ?a j * f (y j)) + a i * f (y i)" |
| 64267 | 802 |
unfolding sum_distrib_left[of "1 - a i" "\<lambda> j. ?a j * f (y j)"] |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
803 |
using i0 by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
804 |
also have "\<dots> = (\<Sum> j \<in> s. a j * f (y j)) + a i * f (y i)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
805 |
using i0 by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
806 |
also have "\<dots> = (\<Sum> j \<in> insert i s. a j * f (y j))" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
807 |
using insert by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
808 |
finally show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
809 |
by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
810 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
811 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
812 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
813 |
lemma convex_on_alt: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
814 |
fixes C :: "'a::real_vector set" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
815 |
assumes "convex C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
816 |
shows "convex_on C f \<longleftrightarrow> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
817 |
(\<forall>x \<in> C. \<forall> y \<in> C. \<forall> \<mu> :: real. \<mu> \<ge> 0 \<and> \<mu> \<le> 1 \<longrightarrow> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
818 |
f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
819 |
proof safe |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
820 |
fix x y |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
821 |
fix \<mu> :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
822 |
assume *: "convex_on C f" "x \<in> C" "y \<in> C" "0 \<le> \<mu>" "\<mu> \<le> 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
823 |
from this[unfolded convex_on_def, rule_format] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
824 |
have "0 \<le> u \<Longrightarrow> 0 \<le> v \<Longrightarrow> u + v = 1 \<Longrightarrow> f (u *\<^sub>R x + v *\<^sub>R y) \<le> u * f x + v * f y" for u v |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
825 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
826 |
from this [of "\<mu>" "1 - \<mu>", simplified] * |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
827 |
show "f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
828 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
829 |
next |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
830 |
assume *: "\<forall>x\<in>C. \<forall>y\<in>C. \<forall>\<mu>. 0 \<le> \<mu> \<and> \<mu> \<le> 1 \<longrightarrow> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
831 |
f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
832 |
{
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
833 |
fix x y |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
834 |
fix u v :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
835 |
assume **: "x \<in> C" "y \<in> C" "u \<ge> 0" "v \<ge> 0" "u + v = 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
836 |
then have[simp]: "1 - u = v" by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
837 |
from *[rule_format, of x y u] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
838 |
have "f (u *\<^sub>R x + v *\<^sub>R y) \<le> u * f x + v * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
839 |
using ** by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
840 |
} |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
841 |
then show "convex_on C f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
842 |
unfolding convex_on_def by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
843 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
844 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
845 |
lemma convex_on_diff: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
846 |
fixes f :: "real \<Rightarrow> real" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
847 |
assumes f: "convex_on I f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
848 |
and I: "x \<in> I" "y \<in> I" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
849 |
and t: "x < t" "t < y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
850 |
shows "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
851 |
and "(f x - f y) / (x - y) \<le> (f t - f y) / (t - y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
852 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
853 |
define a where "a \<equiv> (t - y) / (x - y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
854 |
with t have "0 \<le> a" "0 \<le> 1 - a" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
855 |
by (auto simp: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
856 |
with f \<open>x \<in> I\<close> \<open>y \<in> I\<close> have cvx: "f (a * x + (1 - a) * y) \<le> a * f x + (1 - a) * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
857 |
by (auto simp: convex_on_def) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
858 |
have "a * x + (1 - a) * y = a * (x - y) + y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
859 |
by (simp add: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
860 |
also have "\<dots> = t" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
861 |
unfolding a_def using \<open>x < t\<close> \<open>t < y\<close> by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
862 |
finally have "f t \<le> a * f x + (1 - a) * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
863 |
using cvx by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
864 |
also have "\<dots> = a * (f x - f y) + f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
865 |
by (simp add: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
866 |
finally have "f t - f y \<le> a * (f x - f y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
867 |
by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
868 |
with t show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
869 |
by (simp add: le_divide_eq divide_le_eq field_simps a_def) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
870 |
with t show "(f x - f y) / (x - y) \<le> (f t - f y) / (t - y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
871 |
by (simp add: le_divide_eq divide_le_eq field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
872 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
873 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
874 |
lemma pos_convex_function: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
875 |
fixes f :: "real \<Rightarrow> real" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
876 |
assumes "convex C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
877 |
and leq: "\<And>x y. x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> f' x * (y - x) \<le> f y - f x" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
878 |
shows "convex_on C f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
879 |
unfolding convex_on_alt[OF assms(1)] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
880 |
using assms |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
881 |
proof safe |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
882 |
fix x y \<mu> :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
883 |
let ?x = "\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
884 |
assume *: "convex C" "x \<in> C" "y \<in> C" "\<mu> \<ge> 0" "\<mu> \<le> 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
885 |
then have "1 - \<mu> \<ge> 0" by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
886 |
then have xpos: "?x \<in> C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
887 |
using * unfolding convex_alt by fastforce |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
888 |
have geq: "\<mu> * (f x - f ?x) + (1 - \<mu>) * (f y - f ?x) \<ge> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
889 |
\<mu> * f' ?x * (x - ?x) + (1 - \<mu>) * f' ?x * (y - ?x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
890 |
using add_mono [OF mult_left_mono [OF leq [OF xpos *(2)] \<open>\<mu> \<ge> 0\<close>] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
891 |
mult_left_mono [OF leq [OF xpos *(3)] \<open>1 - \<mu> \<ge> 0\<close>]] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
892 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
893 |
then have "\<mu> * f x + (1 - \<mu>) * f y - f ?x \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
894 |
by (auto simp: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
895 |
then show "f (\<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y) \<le> \<mu> * f x + (1 - \<mu>) * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
896 |
using convex_on_alt by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
897 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
898 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
899 |
lemma atMostAtLeast_subset_convex: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
900 |
fixes C :: "real set" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
901 |
assumes "convex C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
902 |
and "x \<in> C" "y \<in> C" "x < y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
903 |
shows "{x .. y} \<subseteq> C"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
904 |
proof safe |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
905 |
fix z assume z: "z \<in> {x .. y}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
906 |
have less: "z \<in> C" if *: "x < z" "z < y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
907 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
908 |
let ?\<mu> = "(y - z) / (y - x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
909 |
have "0 \<le> ?\<mu>" "?\<mu> \<le> 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
910 |
using assms * by (auto simp: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
911 |
then have comb: "?\<mu> * x + (1 - ?\<mu>) * y \<in> C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
912 |
using assms iffD1[OF convex_alt, rule_format, of C y x ?\<mu>] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
913 |
by (simp add: algebra_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
914 |
have "?\<mu> * x + (1 - ?\<mu>) * y = (y - z) * x / (y - x) + (1 - (y - z) / (y - x)) * y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
915 |
by (auto simp: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
916 |
also have "\<dots> = ((y - z) * x + (y - x - (y - z)) * y) / (y - x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
917 |
using assms by (simp only: add_divide_distrib) (auto simp: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
918 |
also have "\<dots> = z" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
919 |
using assms by (auto simp: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
920 |
finally show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
921 |
using comb by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
922 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
923 |
show "z \<in> C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
924 |
using z less assms by (auto simp: le_less) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
925 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
926 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
927 |
lemma f''_imp_f': |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
928 |
fixes f :: "real \<Rightarrow> real" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
929 |
assumes "convex C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
930 |
and f': "\<And>x. x \<in> C \<Longrightarrow> DERIV f x :> (f' x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
931 |
and f'': "\<And>x. x \<in> C \<Longrightarrow> DERIV f' x :> (f'' x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
932 |
and pos: "\<And>x. x \<in> C \<Longrightarrow> f'' x \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
933 |
and x: "x \<in> C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
934 |
and y: "y \<in> C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
935 |
shows "f' x * (y - x) \<le> f y - f x" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
936 |
using assms |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
937 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
938 |
have less_imp: "f y - f x \<ge> f' x * (y - x)" "f' y * (x - y) \<le> f x - f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
939 |
if *: "x \<in> C" "y \<in> C" "y > x" for x y :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
940 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
941 |
from * have ge: "y - x > 0" "y - x \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
942 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
943 |
from * have le: "x - y < 0" "x - y \<le> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
944 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
945 |
then obtain z1 where z1: "z1 > x" "z1 < y" "f y - f x = (y - x) * f' z1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
946 |
using subsetD[OF atMostAtLeast_subset_convex[OF \<open>convex C\<close> \<open>x \<in> C\<close> \<open>y \<in> C\<close> \<open>x < y\<close>], |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
947 |
THEN f', THEN MVT2[OF \<open>x < y\<close>, rule_format, unfolded atLeastAtMost_iff[symmetric]]] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
948 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
949 |
then have "z1 \<in> C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
950 |
using atMostAtLeast_subset_convex \<open>convex C\<close> \<open>x \<in> C\<close> \<open>y \<in> C\<close> \<open>x < y\<close> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
951 |
by fastforce |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
952 |
from z1 have z1': "f x - f y = (x - y) * f' z1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
953 |
by (simp add: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
954 |
obtain z2 where z2: "z2 > x" "z2 < z1" "f' z1 - f' x = (z1 - x) * f'' z2" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
955 |
using subsetD[OF atMostAtLeast_subset_convex[OF \<open>convex C\<close> \<open>x \<in> C\<close> \<open>z1 \<in> C\<close> \<open>x < z1\<close>], |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
956 |
THEN f'', THEN MVT2[OF \<open>x < z1\<close>, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
957 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
958 |
obtain z3 where z3: "z3 > z1" "z3 < y" "f' y - f' z1 = (y - z1) * f'' z3" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
959 |
using subsetD[OF atMostAtLeast_subset_convex[OF \<open>convex C\<close> \<open>z1 \<in> C\<close> \<open>y \<in> C\<close> \<open>z1 < y\<close>], |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
960 |
THEN f'', THEN MVT2[OF \<open>z1 < y\<close>, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
961 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
962 |
have "f' y - (f x - f y) / (x - y) = f' y - f' z1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
963 |
using * z1' by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
964 |
also have "\<dots> = (y - z1) * f'' z3" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
965 |
using z3 by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
966 |
finally have cool': "f' y - (f x - f y) / (x - y) = (y - z1) * f'' z3" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
967 |
by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
968 |
have A': "y - z1 \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
969 |
using z1 by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
970 |
have "z3 \<in> C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
971 |
using z3 * atMostAtLeast_subset_convex \<open>convex C\<close> \<open>x \<in> C\<close> \<open>z1 \<in> C\<close> \<open>x < z1\<close> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
972 |
by fastforce |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
973 |
then have B': "f'' z3 \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
974 |
using assms by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
975 |
from A' B' have "(y - z1) * f'' z3 \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
976 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
977 |
from cool' this have "f' y - (f x - f y) / (x - y) \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
978 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
979 |
from mult_right_mono_neg[OF this le(2)] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
980 |
have "f' y * (x - y) - (f x - f y) / (x - y) * (x - y) \<le> 0 * (x - y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
981 |
by (simp add: algebra_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
982 |
then have "f' y * (x - y) - (f x - f y) \<le> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
983 |
using le by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
984 |
then have res: "f' y * (x - y) \<le> f x - f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
985 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
986 |
have "(f y - f x) / (y - x) - f' x = f' z1 - f' x" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
987 |
using * z1 by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
988 |
also have "\<dots> = (z1 - x) * f'' z2" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
989 |
using z2 by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
990 |
finally have cool: "(f y - f x) / (y - x) - f' x = (z1 - x) * f'' z2" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
991 |
by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
992 |
have A: "z1 - x \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
993 |
using z1 by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
994 |
have "z2 \<in> C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
995 |
using z2 z1 * atMostAtLeast_subset_convex \<open>convex C\<close> \<open>z1 \<in> C\<close> \<open>y \<in> C\<close> \<open>z1 < y\<close> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
996 |
by fastforce |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
997 |
then have B: "f'' z2 \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
998 |
using assms by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
999 |
from A B have "(z1 - x) * f'' z2 \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1000 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1001 |
with cool have "(f y - f x) / (y - x) - f' x \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1002 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1003 |
from mult_right_mono[OF this ge(2)] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1004 |
have "(f y - f x) / (y - x) * (y - x) - f' x * (y - x) \<ge> 0 * (y - x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1005 |
by (simp add: algebra_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1006 |
then have "f y - f x - f' x * (y - x) \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1007 |
using ge by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1008 |
then show "f y - f x \<ge> f' x * (y - x)" "f' y * (x - y) \<le> f x - f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1009 |
using res by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1010 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1011 |
show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1012 |
proof (cases "x = y") |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1013 |
case True |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1014 |
with x y show ?thesis by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1015 |
next |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1016 |
case False |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1017 |
with less_imp x y show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1018 |
by (auto simp: neq_iff) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1019 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1020 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1021 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1022 |
lemma f''_ge0_imp_convex: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1023 |
fixes f :: "real \<Rightarrow> real" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1024 |
assumes conv: "convex C" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1025 |
and f': "\<And>x. x \<in> C \<Longrightarrow> DERIV f x :> (f' x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1026 |
and f'': "\<And>x. x \<in> C \<Longrightarrow> DERIV f' x :> (f'' x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1027 |
and pos: "\<And>x. x \<in> C \<Longrightarrow> f'' x \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1028 |
shows "convex_on C f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1029 |
using f''_imp_f'[OF conv f' f'' pos] assms pos_convex_function |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1030 |
by fastforce |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1031 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1032 |
lemma minus_log_convex: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1033 |
fixes b :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1034 |
assumes "b > 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1035 |
shows "convex_on {0 <..} (\<lambda> x. - log b x)"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1036 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1037 |
have "\<And>z. z > 0 \<Longrightarrow> DERIV (log b) z :> 1 / (ln b * z)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1038 |
using DERIV_log by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1039 |
then have f': "\<And>z. z > 0 \<Longrightarrow> DERIV (\<lambda> z. - log b z) z :> - 1 / (ln b * z)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1040 |
by (auto simp: DERIV_minus) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1041 |
have "\<And>z::real. z > 0 \<Longrightarrow> DERIV inverse z :> - (inverse z ^ Suc (Suc 0))" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1042 |
using less_imp_neq[THEN not_sym, THEN DERIV_inverse] by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1043 |
from this[THEN DERIV_cmult, of _ "- 1 / ln b"] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1044 |
have "\<And>z::real. z > 0 \<Longrightarrow> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1045 |
DERIV (\<lambda> z. (- 1 / ln b) * inverse z) z :> (- 1 / ln b) * (- (inverse z ^ Suc (Suc 0)))" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1046 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1047 |
then have f''0: "\<And>z::real. z > 0 \<Longrightarrow> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1048 |
DERIV (\<lambda> z. - 1 / (ln b * z)) z :> 1 / (ln b * z * z)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1049 |
unfolding inverse_eq_divide by (auto simp: mult.assoc) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1050 |
have f''_ge0: "\<And>z::real. z > 0 \<Longrightarrow> 1 / (ln b * z * z) \<ge> 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1051 |
using \<open>b > 1\<close> by (auto intro!: less_imp_le) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1052 |
from f''_ge0_imp_convex[OF pos_is_convex, unfolded greaterThan_iff, OF f' f''0 f''_ge0] |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1053 |
show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1054 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1055 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1056 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1057 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1058 |
subsection \<open>Convexity of real functions\<close> |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1059 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1060 |
lemma convex_on_realI: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1061 |
assumes "connected A" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1062 |
and "\<And>x. x \<in> A \<Longrightarrow> (f has_real_derivative f' x) (at x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1063 |
and "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<le> y \<Longrightarrow> f' x \<le> f' y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1064 |
shows "convex_on A f" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1065 |
proof (rule convex_on_linorderI) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1066 |
fix t x y :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1067 |
assume t: "t > 0" "t < 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1068 |
assume xy: "x \<in> A" "y \<in> A" "x < y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1069 |
define z where "z = (1 - t) * x + t * y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1070 |
with \<open>connected A\<close> and xy have ivl: "{x..y} \<subseteq> A"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1071 |
using connected_contains_Icc by blast |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1072 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1073 |
from xy t have xz: "z > x" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1074 |
by (simp add: z_def algebra_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1075 |
have "y - z = (1 - t) * (y - x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1076 |
by (simp add: z_def algebra_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1077 |
also from xy t have "\<dots> > 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1078 |
by (intro mult_pos_pos) simp_all |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1079 |
finally have yz: "z < y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1080 |
by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1081 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1082 |
from assms xz yz ivl t have "\<exists>\<xi>. \<xi> > x \<and> \<xi> < z \<and> f z - f x = (z - x) * f' \<xi>" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1083 |
by (intro MVT2) (auto intro!: assms(2)) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1084 |
then obtain \<xi> where \<xi>: "\<xi> > x" "\<xi> < z" "f' \<xi> = (f z - f x) / (z - x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1085 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1086 |
from assms xz yz ivl t have "\<exists>\<eta>. \<eta> > z \<and> \<eta> < y \<and> f y - f z = (y - z) * f' \<eta>" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1087 |
by (intro MVT2) (auto intro!: assms(2)) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1088 |
then obtain \<eta> where \<eta>: "\<eta> > z" "\<eta> < y" "f' \<eta> = (f y - f z) / (y - z)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1089 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1090 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1091 |
from \<eta>(3) have "(f y - f z) / (y - z) = f' \<eta>" .. |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1092 |
also from \<xi> \<eta> ivl have "\<xi> \<in> A" "\<eta> \<in> A" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1093 |
by auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1094 |
with \<xi> \<eta> have "f' \<eta> \<ge> f' \<xi>" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1095 |
by (intro assms(3)) auto |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1096 |
also from \<xi>(3) have "f' \<xi> = (f z - f x) / (z - x)" . |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1097 |
finally have "(f y - f z) * (z - x) \<ge> (f z - f x) * (y - z)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1098 |
using xz yz by (simp add: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1099 |
also have "z - x = t * (y - x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1100 |
by (simp add: z_def algebra_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1101 |
also have "y - z = (1 - t) * (y - x)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1102 |
by (simp add: z_def algebra_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1103 |
finally have "(f y - f z) * t \<ge> (f z - f x) * (1 - t)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1104 |
using xy by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1105 |
then show "(1 - t) * f x + t * f y \<ge> f ((1 - t) *\<^sub>R x + t *\<^sub>R y)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1106 |
by (simp add: z_def algebra_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1107 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1108 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1109 |
lemma convex_on_inverse: |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1110 |
assumes "A \<subseteq> {0<..}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1111 |
shows "convex_on A (inverse :: real \<Rightarrow> real)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1112 |
proof (rule convex_on_subset[OF _ assms], intro convex_on_realI[of _ _ "\<lambda>x. -inverse (x^2)"]) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1113 |
fix u v :: real |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1114 |
assume "u \<in> {0<..}" "v \<in> {0<..}" "u \<le> v"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1115 |
with assms show "-inverse (u^2) \<le> -inverse (v^2)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1116 |
by (intro le_imp_neg_le le_imp_inverse_le power_mono) (simp_all) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1117 |
qed (insert assms, auto intro!: derivative_eq_intros simp: divide_simps power2_eq_square) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1118 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1119 |
lemma convex_onD_Icc': |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1120 |
assumes "convex_on {x..y} f" "c \<in> {x..y}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1121 |
defines "d \<equiv> y - x" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1122 |
shows "f c \<le> (f y - f x) / d * (c - x) + f x" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1123 |
proof (cases x y rule: linorder_cases) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1124 |
case less |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1125 |
then have d: "d > 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1126 |
by (simp add: d_def) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1127 |
from assms(2) less have A: "0 \<le> (c - x) / d" "(c - x) / d \<le> 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1128 |
by (simp_all add: d_def divide_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1129 |
have "f c = f (x + (c - x) * 1)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1130 |
by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1131 |
also from less have "1 = ((y - x) / d)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1132 |
by (simp add: d_def) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1133 |
also from d have "x + (c - x) * \<dots> = (1 - (c - x) / d) *\<^sub>R x + ((c - x) / d) *\<^sub>R y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1134 |
by (simp add: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1135 |
also have "f \<dots> \<le> (1 - (c - x) / d) * f x + (c - x) / d * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1136 |
using assms less by (intro convex_onD_Icc) simp_all |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1137 |
also from d have "\<dots> = (f y - f x) / d * (c - x) + f x" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1138 |
by (simp add: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1139 |
finally show ?thesis . |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1140 |
qed (insert assms(2), simp_all) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1141 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1142 |
lemma convex_onD_Icc'': |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1143 |
assumes "convex_on {x..y} f" "c \<in> {x..y}"
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1144 |
defines "d \<equiv> y - x" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1145 |
shows "f c \<le> (f x - f y) / d * (y - c) + f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1146 |
proof (cases x y rule: linorder_cases) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1147 |
case less |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1148 |
then have d: "d > 0" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1149 |
by (simp add: d_def) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1150 |
from assms(2) less have A: "0 \<le> (y - c) / d" "(y - c) / d \<le> 1" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1151 |
by (simp_all add: d_def divide_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1152 |
have "f c = f (y - (y - c) * 1)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1153 |
by simp |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1154 |
also from less have "1 = ((y - x) / d)" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1155 |
by (simp add: d_def) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1156 |
also from d have "y - (y - c) * \<dots> = (1 - (1 - (y - c) / d)) *\<^sub>R x + (1 - (y - c) / d) *\<^sub>R y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1157 |
by (simp add: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1158 |
also have "f \<dots> \<le> (1 - (1 - (y - c) / d)) * f x + (1 - (y - c) / d) * f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1159 |
using assms less by (intro convex_onD_Icc) (simp_all add: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1160 |
also from d have "\<dots> = (f x - f y) / d * (y - c) + f y" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1161 |
by (simp add: field_simps) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1162 |
finally show ?thesis . |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1163 |
qed (insert assms(2), simp_all) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1164 |
|
| 64267 | 1165 |
lemma convex_supp_sum: |
1166 |
assumes "convex S" and 1: "supp_sum u I = 1" |
|
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1167 |
and "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> u i \<and> (u i = 0 \<or> f i \<in> S)" |
| 64267 | 1168 |
shows "supp_sum (\<lambda>i. u i *\<^sub>R f i) I \<in> S" |
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1169 |
proof - |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1170 |
have fin: "finite {i \<in> I. u i \<noteq> 0}"
|
| 64267 | 1171 |
using 1 sum.infinite by (force simp: supp_sum_def support_on_def) |
1172 |
then have eq: "supp_sum (\<lambda>i. u i *\<^sub>R f i) I = sum (\<lambda>i. u i *\<^sub>R f i) {i \<in> I. u i \<noteq> 0}"
|
|
1173 |
by (force intro: sum.mono_neutral_left simp: supp_sum_def support_on_def) |
|
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1174 |
show ?thesis |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1175 |
apply (simp add: eq) |
| 64267 | 1176 |
apply (rule convex_sum [OF fin \<open>convex S\<close>]) |
1177 |
using 1 assms apply (auto simp: supp_sum_def support_on_def) |
|
|
63969
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1178 |
done |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1179 |
qed |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1180 |
|
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1181 |
lemma convex_translation_eq [simp]: "convex ((\<lambda>x. a + x) ` s) \<longleftrightarrow> convex s" |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1182 |
by (metis convex_translation translation_galois) |
|
f4b4fba60b1d
HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents:
63967
diff
changeset
|
1183 |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
1184 |
lemma convex_linear_image_eq [simp]: |
|
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
1185 |
fixes f :: "'a::real_vector \<Rightarrow> 'b::real_vector" |
|
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
1186 |
shows "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> convex (f ` s) \<longleftrightarrow> convex s" |
|
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
1187 |
by (metis (no_types) convex_linear_image convex_linear_vimage inj_vimage_image_eq) |
|
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
1188 |
|
| 40377 | 1189 |
lemma basis_to_basis_subspace_isomorphism: |
1190 |
assumes s: "subspace (S:: ('n::euclidean_space) set)"
|
|
| 49529 | 1191 |
and t: "subspace (T :: ('m::euclidean_space) set)"
|
1192 |
and d: "dim S = dim T" |
|
| 53333 | 1193 |
and B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" |
1194 |
and C: "C \<subseteq> T" "independent C" "T \<subseteq> span C" "card C = dim T" |
|
1195 |
shows "\<exists>f. linear f \<and> f ` B = C \<and> f ` S = T \<and> inj_on f S" |
|
| 49529 | 1196 |
proof - |
| 53333 | 1197 |
from B independent_bound have fB: "finite B" |
1198 |
by blast |
|
1199 |
from C independent_bound have fC: "finite C" |
|
1200 |
by blast |
|
| 40377 | 1201 |
from B(4) C(4) card_le_inj[of B C] d obtain f where |
| 60420 | 1202 |
f: "f ` B \<subseteq> C" "inj_on f B" using \<open>finite B\<close> \<open>finite C\<close> by auto |
| 40377 | 1203 |
from linear_independent_extend[OF B(2)] obtain g where |
| 53333 | 1204 |
g: "linear g" "\<forall>x \<in> B. g x = f x" by blast |
| 40377 | 1205 |
from inj_on_iff_eq_card[OF fB, of f] f(2) |
1206 |
have "card (f ` B) = card B" by simp |
|
1207 |
with B(4) C(4) have ceq: "card (f ` B) = card C" using d |
|
1208 |
by simp |
|
1209 |
have "g ` B = f ` B" using g(2) |
|
1210 |
by (auto simp add: image_iff) |
|
1211 |
also have "\<dots> = C" using card_subset_eq[OF fC f(1) ceq] . |
|
1212 |
finally have gBC: "g ` B = C" . |
|
1213 |
have gi: "inj_on g B" using f(2) g(2) |
|
1214 |
by (auto simp add: inj_on_def) |
|
1215 |
note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi] |
|
| 53333 | 1216 |
{
|
1217 |
fix x y |
|
| 49529 | 1218 |
assume x: "x \<in> S" and y: "y \<in> S" and gxy: "g x = g y" |
| 53333 | 1219 |
from B(3) x y have x': "x \<in> span B" and y': "y \<in> span B" |
1220 |
by blast+ |
|
1221 |
from gxy have th0: "g (x - y) = 0" |
|
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
1222 |
by (simp add: linear_diff[OF g(1)]) |
| 53333 | 1223 |
have th1: "x - y \<in> span B" using x' y' |
| 63938 | 1224 |
by (metis span_diff) |
| 53333 | 1225 |
have "x = y" using g0[OF th1 th0] by simp |
1226 |
} |
|
1227 |
then have giS: "inj_on g S" unfolding inj_on_def by blast |
|
| 40377 | 1228 |
from span_subspace[OF B(1,3) s] |
| 53333 | 1229 |
have "g ` S = span (g ` B)" |
1230 |
by (simp add: span_linear_image[OF g(1)]) |
|
1231 |
also have "\<dots> = span C" |
|
1232 |
unfolding gBC .. |
|
1233 |
also have "\<dots> = T" |
|
1234 |
using span_subspace[OF C(1,3) t] . |
|
| 40377 | 1235 |
finally have gS: "g ` S = T" . |
| 53333 | 1236 |
from g(1) gS giS gBC show ?thesis |
1237 |
by blast |
|
| 40377 | 1238 |
qed |
1239 |
||
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1240 |
lemma closure_bounded_linear_image_subset: |
| 44524 | 1241 |
assumes f: "bounded_linear f" |
| 53333 | 1242 |
shows "f ` closure S \<subseteq> closure (f ` S)" |
| 44524 | 1243 |
using linear_continuous_on [OF f] closed_closure closure_subset |
1244 |
by (rule image_closure_subset) |
|
1245 |
||
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1246 |
lemma closure_linear_image_subset: |
| 53339 | 1247 |
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::real_normed_vector" |
| 49529 | 1248 |
assumes "linear f" |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1249 |
shows "f ` (closure S) \<subseteq> closure (f ` S)" |
| 44524 | 1250 |
using assms unfolding linear_conv_bounded_linear |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1251 |
by (rule closure_bounded_linear_image_subset) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1252 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1253 |
lemma closed_injective_linear_image: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1254 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1255 |
assumes S: "closed S" and f: "linear f" "inj f" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1256 |
shows "closed (f ` S)" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1257 |
proof - |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1258 |
obtain g where g: "linear g" "g \<circ> f = id" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1259 |
using linear_injective_left_inverse [OF f] by blast |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1260 |
then have confg: "continuous_on (range f) g" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1261 |
using linear_continuous_on linear_conv_bounded_linear by blast |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1262 |
have [simp]: "g ` f ` S = S" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1263 |
using g by (simp add: image_comp) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1264 |
have cgf: "closed (g ` f ` S)" |
| 61808 | 1265 |
by (simp add: \<open>g \<circ> f = id\<close> S image_comp) |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1266 |
have [simp]: "{x \<in> range f. g x \<in> S} = f ` S"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1267 |
using g by (simp add: o_def id_def image_def) metis |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1268 |
show ?thesis |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1269 |
apply (rule closedin_closed_trans [of "range f"]) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1270 |
apply (rule continuous_closedin_preimage [OF confg cgf, simplified]) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1271 |
apply (rule closed_injective_image_subspace) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1272 |
using f |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1273 |
apply (auto simp: linear_linear linear_injective_0) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1274 |
done |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1275 |
qed |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1276 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1277 |
lemma closed_injective_linear_image_eq: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1278 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1279 |
assumes f: "linear f" "inj f" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1280 |
shows "(closed(image f s) \<longleftrightarrow> closed s)" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1281 |
by (metis closed_injective_linear_image closure_eq closure_linear_image_subset closure_subset_eq f(1) f(2) inj_image_subset_iff) |
| 40377 | 1282 |
|
1283 |
lemma closure_injective_linear_image: |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1284 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1285 |
shows "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1286 |
apply (rule subset_antisym) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1287 |
apply (simp add: closure_linear_image_subset) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1288 |
by (simp add: closure_minimal closed_injective_linear_image closure_subset image_mono) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1289 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1290 |
lemma closure_bounded_linear_image: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1291 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1292 |
shows "\<lbrakk>linear f; bounded S\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1293 |
apply (rule subset_antisym, simp add: closure_linear_image_subset) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1294 |
apply (rule closure_minimal, simp add: closure_subset image_mono) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1295 |
by (meson bounded_closure closed_closure compact_continuous_image compact_eq_bounded_closed linear_continuous_on linear_conv_bounded_linear) |
| 40377 | 1296 |
|
| 44524 | 1297 |
lemma closure_scaleR: |
| 53339 | 1298 |
fixes S :: "'a::real_normed_vector set" |
| 44524 | 1299 |
shows "(op *\<^sub>R c) ` (closure S) = closure ((op *\<^sub>R c) ` S)" |
1300 |
proof |
|
1301 |
show "(op *\<^sub>R c) ` (closure S) \<subseteq> closure ((op *\<^sub>R c) ` S)" |
|
| 53333 | 1302 |
using bounded_linear_scaleR_right |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1303 |
by (rule closure_bounded_linear_image_subset) |
| 44524 | 1304 |
show "closure ((op *\<^sub>R c) ` S) \<subseteq> (op *\<^sub>R c) ` (closure S)" |
| 49529 | 1305 |
by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure) |
1306 |
qed |
|
1307 |
||
1308 |
lemma fst_linear: "linear fst" |
|
|
53600
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
huffman
parents:
53406
diff
changeset
|
1309 |
unfolding linear_iff by (simp add: algebra_simps) |
| 49529 | 1310 |
|
1311 |
lemma snd_linear: "linear snd" |
|
|
53600
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
huffman
parents:
53406
diff
changeset
|
1312 |
unfolding linear_iff by (simp add: algebra_simps) |
| 49529 | 1313 |
|
| 54465 | 1314 |
lemma fst_snd_linear: "linear (\<lambda>(x,y). x + y)" |
|
53600
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
huffman
parents:
53406
diff
changeset
|
1315 |
unfolding linear_iff by (simp add: algebra_simps) |
| 40377 | 1316 |
|
| 49529 | 1317 |
lemma vector_choose_size: |
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1318 |
assumes "0 \<le> c" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1319 |
obtains x :: "'a::{real_normed_vector, perfect_space}" where "norm x = c"
|
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1320 |
proof - |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1321 |
obtain a::'a where "a \<noteq> 0" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1322 |
using UNIV_not_singleton UNIV_eq_I set_zero singletonI by fastforce |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1323 |
then show ?thesis |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1324 |
by (rule_tac x="scaleR (c / norm a) a" in that) (simp add: assms) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1325 |
qed |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1326 |
|
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1327 |
lemma vector_choose_dist: |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1328 |
assumes "0 \<le> c" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1329 |
obtains y :: "'a::{real_normed_vector, perfect_space}" where "dist x y = c"
|
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1330 |
by (metis add_diff_cancel_left' assms dist_commute dist_norm vector_choose_size) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1331 |
|
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1332 |
lemma sphere_eq_empty [simp]: |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1333 |
fixes a :: "'a::{real_normed_vector, perfect_space}"
|
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1334 |
shows "sphere a r = {} \<longleftrightarrow> r < 0"
|
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
1335 |
by (auto simp: sphere_def dist_norm) (metis dist_norm le_less_linear vector_choose_dist) |
| 49529 | 1336 |
|
| 64267 | 1337 |
lemma sum_delta_notmem: |
| 49529 | 1338 |
assumes "x \<notin> s" |
| 64267 | 1339 |
shows "sum (\<lambda>y. if (y = x) then P x else Q y) s = sum Q s" |
1340 |
and "sum (\<lambda>y. if (x = y) then P x else Q y) s = sum Q s" |
|
1341 |
and "sum (\<lambda>y. if (y = x) then P y else Q y) s = sum Q s" |
|
1342 |
and "sum (\<lambda>y. if (x = y) then P y else Q y) s = sum Q s" |
|
1343 |
apply (rule_tac [!] sum.cong) |
|
| 53333 | 1344 |
using assms |
1345 |
apply auto |
|
| 49529 | 1346 |
done |
| 33175 | 1347 |
|
| 64267 | 1348 |
lemma sum_delta'': |
| 49529 | 1349 |
fixes s::"'a::real_vector set" |
1350 |
assumes "finite s" |
|
| 33175 | 1351 |
shows "(\<Sum>x\<in>s. (if y = x then f x else 0) *\<^sub>R x) = (if y\<in>s then (f y) *\<^sub>R y else 0)" |
| 49529 | 1352 |
proof - |
1353 |
have *: "\<And>x y. (if y = x then f x else (0::real)) *\<^sub>R x = (if x=y then (f x) *\<^sub>R x else 0)" |
|
1354 |
by auto |
|
1355 |
show ?thesis |
|
| 64267 | 1356 |
unfolding * using sum.delta[OF assms, of y "\<lambda>x. f x *\<^sub>R x"] by auto |
| 33175 | 1357 |
qed |
1358 |
||
| 53333 | 1359 |
lemma if_smult: "(if P then x else (y::real)) *\<^sub>R v = (if P then x *\<^sub>R v else y *\<^sub>R v)" |
| 57418 | 1360 |
by (fact if_distrib) |
| 33175 | 1361 |
|
1362 |
lemma dist_triangle_eq: |
|
|
44361
75ec83d45303
remove unnecessary euclidean_space class constraints
huffman
parents:
44349
diff
changeset
|
1363 |
fixes x y z :: "'a::real_inner" |
| 53333 | 1364 |
shows "dist x z = dist x y + dist y z \<longleftrightarrow> |
1365 |
norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)" |
|
| 49529 | 1366 |
proof - |
1367 |
have *: "x - y + (y - z) = x - z" by auto |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
1368 |
show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *] |
| 49529 | 1369 |
by (auto simp add:norm_minus_commute) |
1370 |
qed |
|
| 33175 | 1371 |
|
| 53406 | 1372 |
lemma norm_minus_eqI: "x = - y \<Longrightarrow> norm x = norm y" by auto |
| 33175 | 1373 |
|
| 49529 | 1374 |
lemma Min_grI: |
1375 |
assumes "finite A" "A \<noteq> {}" "\<forall>a\<in>A. x < a"
|
|
1376 |
shows "x < Min A" |
|
| 33175 | 1377 |
unfolding Min_gr_iff[OF assms(1,2)] using assms(3) by auto |
1378 |
||
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
1379 |
lemma norm_lt: "norm x < norm y \<longleftrightarrow> inner x x < inner y y" |
|
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
1380 |
unfolding norm_eq_sqrt_inner by simp |
| 33175 | 1381 |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
1382 |
lemma norm_le: "norm x \<le> norm y \<longleftrightarrow> inner x x \<le> inner y y" |
|
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
1383 |
unfolding norm_eq_sqrt_inner by simp |
|
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
1384 |
|
|
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
1385 |
|
| 60420 | 1386 |
subsection \<open>Affine set and affine hull\<close> |
| 33175 | 1387 |
|
| 49529 | 1388 |
definition affine :: "'a::real_vector set \<Rightarrow> bool" |
1389 |
where "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s)" |
|
| 33175 | 1390 |
|
1391 |
lemma affine_alt: "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u::real. (1 - u) *\<^sub>R x + u *\<^sub>R y \<in> s)" |
|
| 49529 | 1392 |
unfolding affine_def by (metis eq_diff_eq') |
| 33175 | 1393 |
|
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
1394 |
lemma affine_empty [iff]: "affine {}"
|
| 33175 | 1395 |
unfolding affine_def by auto |
1396 |
||
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
1397 |
lemma affine_sing [iff]: "affine {x}"
|
| 33175 | 1398 |
unfolding affine_alt by (auto simp add: scaleR_left_distrib [symmetric]) |
1399 |
||
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
1400 |
lemma affine_UNIV [iff]: "affine UNIV" |
| 33175 | 1401 |
unfolding affine_def by auto |
1402 |
||
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
1403 |
lemma affine_Inter [intro]: "(\<And>s. s\<in>f \<Longrightarrow> affine s) \<Longrightarrow> affine (\<Inter>f)" |
| 49531 | 1404 |
unfolding affine_def by auto |
| 33175 | 1405 |
|
| 60303 | 1406 |
lemma affine_Int[intro]: "affine s \<Longrightarrow> affine t \<Longrightarrow> affine (s \<inter> t)" |
| 33175 | 1407 |
unfolding affine_def by auto |
1408 |
||
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
1409 |
lemma affine_scaling: "affine s \<Longrightarrow> affine (image (\<lambda>x. c *\<^sub>R x) s)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
1410 |
apply (clarsimp simp add: affine_def) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
1411 |
apply (rule_tac x="u *\<^sub>R x + v *\<^sub>R y" in image_eqI) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
1412 |
apply (auto simp: algebra_simps) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
1413 |
done |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
1414 |
|
| 60303 | 1415 |
lemma affine_affine_hull [simp]: "affine(affine hull s)" |
| 49529 | 1416 |
unfolding hull_def |
1417 |
using affine_Inter[of "{t. affine t \<and> s \<subseteq> t}"] by auto
|
|
| 33175 | 1418 |
|
1419 |
lemma affine_hull_eq[simp]: "(affine hull s = s) \<longleftrightarrow> affine s" |
|
| 49529 | 1420 |
by (metis affine_affine_hull hull_same) |
1421 |
||
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
1422 |
lemma affine_hyperplane: "affine {x. a \<bullet> x = b}"
|
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
1423 |
by (simp add: affine_def algebra_simps) (metis distrib_right mult.left_neutral) |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
1424 |
|
| 33175 | 1425 |
|
| 60420 | 1426 |
subsubsection \<open>Some explicit formulations (from Lars Schewe)\<close> |
| 33175 | 1427 |
|
| 49529 | 1428 |
lemma affine: |
1429 |
fixes V::"'a::real_vector set" |
|
1430 |
shows "affine V \<longleftrightarrow> |
|
| 64267 | 1431 |
(\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> sum u s = 1 \<longrightarrow> (sum (\<lambda>x. (u x) *\<^sub>R x)) s \<in> V)"
|
| 49529 | 1432 |
unfolding affine_def |
1433 |
apply rule |
|
1434 |
apply(rule, rule, rule) |
|
| 49531 | 1435 |
apply(erule conjE)+ |
| 49529 | 1436 |
defer |
1437 |
apply (rule, rule, rule, rule, rule) |
|
1438 |
proof - |
|
1439 |
fix x y u v |
|
1440 |
assume as: "x \<in> V" "y \<in> V" "u + v = (1::real)" |
|
| 64267 | 1441 |
"\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> sum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
|
| 49529 | 1442 |
then show "u *\<^sub>R x + v *\<^sub>R y \<in> V" |
1443 |
apply (cases "x = y") |
|
1444 |
using as(4)[THEN spec[where x="{x,y}"], THEN spec[where x="\<lambda>w. if w = x then u else v"]]
|
|
1445 |
and as(1-3) |
|
| 53333 | 1446 |
apply (auto simp add: scaleR_left_distrib[symmetric]) |
1447 |
done |
|
| 33175 | 1448 |
next |
| 49529 | 1449 |
fix s u |
1450 |
assume as: "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V" |
|
| 64267 | 1451 |
"finite s" "s \<noteq> {}" "s \<subseteq> V" "sum u s = (1::real)"
|
| 63040 | 1452 |
define n where "n = card s" |
| 33175 | 1453 |
have "card s = 0 \<or> card s = 1 \<or> card s = 2 \<or> card s > 2" by auto |
| 49529 | 1454 |
then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" |
1455 |
proof (auto simp only: disjE) |
|
1456 |
assume "card s = 2" |
|
| 53333 | 1457 |
then have "card s = Suc (Suc 0)" |
1458 |
by auto |
|
1459 |
then obtain a b where "s = {a, b}"
|
|
1460 |
unfolding card_Suc_eq by auto |
|
| 49529 | 1461 |
then show ?thesis |
1462 |
using as(1)[THEN bspec[where x=a], THEN bspec[where x=b]] using as(4,5) |
|
| 64267 | 1463 |
by (auto simp add: sum_clauses(2)) |
| 49529 | 1464 |
next |
1465 |
assume "card s > 2" |
|
1466 |
then show ?thesis using as and n_def |
|
1467 |
proof (induct n arbitrary: u s) |
|
1468 |
case 0 |
|
1469 |
then show ?case by auto |
|
1470 |
next |
|
1471 |
case (Suc n) |
|
1472 |
fix s :: "'a set" and u :: "'a \<Rightarrow> real" |
|
1473 |
assume IA: |
|
1474 |
"\<And>u s. \<lbrakk>2 < card s; \<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V; finite s; |
|
| 64267 | 1475 |
s \<noteq> {}; s \<subseteq> V; sum u s = 1; n = card s \<rbrakk> \<Longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
|
| 49529 | 1476 |
and as: |
1477 |
"Suc n = card s" "2 < card s" "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V" |
|
| 64267 | 1478 |
"finite s" "s \<noteq> {}" "s \<subseteq> V" "sum u s = 1"
|
| 49529 | 1479 |
have "\<exists>x\<in>s. u x \<noteq> 1" |
1480 |
proof (rule ccontr) |
|
1481 |
assume "\<not> ?thesis" |
|
| 64267 | 1482 |
then have "sum u s = real_of_nat (card s)" |
1483 |
unfolding card_eq_sum by auto |
|
| 49529 | 1484 |
then show False |
| 60420 | 1485 |
using as(7) and \<open>card s > 2\<close> |
| 49529 | 1486 |
by (metis One_nat_def less_Suc0 Zero_not_Suc of_nat_1 of_nat_eq_iff numeral_2_eq_2) |
|
45498
2dc373f1867a
avoid numeral-representation-specific rules in metis proof
huffman
parents:
45051
diff
changeset
|
1487 |
qed |
| 53339 | 1488 |
then obtain x where x:"x \<in> s" "u x \<noteq> 1" by auto |
| 33175 | 1489 |
|
| 49529 | 1490 |
have c: "card (s - {x}) = card s - 1"
|
| 53333 | 1491 |
apply (rule card_Diff_singleton) |
| 60420 | 1492 |
using \<open>x\<in>s\<close> as(4) |
| 53333 | 1493 |
apply auto |
1494 |
done |
|
| 49529 | 1495 |
have *: "s = insert x (s - {x})" "finite (s - {x})"
|
| 60420 | 1496 |
using \<open>x\<in>s\<close> and as(4) by auto |
| 64267 | 1497 |
have **: "sum u (s - {x}) = 1 - u x"
|
1498 |
using sum_clauses(2)[OF *(2), of u x, unfolded *(1)[symmetric] as(7)] by auto |
|
1499 |
have ***: "inverse (1 - u x) * sum u (s - {x}) = 1"
|
|
| 60420 | 1500 |
unfolding ** using \<open>u x \<noteq> 1\<close> by auto |
| 49529 | 1501 |
have "(\<Sum>xa\<in>s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \<in> V"
|
1502 |
proof (cases "card (s - {x}) > 2")
|
|
1503 |
case True |
|
1504 |
then have "s - {x} \<noteq> {}" "card (s - {x}) = n"
|
|
1505 |
unfolding c and as(1)[symmetric] |
|
| 49531 | 1506 |
proof (rule_tac ccontr) |
| 49529 | 1507 |
assume "\<not> s - {x} \<noteq> {}"
|
| 49531 | 1508 |
then have "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp
|
| 49529 | 1509 |
then show False using True by auto |
1510 |
qed auto |
|
1511 |
then show ?thesis |
|
1512 |
apply (rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"])
|
|
| 64267 | 1513 |
unfolding sum_distrib_left[symmetric] |
| 53333 | 1514 |
using as and *** and True |
| 49529 | 1515 |
apply auto |
1516 |
done |
|
1517 |
next |
|
1518 |
case False |
|
| 53333 | 1519 |
then have "card (s - {x}) = Suc (Suc 0)"
|
1520 |
using as(2) and c by auto |
|
1521 |
then obtain a b where "(s - {x}) = {a, b}" "a\<noteq>b"
|
|
1522 |
unfolding card_Suc_eq by auto |
|
1523 |
then show ?thesis |
|
1524 |
using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]] |
|
| 60420 | 1525 |
using *** *(2) and \<open>s \<subseteq> V\<close> |
| 64267 | 1526 |
unfolding sum_distrib_left |
1527 |
by (auto simp add: sum_clauses(2)) |
|
| 49529 | 1528 |
qed |
1529 |
then have "u x + (1 - u x) = 1 \<Longrightarrow> |
|
1530 |
u x *\<^sub>R x + (1 - u x) *\<^sub>R ((\<Sum>xa\<in>s - {x}. u xa *\<^sub>R xa) /\<^sub>R (1 - u x)) \<in> V"
|
|
1531 |
apply - |
|
1532 |
apply (rule as(3)[rule_format]) |
|
| 64267 | 1533 |
unfolding Real_Vector_Spaces.scaleR_right.sum |
| 53333 | 1534 |
using x(1) as(6) |
1535 |
apply auto |
|
| 49529 | 1536 |
done |
1537 |
then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V" |
|
| 64267 | 1538 |
unfolding scaleR_scaleR[symmetric] and scaleR_right.sum [symmetric] |
| 49529 | 1539 |
apply (subst *) |
| 64267 | 1540 |
unfolding sum_clauses(2)[OF *(2)] |
| 60420 | 1541 |
using \<open>u x \<noteq> 1\<close> |
| 53333 | 1542 |
apply auto |
| 49529 | 1543 |
done |
1544 |
qed |
|
1545 |
next |
|
1546 |
assume "card s = 1" |
|
| 53333 | 1547 |
then obtain a where "s={a}"
|
1548 |
by (auto simp add: card_Suc_eq) |
|
1549 |
then show ?thesis |
|
1550 |
using as(4,5) by simp |
|
| 60420 | 1551 |
qed (insert \<open>s\<noteq>{}\<close> \<open>finite s\<close>, auto)
|
| 33175 | 1552 |
qed |
1553 |
||
1554 |
lemma affine_hull_explicit: |
|
| 53333 | 1555 |
"affine hull p = |
| 64267 | 1556 |
{y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> sum (\<lambda>v. (u v) *\<^sub>R v) s = y}"
|
| 49529 | 1557 |
apply (rule hull_unique) |
1558 |
apply (subst subset_eq) |
|
1559 |
prefer 3 |
|
1560 |
apply rule |
|
1561 |
unfolding mem_Collect_eq |
|
1562 |
apply (erule exE)+ |
|
1563 |
apply (erule conjE)+ |
|
1564 |
prefer 2 |
|
1565 |
apply rule |
|
1566 |
proof - |
|
1567 |
fix x |
|
1568 |
assume "x\<in>p" |
|
| 64267 | 1569 |
then show "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
|
| 53333 | 1570 |
apply (rule_tac x="{x}" in exI)
|
1571 |
apply (rule_tac x="\<lambda>x. 1" in exI) |
|
| 49529 | 1572 |
apply auto |
1573 |
done |
|
| 33175 | 1574 |
next |
| 49529 | 1575 |
fix t x s u |
| 53333 | 1576 |
assume as: "p \<subseteq> t" "affine t" "finite s" "s \<noteq> {}"
|
| 64267 | 1577 |
"s \<subseteq> p" "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x" |
| 49529 | 1578 |
then show "x \<in> t" |
| 53333 | 1579 |
using as(2)[unfolded affine, THEN spec[where x=s], THEN spec[where x=u]] |
1580 |
by auto |
|
| 33175 | 1581 |
next |
| 64267 | 1582 |
show "affine {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y}"
|
| 49529 | 1583 |
unfolding affine_def |
1584 |
apply (rule, rule, rule, rule, rule) |
|
1585 |
unfolding mem_Collect_eq |
|
1586 |
proof - |
|
1587 |
fix u v :: real |
|
1588 |
assume uv: "u + v = 1" |
|
1589 |
fix x |
|
| 64267 | 1590 |
assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
|
| 49529 | 1591 |
then obtain sx ux where |
| 64267 | 1592 |
x: "finite sx" "sx \<noteq> {}" "sx \<subseteq> p" "sum ux sx = 1" "(\<Sum>v\<in>sx. ux v *\<^sub>R v) = x"
|
| 53333 | 1593 |
by auto |
1594 |
fix y |
|
| 64267 | 1595 |
assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
|
| 49529 | 1596 |
then obtain sy uy where |
| 64267 | 1597 |
y: "finite sy" "sy \<noteq> {}" "sy \<subseteq> p" "sum uy sy = 1" "(\<Sum>v\<in>sy. uy v *\<^sub>R v) = y" by auto
|
| 53333 | 1598 |
have xy: "finite (sx \<union> sy)" |
1599 |
using x(1) y(1) by auto |
|
1600 |
have **: "(sx \<union> sy) \<inter> sx = sx" "(sx \<union> sy) \<inter> sy = sy" |
|
1601 |
by auto |
|
| 49529 | 1602 |
show "\<exists>s ua. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and>
|
| 64267 | 1603 |
sum ua s = 1 \<and> (\<Sum>v\<in>s. ua v *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y" |
| 49529 | 1604 |
apply (rule_tac x="sx \<union> sy" in exI) |
1605 |
apply (rule_tac x="\<lambda>a. (if a\<in>sx then u * ux a else 0) + (if a\<in>sy then v * uy a else 0)" in exI) |
|
| 64267 | 1606 |
unfolding scaleR_left_distrib sum.distrib if_smult scaleR_zero_left |
1607 |
** sum.inter_restrict[OF xy, symmetric] |
|
1608 |
unfolding scaleR_scaleR[symmetric] Real_Vector_Spaces.scaleR_right.sum [symmetric] |
|
1609 |
and sum_distrib_left[symmetric] |
|
| 49529 | 1610 |
unfolding x y |
| 53333 | 1611 |
using x(1-3) y(1-3) uv |
1612 |
apply simp |
|
| 49529 | 1613 |
done |
1614 |
qed |
|
1615 |
qed |
|
| 33175 | 1616 |
|
1617 |
lemma affine_hull_finite: |
|
1618 |
assumes "finite s" |
|
| 64267 | 1619 |
shows "affine hull s = {y. \<exists>u. sum u s = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) s = y}"
|
| 53333 | 1620 |
unfolding affine_hull_explicit and set_eq_iff and mem_Collect_eq |
1621 |
apply (rule, rule) |
|
1622 |
apply (erule exE)+ |
|
1623 |
apply (erule conjE)+ |
|
| 49529 | 1624 |
defer |
1625 |
apply (erule exE) |
|
1626 |
apply (erule conjE) |
|
1627 |
proof - |
|
1628 |
fix x u |
|
| 64267 | 1629 |
assume "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x" |
| 49529 | 1630 |
then show "\<exists>sa u. finite sa \<and> |
| 64267 | 1631 |
\<not> (\<forall>x. (x \<in> sa) = (x \<in> {})) \<and> sa \<subseteq> s \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = x"
|
| 49529 | 1632 |
apply (rule_tac x=s in exI, rule_tac x=u in exI) |
| 53333 | 1633 |
using assms |
1634 |
apply auto |
|
| 49529 | 1635 |
done |
| 33175 | 1636 |
next |
| 49529 | 1637 |
fix x t u |
1638 |
assume "t \<subseteq> s" |
|
| 53333 | 1639 |
then have *: "s \<inter> t = t" |
1640 |
by auto |
|
| 64267 | 1641 |
assume "finite t" "\<not> (\<forall>x. (x \<in> t) = (x \<in> {}))" "sum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
|
1642 |
then show "\<exists>u. sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x" |
|
| 49529 | 1643 |
apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI) |
| 64267 | 1644 |
unfolding if_smult scaleR_zero_left and sum.inter_restrict[OF assms, symmetric] and * |
| 49529 | 1645 |
apply auto |
1646 |
done |
|
1647 |
qed |
|
1648 |
||
| 33175 | 1649 |
|
| 60420 | 1650 |
subsubsection \<open>Stepping theorems and hence small special cases\<close> |
| 33175 | 1651 |
|
1652 |
lemma affine_hull_empty[simp]: "affine hull {} = {}"
|
|
| 49529 | 1653 |
by (rule hull_unique) auto |
| 33175 | 1654 |
|
| 64267 | 1655 |
(*could delete: it simply rewrites sum expressions, but it's used twice*) |
| 33175 | 1656 |
lemma affine_hull_finite_step: |
1657 |
fixes y :: "'a::real_vector" |
|
| 49529 | 1658 |
shows |
| 64267 | 1659 |
"(\<exists>u. sum u {} = w \<and> sum (\<lambda>x. u x *\<^sub>R x) {} = y) \<longleftrightarrow> w = 0 \<and> y = 0" (is ?th1)
|
| 53347 | 1660 |
and |
| 49529 | 1661 |
"finite s \<Longrightarrow> |
| 64267 | 1662 |
(\<exists>u. sum u (insert a s) = w \<and> sum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y) \<longleftrightarrow> |
1663 |
(\<exists>v u. sum u s = w - v \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" (is "_ \<Longrightarrow> ?lhs = ?rhs") |
|
| 49529 | 1664 |
proof - |
| 33175 | 1665 |
show ?th1 by simp |
| 53347 | 1666 |
assume fin: "finite s" |
1667 |
show "?lhs = ?rhs" |
|
1668 |
proof |
|
| 53302 | 1669 |
assume ?lhs |
| 64267 | 1670 |
then obtain u where u: "sum u (insert a s) = w \<and> (\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y" |
| 53302 | 1671 |
by auto |
| 53347 | 1672 |
show ?rhs |
| 49529 | 1673 |
proof (cases "a \<in> s") |
1674 |
case True |
|
1675 |
then have *: "insert a s = s" by auto |
|
| 53302 | 1676 |
show ?thesis |
1677 |
using u[unfolded *] |
|
1678 |
apply(rule_tac x=0 in exI) |
|
1679 |
apply auto |
|
1680 |
done |
|
| 33175 | 1681 |
next |
| 49529 | 1682 |
case False |
1683 |
then show ?thesis |
|
1684 |
apply (rule_tac x="u a" in exI) |
|
| 53347 | 1685 |
using u and fin |
| 53302 | 1686 |
apply auto |
| 49529 | 1687 |
done |
| 53302 | 1688 |
qed |
| 53347 | 1689 |
next |
| 53302 | 1690 |
assume ?rhs |
| 64267 | 1691 |
then obtain v u where vu: "sum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a" |
| 53302 | 1692 |
by auto |
1693 |
have *: "\<And>x M. (if x = a then v else M) *\<^sub>R x = (if x = a then v *\<^sub>R x else M *\<^sub>R x)" |
|
1694 |
by auto |
|
| 53347 | 1695 |
show ?lhs |
| 49529 | 1696 |
proof (cases "a \<in> s") |
1697 |
case True |
|
1698 |
then show ?thesis |
|
1699 |
apply (rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI) |
|
| 64267 | 1700 |
unfolding sum_clauses(2)[OF fin] |
| 53333 | 1701 |
apply simp |
| 64267 | 1702 |
unfolding scaleR_left_distrib and sum.distrib |
| 33175 | 1703 |
unfolding vu and * and scaleR_zero_left |
| 64267 | 1704 |
apply (auto simp add: sum.delta[OF fin]) |
| 49529 | 1705 |
done |
| 33175 | 1706 |
next |
| 49531 | 1707 |
case False |
| 49529 | 1708 |
then have **: |
1709 |
"\<And>x. x \<in> s \<Longrightarrow> u x = (if x = a then v else u x)" |
|
1710 |
"\<And>x. x \<in> s \<Longrightarrow> u x *\<^sub>R x = (if x = a then v *\<^sub>R x else u x *\<^sub>R x)" by auto |
|
| 33175 | 1711 |
from False show ?thesis |
| 49529 | 1712 |
apply (rule_tac x="\<lambda>x. if x=a then v else u x" in exI) |
| 64267 | 1713 |
unfolding sum_clauses(2)[OF fin] and * using vu |
1714 |
using sum.cong [of s _ "\<lambda>x. u x *\<^sub>R x" "\<lambda>x. if x = a then v *\<^sub>R x else u x *\<^sub>R x", OF _ **(2)] |
|
1715 |
using sum.cong [of s _ u "\<lambda>x. if x = a then v else u x", OF _ **(1)] |
|
| 49529 | 1716 |
apply auto |
1717 |
done |
|
1718 |
qed |
|
| 53347 | 1719 |
qed |
| 33175 | 1720 |
qed |
1721 |
||
1722 |
lemma affine_hull_2: |
|
1723 |
fixes a b :: "'a::real_vector" |
|
| 53302 | 1724 |
shows "affine hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b| u v. (u + v = 1)}"
|
1725 |
(is "?lhs = ?rhs") |
|
| 49529 | 1726 |
proof - |
1727 |
have *: |
|
| 49531 | 1728 |
"\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)" |
| 49529 | 1729 |
"\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto |
| 64267 | 1730 |
have "?lhs = {y. \<exists>u. sum u {a, b} = 1 \<and> (\<Sum>v\<in>{a, b}. u v *\<^sub>R v) = y}"
|
| 33175 | 1731 |
using affine_hull_finite[of "{a,b}"] by auto
|
1732 |
also have "\<dots> = {y. \<exists>v u. u b = 1 - v \<and> u b *\<^sub>R b = y - v *\<^sub>R a}"
|
|
| 49529 | 1733 |
by (simp add: affine_hull_finite_step(2)[of "{b}" a])
|
| 33175 | 1734 |
also have "\<dots> = ?rhs" unfolding * by auto |
1735 |
finally show ?thesis by auto |
|
1736 |
qed |
|
1737 |
||
1738 |
lemma affine_hull_3: |
|
1739 |
fixes a b c :: "'a::real_vector" |
|
| 53302 | 1740 |
shows "affine hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c| u v w. u + v + w = 1}"
|
| 49529 | 1741 |
proof - |
1742 |
have *: |
|
| 49531 | 1743 |
"\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)" |
| 49529 | 1744 |
"\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto |
1745 |
show ?thesis |
|
1746 |
apply (simp add: affine_hull_finite affine_hull_finite_step) |
|
1747 |
unfolding * |
|
1748 |
apply auto |
|
| 53302 | 1749 |
apply (rule_tac x=v in exI) |
1750 |
apply (rule_tac x=va in exI) |
|
1751 |
apply auto |
|
1752 |
apply (rule_tac x=u in exI) |
|
1753 |
apply force |
|
| 49529 | 1754 |
done |
| 33175 | 1755 |
qed |
1756 |
||
| 40377 | 1757 |
lemma mem_affine: |
| 53333 | 1758 |
assumes "affine S" "x \<in> S" "y \<in> S" "u + v = 1" |
| 53347 | 1759 |
shows "u *\<^sub>R x + v *\<^sub>R y \<in> S" |
| 40377 | 1760 |
using assms affine_def[of S] by auto |
1761 |
||
1762 |
lemma mem_affine_3: |
|
| 53333 | 1763 |
assumes "affine S" "x \<in> S" "y \<in> S" "z \<in> S" "u + v + w = 1" |
| 53347 | 1764 |
shows "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \<in> S" |
| 49529 | 1765 |
proof - |
| 53347 | 1766 |
have "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \<in> affine hull {x, y, z}"
|
| 49529 | 1767 |
using affine_hull_3[of x y z] assms by auto |
1768 |
moreover |
|
| 53347 | 1769 |
have "affine hull {x, y, z} \<subseteq> affine hull S"
|
| 49529 | 1770 |
using hull_mono[of "{x, y, z}" "S"] assms by auto
|
1771 |
moreover |
|
| 53347 | 1772 |
have "affine hull S = S" |
1773 |
using assms affine_hull_eq[of S] by auto |
|
| 49531 | 1774 |
ultimately show ?thesis by auto |
| 40377 | 1775 |
qed |
1776 |
||
1777 |
lemma mem_affine_3_minus: |
|
| 53333 | 1778 |
assumes "affine S" "x \<in> S" "y \<in> S" "z \<in> S" |
1779 |
shows "x + v *\<^sub>R (y-z) \<in> S" |
|
1780 |
using mem_affine_3[of S x y z 1 v "-v"] assms |
|
1781 |
by (simp add: algebra_simps) |
|
| 40377 | 1782 |
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
1783 |
corollary mem_affine_3_minus2: |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
1784 |
"\<lbrakk>affine S; x \<in> S; y \<in> S; z \<in> S\<rbrakk> \<Longrightarrow> x - v *\<^sub>R (y-z) \<in> S" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
1785 |
by (metis add_uminus_conv_diff mem_affine_3_minus real_vector.scale_minus_left) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
1786 |
|
| 40377 | 1787 |
|
| 60420 | 1788 |
subsubsection \<open>Some relations between affine hull and subspaces\<close> |
| 33175 | 1789 |
|
1790 |
lemma affine_hull_insert_subset_span: |
|
| 49529 | 1791 |
"affine hull (insert a s) \<subseteq> {a + v| v . v \<in> span {x - a | x . x \<in> s}}"
|
1792 |
unfolding subset_eq Ball_def |
|
1793 |
unfolding affine_hull_explicit span_explicit mem_Collect_eq |
|
| 50804 | 1794 |
apply (rule, rule) |
1795 |
apply (erule exE)+ |
|
1796 |
apply (erule conjE)+ |
|
| 49529 | 1797 |
proof - |
1798 |
fix x t u |
|
| 64267 | 1799 |
assume as: "finite t" "t \<noteq> {}" "t \<subseteq> insert a s" "sum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
|
| 53333 | 1800 |
have "(\<lambda>x. x - a) ` (t - {a}) \<subseteq> {x - a |x. x \<in> s}"
|
1801 |
using as(3) by auto |
|
| 49529 | 1802 |
then show "\<exists>v. x = a + v \<and> (\<exists>S u. finite S \<and> S \<subseteq> {x - a |x. x \<in> s} \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = v)"
|
1803 |
apply (rule_tac x="x - a" in exI) |
|
| 33175 | 1804 |
apply (rule conjI, simp) |
| 49529 | 1805 |
apply (rule_tac x="(\<lambda>x. x - a) ` (t - {a})" in exI)
|
1806 |
apply (rule_tac x="\<lambda>x. u (x + a)" in exI) |
|
| 33175 | 1807 |
apply (rule conjI) using as(1) apply simp |
1808 |
apply (erule conjI) |
|
1809 |
using as(1) |
|
| 64267 | 1810 |
apply (simp add: sum.reindex[unfolded inj_on_def] scaleR_right_diff_distrib |
1811 |
sum_subtractf scaleR_left.sum[symmetric] sum_diff1 scaleR_left_diff_distrib) |
|
| 49529 | 1812 |
unfolding as |
1813 |
apply simp |
|
1814 |
done |
|
1815 |
qed |
|
| 33175 | 1816 |
|
1817 |
lemma affine_hull_insert_span: |
|
1818 |
assumes "a \<notin> s" |
|
| 49529 | 1819 |
shows "affine hull (insert a s) = {a + v | v . v \<in> span {x - a | x. x \<in> s}}"
|
1820 |
apply (rule, rule affine_hull_insert_subset_span) |
|
1821 |
unfolding subset_eq Ball_def |
|
1822 |
unfolding affine_hull_explicit and mem_Collect_eq |
|
1823 |
proof (rule, rule, erule exE, erule conjE) |
|
| 49531 | 1824 |
fix y v |
| 49529 | 1825 |
assume "y = a + v" "v \<in> span {x - a |x. x \<in> s}"
|
| 53339 | 1826 |
then obtain t u where obt: "finite t" "t \<subseteq> {x - a |x. x \<in> s}" "a + (\<Sum>v\<in>t. u v *\<^sub>R v) = y"
|
| 49529 | 1827 |
unfolding span_explicit by auto |
| 63040 | 1828 |
define f where "f = (\<lambda>x. x + a) ` t" |
| 53333 | 1829 |
have f: "finite f" "f \<subseteq> s" "(\<Sum>v\<in>f. u (v - a) *\<^sub>R (v - a)) = y - a" |
| 64267 | 1830 |
unfolding f_def using obt by (auto simp add: sum.reindex[unfolded inj_on_def]) |
| 53333 | 1831 |
have *: "f \<inter> {a} = {}" "f \<inter> - {a} = f"
|
1832 |
using f(2) assms by auto |
|
| 64267 | 1833 |
show "\<exists>sa u. finite sa \<and> sa \<noteq> {} \<and> sa \<subseteq> insert a s \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
|
| 49529 | 1834 |
apply (rule_tac x = "insert a f" in exI) |
| 64267 | 1835 |
apply (rule_tac x = "\<lambda>x. if x=a then 1 - sum (\<lambda>x. u (x - a)) f else u (x - a)" in exI) |
| 53339 | 1836 |
using assms and f |
| 64267 | 1837 |
unfolding sum_clauses(2)[OF f(1)] and if_smult |
1838 |
unfolding sum.If_cases[OF f(1), of "\<lambda>x. x = a"] |
|
1839 |
apply (auto simp add: sum_subtractf scaleR_left.sum algebra_simps *) |
|
| 49529 | 1840 |
done |
1841 |
qed |
|
| 33175 | 1842 |
|
1843 |
lemma affine_hull_span: |
|
1844 |
assumes "a \<in> s" |
|
1845 |
shows "affine hull s = {a + v | v. v \<in> span {x - a | x. x \<in> s - {a}}}"
|
|
1846 |
using affine_hull_insert_span[of a "s - {a}", unfolded insert_Diff[OF assms]] by auto
|
|
1847 |
||
| 49529 | 1848 |
|
| 60420 | 1849 |
subsubsection \<open>Parallel affine sets\<close> |
| 40377 | 1850 |
|
| 53347 | 1851 |
definition affine_parallel :: "'a::real_vector set \<Rightarrow> 'a::real_vector set \<Rightarrow> bool" |
| 53339 | 1852 |
where "affine_parallel S T \<longleftrightarrow> (\<exists>a. T = (\<lambda>x. a + x) ` S)" |
| 40377 | 1853 |
|
1854 |
lemma affine_parallel_expl_aux: |
|
| 49529 | 1855 |
fixes S T :: "'a::real_vector set" |
| 53339 | 1856 |
assumes "\<forall>x. x \<in> S \<longleftrightarrow> a + x \<in> T" |
1857 |
shows "T = (\<lambda>x. a + x) ` S" |
|
| 49529 | 1858 |
proof - |
| 53302 | 1859 |
{
|
1860 |
fix x |
|
| 53339 | 1861 |
assume "x \<in> T" |
1862 |
then have "( - a) + x \<in> S" |
|
1863 |
using assms by auto |
|
1864 |
then have "x \<in> ((\<lambda>x. a + x) ` S)" |
|
| 53333 | 1865 |
using imageI[of "-a+x" S "(\<lambda>x. a+x)"] by auto |
| 53302 | 1866 |
} |
| 53339 | 1867 |
moreover have "T \<ge> (\<lambda>x. a + x) ` S" |
| 53333 | 1868 |
using assms by auto |
| 49529 | 1869 |
ultimately show ?thesis by auto |
1870 |
qed |
|
1871 |
||
| 53339 | 1872 |
lemma affine_parallel_expl: "affine_parallel S T \<longleftrightarrow> (\<exists>a. \<forall>x. x \<in> S \<longleftrightarrow> a + x \<in> T)" |
| 49529 | 1873 |
unfolding affine_parallel_def |
1874 |
using affine_parallel_expl_aux[of S _ T] by auto |
|
1875 |
||
1876 |
lemma affine_parallel_reflex: "affine_parallel S S" |
|
| 53302 | 1877 |
unfolding affine_parallel_def |
1878 |
apply (rule exI[of _ "0"]) |
|
1879 |
apply auto |
|
1880 |
done |
|
| 40377 | 1881 |
|
1882 |
lemma affine_parallel_commut: |
|
| 49529 | 1883 |
assumes "affine_parallel A B" |
1884 |
shows "affine_parallel B A" |
|
1885 |
proof - |
|
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53676
diff
changeset
|
1886 |
from assms obtain a where B: "B = (\<lambda>x. a + x) ` A" |
| 49529 | 1887 |
unfolding affine_parallel_def by auto |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53676
diff
changeset
|
1888 |
have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff) |
|
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53676
diff
changeset
|
1889 |
from B show ?thesis |
| 53333 | 1890 |
using translation_galois [of B a A] |
1891 |
unfolding affine_parallel_def by auto |
|
| 40377 | 1892 |
qed |
1893 |
||
1894 |
lemma affine_parallel_assoc: |
|
| 53339 | 1895 |
assumes "affine_parallel A B" |
1896 |
and "affine_parallel B C" |
|
| 49531 | 1897 |
shows "affine_parallel A C" |
| 49529 | 1898 |
proof - |
| 53333 | 1899 |
from assms obtain ab where "B = (\<lambda>x. ab + x) ` A" |
| 49531 | 1900 |
unfolding affine_parallel_def by auto |
1901 |
moreover |
|
| 53333 | 1902 |
from assms obtain bc where "C = (\<lambda>x. bc + x) ` B" |
| 49529 | 1903 |
unfolding affine_parallel_def by auto |
1904 |
ultimately show ?thesis |
|
1905 |
using translation_assoc[of bc ab A] unfolding affine_parallel_def by auto |
|
| 40377 | 1906 |
qed |
1907 |
||
1908 |
lemma affine_translation_aux: |
|
1909 |
fixes a :: "'a::real_vector" |
|
| 53333 | 1910 |
assumes "affine ((\<lambda>x. a + x) ` S)" |
1911 |
shows "affine S" |
|
| 53302 | 1912 |
proof - |
1913 |
{
|
|
1914 |
fix x y u v |
|
| 53333 | 1915 |
assume xy: "x \<in> S" "y \<in> S" "(u :: real) + v = 1" |
1916 |
then have "(a + x) \<in> ((\<lambda>x. a + x) ` S)" "(a + y) \<in> ((\<lambda>x. a + x) ` S)" |
|
1917 |
by auto |
|
| 53339 | 1918 |
then have h1: "u *\<^sub>R (a + x) + v *\<^sub>R (a + y) \<in> (\<lambda>x. a + x) ` S" |
| 49529 | 1919 |
using xy assms unfolding affine_def by auto |
| 53339 | 1920 |
have "u *\<^sub>R (a + x) + v *\<^sub>R (a + y) = (u + v) *\<^sub>R a + (u *\<^sub>R x + v *\<^sub>R y)" |
| 49529 | 1921 |
by (simp add: algebra_simps) |
| 53339 | 1922 |
also have "\<dots> = a + (u *\<^sub>R x + v *\<^sub>R y)" |
| 60420 | 1923 |
using \<open>u + v = 1\<close> by auto |
| 53339 | 1924 |
ultimately have "a + (u *\<^sub>R x + v *\<^sub>R y) \<in> (\<lambda>x. a + x) ` S" |
| 53333 | 1925 |
using h1 by auto |
| 49529 | 1926 |
then have "u *\<^sub>R x + v *\<^sub>R y : S" by auto |
1927 |
} |
|
1928 |
then show ?thesis unfolding affine_def by auto |
|
| 40377 | 1929 |
qed |
1930 |
||
1931 |
lemma affine_translation: |
|
1932 |
fixes a :: "'a::real_vector" |
|
| 53339 | 1933 |
shows "affine S \<longleftrightarrow> affine ((\<lambda>x. a + x) ` S)" |
| 49529 | 1934 |
proof - |
| 53339 | 1935 |
have "affine S \<Longrightarrow> affine ((\<lambda>x. a + x) ` S)" |
1936 |
using affine_translation_aux[of "-a" "((\<lambda>x. a + x) ` S)"] |
|
| 49529 | 1937 |
using translation_assoc[of "-a" a S] by auto |
1938 |
then show ?thesis using affine_translation_aux by auto |
|
| 40377 | 1939 |
qed |
1940 |
||
1941 |
lemma parallel_is_affine: |
|
| 49529 | 1942 |
fixes S T :: "'a::real_vector set" |
1943 |
assumes "affine S" "affine_parallel S T" |
|
1944 |
shows "affine T" |
|
1945 |
proof - |
|
| 53339 | 1946 |
from assms obtain a where "T = (\<lambda>x. a + x) ` S" |
| 49531 | 1947 |
unfolding affine_parallel_def by auto |
| 53339 | 1948 |
then show ?thesis |
1949 |
using affine_translation assms by auto |
|
| 40377 | 1950 |
qed |
1951 |
||
|
44361
75ec83d45303
remove unnecessary euclidean_space class constraints
huffman
parents:
44349
diff
changeset
|
1952 |
lemma subspace_imp_affine: "subspace s \<Longrightarrow> affine s" |
| 40377 | 1953 |
unfolding subspace_def affine_def by auto |
1954 |
||
| 49529 | 1955 |
|
| 60420 | 1956 |
subsubsection \<open>Subspace parallel to an affine set\<close> |
| 40377 | 1957 |
|
| 53339 | 1958 |
lemma subspace_affine: "subspace S \<longleftrightarrow> affine S \<and> 0 \<in> S" |
| 49529 | 1959 |
proof - |
| 53333 | 1960 |
have h0: "subspace S \<Longrightarrow> affine S \<and> 0 \<in> S" |
| 49529 | 1961 |
using subspace_imp_affine[of S] subspace_0 by auto |
| 53302 | 1962 |
{
|
| 53333 | 1963 |
assume assm: "affine S \<and> 0 \<in> S" |
| 53302 | 1964 |
{
|
1965 |
fix c :: real |
|
| 54465 | 1966 |
fix x |
1967 |
assume x: "x \<in> S" |
|
| 49529 | 1968 |
have "c *\<^sub>R x = (1-c) *\<^sub>R 0 + c *\<^sub>R x" by auto |
1969 |
moreover |
|
| 53339 | 1970 |
have "(1 - c) *\<^sub>R 0 + c *\<^sub>R x \<in> S" |
| 54465 | 1971 |
using affine_alt[of S] assm x by auto |
| 53333 | 1972 |
ultimately have "c *\<^sub>R x \<in> S" by auto |
| 49529 | 1973 |
} |
| 53333 | 1974 |
then have h1: "\<forall>c. \<forall>x \<in> S. c *\<^sub>R x \<in> S" by auto |
| 49529 | 1975 |
|
| 53302 | 1976 |
{
|
1977 |
fix x y |
|
| 54465 | 1978 |
assume xy: "x \<in> S" "y \<in> S" |
| 63040 | 1979 |
define u where "u = (1 :: real)/2" |
| 53302 | 1980 |
have "(1/2) *\<^sub>R (x+y) = (1/2) *\<^sub>R (x+y)" |
1981 |
by auto |
|
| 49529 | 1982 |
moreover |
| 53302 | 1983 |
have "(1/2) *\<^sub>R (x+y)=(1/2) *\<^sub>R x + (1-(1/2)) *\<^sub>R y" |
1984 |
by (simp add: algebra_simps) |
|
| 49529 | 1985 |
moreover |
| 54465 | 1986 |
have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> S" |
1987 |
using affine_alt[of S] assm xy by auto |
|
| 49529 | 1988 |
ultimately |
| 53333 | 1989 |
have "(1/2) *\<^sub>R (x+y) \<in> S" |
| 53302 | 1990 |
using u_def by auto |
| 49529 | 1991 |
moreover |
| 54465 | 1992 |
have "x + y = 2 *\<^sub>R ((1/2) *\<^sub>R (x+y))" |
| 53302 | 1993 |
by auto |
| 49529 | 1994 |
ultimately |
| 54465 | 1995 |
have "x + y \<in> S" |
| 53302 | 1996 |
using h1[rule_format, of "(1/2) *\<^sub>R (x+y)" "2"] by auto |
| 49529 | 1997 |
} |
| 53302 | 1998 |
then have "\<forall>x \<in> S. \<forall>y \<in> S. x + y \<in> S" |
1999 |
by auto |
|
2000 |
then have "subspace S" |
|
2001 |
using h1 assm unfolding subspace_def by auto |
|
| 49529 | 2002 |
} |
2003 |
then show ?thesis using h0 by metis |
|
| 40377 | 2004 |
qed |
2005 |
||
2006 |
lemma affine_diffs_subspace: |
|
| 53333 | 2007 |
assumes "affine S" "a \<in> S" |
| 53302 | 2008 |
shows "subspace ((\<lambda>x. (-a)+x) ` S)" |
| 49529 | 2009 |
proof - |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53676
diff
changeset
|
2010 |
have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff) |
| 53302 | 2011 |
have "affine ((\<lambda>x. (-a)+x) ` S)" |
| 49531 | 2012 |
using affine_translation assms by auto |
| 53302 | 2013 |
moreover have "0 : ((\<lambda>x. (-a)+x) ` S)" |
| 53333 | 2014 |
using assms exI[of "(\<lambda>x. x\<in>S \<and> -a+x = 0)" a] by auto |
| 49531 | 2015 |
ultimately show ?thesis using subspace_affine by auto |
| 40377 | 2016 |
qed |
2017 |
||
2018 |
lemma parallel_subspace_explicit: |
|
| 54465 | 2019 |
assumes "affine S" |
2020 |
and "a \<in> S" |
|
2021 |
assumes "L \<equiv> {y. \<exists>x \<in> S. (-a) + x = y}"
|
|
2022 |
shows "subspace L \<and> affine_parallel S L" |
|
| 49529 | 2023 |
proof - |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53676
diff
changeset
|
2024 |
from assms have "L = plus (- a) ` S" by auto |
|
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53676
diff
changeset
|
2025 |
then have par: "affine_parallel S L" |
| 54465 | 2026 |
unfolding affine_parallel_def .. |
| 49531 | 2027 |
then have "affine L" using assms parallel_is_affine by auto |
| 53302 | 2028 |
moreover have "0 \<in> L" |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53676
diff
changeset
|
2029 |
using assms by auto |
| 53302 | 2030 |
ultimately show ?thesis |
2031 |
using subspace_affine par by auto |
|
| 40377 | 2032 |
qed |
2033 |
||
2034 |
lemma parallel_subspace_aux: |
|
| 53302 | 2035 |
assumes "subspace A" |
2036 |
and "subspace B" |
|
2037 |
and "affine_parallel A B" |
|
2038 |
shows "A \<supseteq> B" |
|
| 49529 | 2039 |
proof - |
| 54465 | 2040 |
from assms obtain a where a: "\<forall>x. x \<in> A \<longleftrightarrow> a + x \<in> B" |
| 49529 | 2041 |
using affine_parallel_expl[of A B] by auto |
| 53302 | 2042 |
then have "-a \<in> A" |
2043 |
using assms subspace_0[of B] by auto |
|
2044 |
then have "a \<in> A" |
|
2045 |
using assms subspace_neg[of A "-a"] by auto |
|
2046 |
then show ?thesis |
|
| 54465 | 2047 |
using assms a unfolding subspace_def by auto |
| 40377 | 2048 |
qed |
2049 |
||
2050 |
lemma parallel_subspace: |
|
| 53302 | 2051 |
assumes "subspace A" |
2052 |
and "subspace B" |
|
2053 |
and "affine_parallel A B" |
|
| 49529 | 2054 |
shows "A = B" |
2055 |
proof |
|
| 53302 | 2056 |
show "A \<supseteq> B" |
| 49529 | 2057 |
using assms parallel_subspace_aux by auto |
| 53302 | 2058 |
show "A \<subseteq> B" |
| 49529 | 2059 |
using assms parallel_subspace_aux[of B A] affine_parallel_commut by auto |
| 40377 | 2060 |
qed |
2061 |
||
2062 |
lemma affine_parallel_subspace: |
|
| 53302 | 2063 |
assumes "affine S" "S \<noteq> {}"
|
| 53339 | 2064 |
shows "\<exists>!L. subspace L \<and> affine_parallel S L" |
| 49529 | 2065 |
proof - |
| 53339 | 2066 |
have ex: "\<exists>L. subspace L \<and> affine_parallel S L" |
| 49531 | 2067 |
using assms parallel_subspace_explicit by auto |
| 53302 | 2068 |
{
|
2069 |
fix L1 L2 |
|
| 53339 | 2070 |
assume ass: "subspace L1 \<and> affine_parallel S L1" "subspace L2 \<and> affine_parallel S L2" |
| 49529 | 2071 |
then have "affine_parallel L1 L2" |
2072 |
using affine_parallel_commut[of S L1] affine_parallel_assoc[of L1 S L2] by auto |
|
2073 |
then have "L1 = L2" |
|
2074 |
using ass parallel_subspace by auto |
|
2075 |
} |
|
2076 |
then show ?thesis using ex by auto |
|
2077 |
qed |
|
2078 |
||
| 40377 | 2079 |
|
| 60420 | 2080 |
subsection \<open>Cones\<close> |
| 33175 | 2081 |
|
| 49529 | 2082 |
definition cone :: "'a::real_vector set \<Rightarrow> bool" |
| 53339 | 2083 |
where "cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>c\<ge>0. c *\<^sub>R x \<in> s)" |
| 33175 | 2084 |
|
2085 |
lemma cone_empty[intro, simp]: "cone {}"
|
|
2086 |
unfolding cone_def by auto |
|
2087 |
||
2088 |
lemma cone_univ[intro, simp]: "cone UNIV" |
|
2089 |
unfolding cone_def by auto |
|
2090 |
||
| 53339 | 2091 |
lemma cone_Inter[intro]: "\<forall>s\<in>f. cone s \<Longrightarrow> cone (\<Inter>f)" |
| 33175 | 2092 |
unfolding cone_def by auto |
2093 |
||
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
2094 |
lemma subspace_imp_cone: "subspace S \<Longrightarrow> cone S" |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
2095 |
by (simp add: cone_def subspace_mul) |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
2096 |
|
| 49529 | 2097 |
|
| 60420 | 2098 |
subsubsection \<open>Conic hull\<close> |
| 33175 | 2099 |
|
2100 |
lemma cone_cone_hull: "cone (cone hull s)" |
|
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44142
diff
changeset
|
2101 |
unfolding hull_def by auto |
| 33175 | 2102 |
|
| 53302 | 2103 |
lemma cone_hull_eq: "cone hull s = s \<longleftrightarrow> cone s" |
| 49529 | 2104 |
apply (rule hull_eq) |
| 53302 | 2105 |
using cone_Inter |
2106 |
unfolding subset_eq |
|
2107 |
apply auto |
|
| 49529 | 2108 |
done |
| 33175 | 2109 |
|
| 40377 | 2110 |
lemma mem_cone: |
| 53302 | 2111 |
assumes "cone S" "x \<in> S" "c \<ge> 0" |
| 40377 | 2112 |
shows "c *\<^sub>R x : S" |
2113 |
using assms cone_def[of S] by auto |
|
2114 |
||
2115 |
lemma cone_contains_0: |
|
| 49529 | 2116 |
assumes "cone S" |
| 53302 | 2117 |
shows "S \<noteq> {} \<longleftrightarrow> 0 \<in> S"
|
| 49529 | 2118 |
proof - |
| 53302 | 2119 |
{
|
2120 |
assume "S \<noteq> {}"
|
|
2121 |
then obtain a where "a \<in> S" by auto |
|
2122 |
then have "0 \<in> S" |
|
2123 |
using assms mem_cone[of S a 0] by auto |
|
| 49529 | 2124 |
} |
2125 |
then show ?thesis by auto |
|
| 40377 | 2126 |
qed |
2127 |
||
|
44361
75ec83d45303
remove unnecessary euclidean_space class constraints
huffman
parents:
44349
diff
changeset
|
2128 |
lemma cone_0: "cone {0}"
|
| 49529 | 2129 |
unfolding cone_def by auto |
| 40377 | 2130 |
|
| 61952 | 2131 |
lemma cone_Union[intro]: "(\<forall>s\<in>f. cone s) \<longrightarrow> cone (\<Union>f)" |
| 40377 | 2132 |
unfolding cone_def by blast |
2133 |
||
2134 |
lemma cone_iff: |
|
| 53347 | 2135 |
assumes "S \<noteq> {}"
|
2136 |
shows "cone S \<longleftrightarrow> 0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)" |
|
| 49529 | 2137 |
proof - |
| 53302 | 2138 |
{
|
2139 |
assume "cone S" |
|
2140 |
{
|
|
| 53347 | 2141 |
fix c :: real |
2142 |
assume "c > 0" |
|
| 53302 | 2143 |
{
|
2144 |
fix x |
|
| 53347 | 2145 |
assume "x \<in> S" |
2146 |
then have "x \<in> (op *\<^sub>R c) ` S" |
|
| 49529 | 2147 |
unfolding image_def |
| 60420 | 2148 |
using \<open>cone S\<close> \<open>c>0\<close> mem_cone[of S x "1/c"] |
| 54465 | 2149 |
exI[of "(\<lambda>t. t \<in> S \<and> x = c *\<^sub>R t)" "(1 / c) *\<^sub>R x"] |
| 53347 | 2150 |
by auto |
| 49529 | 2151 |
} |
2152 |
moreover |
|
| 53302 | 2153 |
{
|
2154 |
fix x |
|
| 53347 | 2155 |
assume "x \<in> (op *\<^sub>R c) ` S" |
2156 |
then have "x \<in> S" |
|
| 60420 | 2157 |
using \<open>cone S\<close> \<open>c > 0\<close> |
2158 |
unfolding cone_def image_def \<open>c > 0\<close> by auto |
|
| 49529 | 2159 |
} |
| 53302 | 2160 |
ultimately have "(op *\<^sub>R c) ` S = S" by auto |
| 40377 | 2161 |
} |
| 53339 | 2162 |
then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)" |
| 60420 | 2163 |
using \<open>cone S\<close> cone_contains_0[of S] assms by auto |
| 49529 | 2164 |
} |
2165 |
moreover |
|
| 53302 | 2166 |
{
|
| 53339 | 2167 |
assume a: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)" |
| 53302 | 2168 |
{
|
2169 |
fix x |
|
2170 |
assume "x \<in> S" |
|
| 53347 | 2171 |
fix c1 :: real |
2172 |
assume "c1 \<ge> 0" |
|
2173 |
then have "c1 = 0 \<or> c1 > 0" by auto |
|
| 60420 | 2174 |
then have "c1 *\<^sub>R x \<in> S" using a \<open>x \<in> S\<close> by auto |
| 49529 | 2175 |
} |
2176 |
then have "cone S" unfolding cone_def by auto |
|
| 40377 | 2177 |
} |
| 49529 | 2178 |
ultimately show ?thesis by blast |
2179 |
qed |
|
2180 |
||
2181 |
lemma cone_hull_empty: "cone hull {} = {}"
|
|
2182 |
by (metis cone_empty cone_hull_eq) |
|
2183 |
||
| 53302 | 2184 |
lemma cone_hull_empty_iff: "S = {} \<longleftrightarrow> cone hull S = {}"
|
| 49529 | 2185 |
by (metis bot_least cone_hull_empty hull_subset xtrans(5)) |
2186 |
||
| 53302 | 2187 |
lemma cone_hull_contains_0: "S \<noteq> {} \<longleftrightarrow> 0 \<in> cone hull S"
|
| 49529 | 2188 |
using cone_cone_hull[of S] cone_contains_0[of "cone hull S"] cone_hull_empty_iff[of S] |
2189 |
by auto |
|
| 40377 | 2190 |
|
2191 |
lemma mem_cone_hull: |
|
| 53347 | 2192 |
assumes "x \<in> S" "c \<ge> 0" |
| 53302 | 2193 |
shows "c *\<^sub>R x \<in> cone hull S" |
| 49529 | 2194 |
by (metis assms cone_cone_hull hull_inc mem_cone) |
2195 |
||
| 53339 | 2196 |
lemma cone_hull_expl: "cone hull S = {c *\<^sub>R x | c x. c \<ge> 0 \<and> x \<in> S}"
|
2197 |
(is "?lhs = ?rhs") |
|
| 49529 | 2198 |
proof - |
| 53302 | 2199 |
{
|
2200 |
fix x |
|
2201 |
assume "x \<in> ?rhs" |
|
| 54465 | 2202 |
then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S" |
| 49529 | 2203 |
by auto |
| 53347 | 2204 |
fix c :: real |
2205 |
assume c: "c \<ge> 0" |
|
| 53339 | 2206 |
then have "c *\<^sub>R x = (c * cx) *\<^sub>R xx" |
| 54465 | 2207 |
using x by (simp add: algebra_simps) |
| 49529 | 2208 |
moreover |
| 56536 | 2209 |
have "c * cx \<ge> 0" using c x by auto |
| 49529 | 2210 |
ultimately |
| 54465 | 2211 |
have "c *\<^sub>R x \<in> ?rhs" using x by auto |
| 53302 | 2212 |
} |
| 53347 | 2213 |
then have "cone ?rhs" |
2214 |
unfolding cone_def by auto |
|
2215 |
then have "?rhs \<in> Collect cone" |
|
2216 |
unfolding mem_Collect_eq by auto |
|
| 53302 | 2217 |
{
|
2218 |
fix x |
|
2219 |
assume "x \<in> S" |
|
2220 |
then have "1 *\<^sub>R x \<in> ?rhs" |
|
| 49531 | 2221 |
apply auto |
| 53347 | 2222 |
apply (rule_tac x = 1 in exI) |
| 49529 | 2223 |
apply auto |
2224 |
done |
|
| 53302 | 2225 |
then have "x \<in> ?rhs" by auto |
| 53347 | 2226 |
} |
2227 |
then have "S \<subseteq> ?rhs" by auto |
|
| 53302 | 2228 |
then have "?lhs \<subseteq> ?rhs" |
| 60420 | 2229 |
using \<open>?rhs \<in> Collect cone\<close> hull_minimal[of S "?rhs" "cone"] by auto |
| 49529 | 2230 |
moreover |
| 53302 | 2231 |
{
|
2232 |
fix x |
|
2233 |
assume "x \<in> ?rhs" |
|
| 54465 | 2234 |
then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S" |
| 53339 | 2235 |
by auto |
2236 |
then have "xx \<in> cone hull S" |
|
2237 |
using hull_subset[of S] by auto |
|
| 53302 | 2238 |
then have "x \<in> ?lhs" |
| 54465 | 2239 |
using x cone_cone_hull[of S] cone_def[of "cone hull S"] by auto |
| 49529 | 2240 |
} |
2241 |
ultimately show ?thesis by auto |
|
| 40377 | 2242 |
qed |
2243 |
||
2244 |
lemma cone_closure: |
|
| 53347 | 2245 |
fixes S :: "'a::real_normed_vector set" |
| 49529 | 2246 |
assumes "cone S" |
2247 |
shows "cone (closure S)" |
|
2248 |
proof (cases "S = {}")
|
|
2249 |
case True |
|
2250 |
then show ?thesis by auto |
|
2251 |
next |
|
2252 |
case False |
|
| 53339 | 2253 |
then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)" |
| 49529 | 2254 |
using cone_iff[of S] assms by auto |
| 53339 | 2255 |
then have "0 \<in> closure S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` closure S = closure S)" |
| 49529 | 2256 |
using closure_subset by (auto simp add: closure_scaleR) |
| 53339 | 2257 |
then show ?thesis |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
2258 |
using False cone_iff[of "closure S"] by auto |
| 49529 | 2259 |
qed |
2260 |
||
| 40377 | 2261 |
|
| 60420 | 2262 |
subsection \<open>Affine dependence and consequential theorems (from Lars Schewe)\<close> |
| 33175 | 2263 |
|
| 49529 | 2264 |
definition affine_dependent :: "'a::real_vector set \<Rightarrow> bool" |
| 53339 | 2265 |
where "affine_dependent s \<longleftrightarrow> (\<exists>x\<in>s. x \<in> affine hull (s - {x}))"
|
| 33175 | 2266 |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2267 |
lemma affine_dependent_subset: |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2268 |
"\<lbrakk>affine_dependent s; s \<subseteq> t\<rbrakk> \<Longrightarrow> affine_dependent t" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2269 |
apply (simp add: affine_dependent_def Bex_def) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2270 |
apply (blast dest: hull_mono [OF Diff_mono [OF _ subset_refl]]) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2271 |
done |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2272 |
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2273 |
lemma affine_independent_subset: |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2274 |
shows "\<lbrakk>~ affine_dependent t; s \<subseteq> t\<rbrakk> \<Longrightarrow> ~ affine_dependent s" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2275 |
by (metis affine_dependent_subset) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2276 |
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2277 |
lemma affine_independent_Diff: |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2278 |
"~ affine_dependent s \<Longrightarrow> ~ affine_dependent(s - t)" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2279 |
by (meson Diff_subset affine_dependent_subset) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2280 |
|
| 33175 | 2281 |
lemma affine_dependent_explicit: |
2282 |
"affine_dependent p \<longleftrightarrow> |
|
| 64267 | 2283 |
(\<exists>s u. finite s \<and> s \<subseteq> p \<and> sum u s = 0 \<and> |
2284 |
(\<exists>v\<in>s. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) s = 0)" |
|
| 49529 | 2285 |
unfolding affine_dependent_def affine_hull_explicit mem_Collect_eq |
2286 |
apply rule |
|
2287 |
apply (erule bexE, erule exE, erule exE) |
|
2288 |
apply (erule conjE)+ |
|
2289 |
defer |
|
2290 |
apply (erule exE, erule exE) |
|
2291 |
apply (erule conjE)+ |
|
2292 |
apply (erule bexE) |
|
2293 |
proof - |
|
2294 |
fix x s u |
|
| 64267 | 2295 |
assume as: "x \<in> p" "finite s" "s \<noteq> {}" "s \<subseteq> p - {x}" "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
|
| 53302 | 2296 |
have "x \<notin> s" using as(1,4) by auto |
| 64267 | 2297 |
show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> sum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0" |
| 49529 | 2298 |
apply (rule_tac x="insert x s" in exI, rule_tac x="\<lambda>v. if v = x then - 1 else u v" in exI) |
| 64267 | 2299 |
unfolding if_smult and sum_clauses(2)[OF as(2)] and sum_delta_notmem[OF \<open>x\<notin>s\<close>] and as |
| 53339 | 2300 |
using as |
2301 |
apply auto |
|
| 49529 | 2302 |
done |
| 33175 | 2303 |
next |
| 49529 | 2304 |
fix s u v |
| 64267 | 2305 |
assume as: "finite s" "s \<subseteq> p" "sum u s = 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" "v \<in> s" "u v \<noteq> 0" |
| 53339 | 2306 |
have "s \<noteq> {v}"
|
2307 |
using as(3,6) by auto |
|
| 64267 | 2308 |
then show "\<exists>x\<in>p. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p - {x} \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
|
| 53302 | 2309 |
apply (rule_tac x=v in bexI) |
2310 |
apply (rule_tac x="s - {v}" in exI)
|
|
2311 |
apply (rule_tac x="\<lambda>x. - (1 / u v) * u x" in exI) |
|
| 64267 | 2312 |
unfolding scaleR_scaleR[symmetric] and scaleR_right.sum [symmetric] |
2313 |
unfolding sum_distrib_left[symmetric] and sum_diff1[OF as(1)] |
|
| 53302 | 2314 |
using as |
2315 |
apply auto |
|
| 49529 | 2316 |
done |
| 33175 | 2317 |
qed |
2318 |
||
2319 |
lemma affine_dependent_explicit_finite: |
|
| 49529 | 2320 |
fixes s :: "'a::real_vector set" |
2321 |
assumes "finite s" |
|
| 53302 | 2322 |
shows "affine_dependent s \<longleftrightarrow> |
| 64267 | 2323 |
(\<exists>u. sum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) s = 0)" |
| 33175 | 2324 |
(is "?lhs = ?rhs") |
2325 |
proof |
|
| 53347 | 2326 |
have *: "\<And>vt u v. (if vt then u v else 0) *\<^sub>R v = (if vt then (u v) *\<^sub>R v else 0::'a)" |
| 49529 | 2327 |
by auto |
| 33175 | 2328 |
assume ?lhs |
| 49529 | 2329 |
then obtain t u v where |
| 64267 | 2330 |
"finite t" "t \<subseteq> s" "sum u t = 0" "v\<in>t" "u v \<noteq> 0" "(\<Sum>v\<in>t. u v *\<^sub>R v) = 0" |
| 33175 | 2331 |
unfolding affine_dependent_explicit by auto |
| 49529 | 2332 |
then show ?rhs |
2333 |
apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI) |
|
| 64267 | 2334 |
apply auto unfolding * and sum.inter_restrict[OF assms, symmetric] |
| 60420 | 2335 |
unfolding Int_absorb1[OF \<open>t\<subseteq>s\<close>] |
| 49529 | 2336 |
apply auto |
2337 |
done |
|
| 33175 | 2338 |
next |
2339 |
assume ?rhs |
|
| 64267 | 2340 |
then obtain u v where "sum u s = 0" "v\<in>s" "u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" |
| 53339 | 2341 |
by auto |
| 49529 | 2342 |
then show ?lhs unfolding affine_dependent_explicit |
2343 |
using assms by auto |
|
2344 |
qed |
|
2345 |
||
| 33175 | 2346 |
|
| 60420 | 2347 |
subsection \<open>Connectedness of convex sets\<close> |
|
44465
fa56622bb7bc
move connected_real_lemma to the one place it is used
huffman
parents:
44457
diff
changeset
|
2348 |
|
|
51480
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2349 |
lemma connectedD: |
|
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2350 |
"connected S \<Longrightarrow> open A \<Longrightarrow> open B \<Longrightarrow> S \<subseteq> A \<union> B \<Longrightarrow> A \<inter> B \<inter> S = {} \<Longrightarrow> A \<inter> S = {} \<or> B \<inter> 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
|
2351 |
by (rule Topological_Spaces.topological_space_class.connectedD) |
| 33175 | 2352 |
|
2353 |
lemma convex_connected: |
|
2354 |
fixes s :: "'a::real_normed_vector set" |
|
| 53302 | 2355 |
assumes "convex s" |
2356 |
shows "connected s" |
|
|
51480
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2357 |
proof (rule connectedI) |
|
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2358 |
fix A B |
|
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2359 |
assume "open A" "open B" "A \<inter> B \<inter> s = {}" "s \<subseteq> A \<union> B"
|
|
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2360 |
moreover |
|
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2361 |
assume "A \<inter> s \<noteq> {}" "B \<inter> s \<noteq> {}"
|
|
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2362 |
then obtain a b where a: "a \<in> A" "a \<in> s" and b: "b \<in> B" "b \<in> s" by auto |
| 63040 | 2363 |
define f where [abs_def]: "f u = u *\<^sub>R a + (1 - u) *\<^sub>R b" for u |
|
51480
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2364 |
then have "continuous_on {0 .. 1} f"
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56369
diff
changeset
|
2365 |
by (auto intro!: continuous_intros) |
|
51480
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2366 |
then have "connected (f ` {0 .. 1})"
|
|
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2367 |
by (auto intro!: connected_continuous_image) |
|
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2368 |
note connectedD[OF this, of A B] |
|
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2369 |
moreover have "a \<in> A \<inter> f ` {0 .. 1}"
|
|
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2370 |
using a by (auto intro!: image_eqI[of _ _ 1] simp: f_def) |
|
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2371 |
moreover have "b \<in> B \<inter> f ` {0 .. 1}"
|
|
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2372 |
using b by (auto intro!: image_eqI[of _ _ 0] simp: f_def) |
|
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2373 |
moreover have "f ` {0 .. 1} \<subseteq> s"
|
| 60420 | 2374 |
using \<open>convex s\<close> a b unfolding convex_def f_def by auto |
|
51480
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
hoelzl
parents:
51475
diff
changeset
|
2375 |
ultimately show False by auto |
| 33175 | 2376 |
qed |
2377 |
||
|
61426
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61222
diff
changeset
|
2378 |
corollary connected_UNIV[intro]: "connected (UNIV :: 'a::real_normed_vector set)" |
|
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61222
diff
changeset
|
2379 |
by(simp add: convex_connected) |
| 33175 | 2380 |
|
|
62131
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62097
diff
changeset
|
2381 |
proposition clopen: |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62097
diff
changeset
|
2382 |
fixes s :: "'a :: real_normed_vector set" |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62097
diff
changeset
|
2383 |
shows "closed s \<and> open s \<longleftrightarrow> s = {} \<or> s = UNIV"
|
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62097
diff
changeset
|
2384 |
apply (rule iffI) |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62097
diff
changeset
|
2385 |
apply (rule connected_UNIV [unfolded connected_clopen, rule_format]) |
| 64122 | 2386 |
apply (force simp add: closed_closedin, force) |
|
62131
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62097
diff
changeset
|
2387 |
done |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62097
diff
changeset
|
2388 |
|
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62097
diff
changeset
|
2389 |
corollary compact_open: |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62097
diff
changeset
|
2390 |
fixes s :: "'a :: euclidean_space set" |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62097
diff
changeset
|
2391 |
shows "compact s \<and> open s \<longleftrightarrow> s = {}"
|
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62097
diff
changeset
|
2392 |
by (auto simp: compact_eq_bounded_closed clopen) |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62097
diff
changeset
|
2393 |
|
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
2394 |
corollary finite_imp_not_open: |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
2395 |
fixes S :: "'a::{real_normed_vector, perfect_space} set"
|
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
2396 |
shows "\<lbrakk>finite S; open S\<rbrakk> \<Longrightarrow> S={}"
|
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
2397 |
using clopen [of S] finite_imp_closed not_bounded_UNIV by blast |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
2398 |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2399 |
corollary empty_interior_finite: |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2400 |
fixes S :: "'a::{real_normed_vector, perfect_space} set"
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2401 |
shows "finite S \<Longrightarrow> interior S = {}"
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2402 |
by (metis interior_subset finite_subset open_interior [of S] finite_imp_not_open) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
2403 |
|
| 60420 | 2404 |
text \<open>Balls, being convex, are connected.\<close> |
| 33175 | 2405 |
|
| 56188 | 2406 |
lemma convex_prod: |
| 53347 | 2407 |
assumes "\<And>i. i \<in> Basis \<Longrightarrow> convex {x. P i x}"
|
|
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
|
2408 |
shows "convex {x. \<forall>i\<in>Basis. P i (x\<bullet>i)}"
|
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
2409 |
using assms unfolding convex_def |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
2410 |
by (auto simp: inner_add_left) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
2411 |
|
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
2412 |
lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i)}"
|
| 56188 | 2413 |
by (rule convex_prod) (simp add: atLeast_def[symmetric] convex_real_interval) |
| 33175 | 2414 |
|
2415 |
lemma convex_local_global_minimum: |
|
2416 |
fixes s :: "'a::real_normed_vector set" |
|
| 53347 | 2417 |
assumes "e > 0" |
2418 |
and "convex_on s f" |
|
2419 |
and "ball x e \<subseteq> s" |
|
2420 |
and "\<forall>y\<in>ball x e. f x \<le> f y" |
|
| 33175 | 2421 |
shows "\<forall>y\<in>s. f x \<le> f y" |
| 53302 | 2422 |
proof (rule ccontr) |
2423 |
have "x \<in> s" using assms(1,3) by auto |
|
2424 |
assume "\<not> ?thesis" |
|
2425 |
then obtain y where "y\<in>s" and y: "f x > f y" by auto |
|
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61952
diff
changeset
|
2426 |
then have xy: "0 < dist x y" by auto |
| 53347 | 2427 |
then obtain u where "0 < u" "u \<le> 1" and u: "u < e / dist x y" |
| 60420 | 2428 |
using real_lbound_gt_zero[of 1 "e / dist x y"] xy \<open>e>0\<close> by auto |
| 53302 | 2429 |
then have "f ((1-u) *\<^sub>R x + u *\<^sub>R y) \<le> (1-u) * f x + u * f y" |
| 60420 | 2430 |
using \<open>x\<in>s\<close> \<open>y\<in>s\<close> |
| 53302 | 2431 |
using assms(2)[unfolded convex_on_def, |
2432 |
THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x="1-u"]] |
|
| 50804 | 2433 |
by auto |
| 33175 | 2434 |
moreover |
| 50804 | 2435 |
have *: "x - ((1 - u) *\<^sub>R x + u *\<^sub>R y) = u *\<^sub>R (x - y)" |
2436 |
by (simp add: algebra_simps) |
|
2437 |
have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> ball x e" |
|
| 53302 | 2438 |
unfolding mem_ball dist_norm |
| 60420 | 2439 |
unfolding * and norm_scaleR and abs_of_pos[OF \<open>0<u\<close>] |
| 50804 | 2440 |
unfolding dist_norm[symmetric] |
| 53302 | 2441 |
using u |
2442 |
unfolding pos_less_divide_eq[OF xy] |
|
2443 |
by auto |
|
2444 |
then have "f x \<le> f ((1 - u) *\<^sub>R x + u *\<^sub>R y)" |
|
2445 |
using assms(4) by auto |
|
| 50804 | 2446 |
ultimately show False |
| 60420 | 2447 |
using mult_strict_left_mono[OF y \<open>u>0\<close>] |
| 53302 | 2448 |
unfolding left_diff_distrib |
2449 |
by auto |
|
| 33175 | 2450 |
qed |
2451 |
||
|
60800
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
2452 |
lemma convex_ball [iff]: |
| 33175 | 2453 |
fixes x :: "'a::real_normed_vector" |
| 49531 | 2454 |
shows "convex (ball x e)" |
| 50804 | 2455 |
proof (auto simp add: convex_def) |
2456 |
fix y z |
|
2457 |
assume yz: "dist x y < e" "dist x z < e" |
|
2458 |
fix u v :: real |
|
2459 |
assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" |
|
2460 |
have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" |
|
2461 |
using uv yz |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2462 |
using convex_on_dist [of "ball x e" x, unfolded convex_on_def, |
| 53302 | 2463 |
THEN bspec[where x=y], THEN bspec[where x=z]] |
| 50804 | 2464 |
by auto |
2465 |
then show "dist x (u *\<^sub>R y + v *\<^sub>R z) < e" |
|
2466 |
using convex_bound_lt[OF yz uv] by auto |
|
| 33175 | 2467 |
qed |
2468 |
||
|
60800
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
2469 |
lemma convex_cball [iff]: |
| 33175 | 2470 |
fixes x :: "'a::real_normed_vector" |
| 53347 | 2471 |
shows "convex (cball x e)" |
2472 |
proof - |
|
2473 |
{
|
|
2474 |
fix y z |
|
2475 |
assume yz: "dist x y \<le> e" "dist x z \<le> e" |
|
2476 |
fix u v :: real |
|
2477 |
assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" |
|
2478 |
have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" |
|
2479 |
using uv yz |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2480 |
using convex_on_dist [of "cball x e" x, unfolded convex_on_def, |
| 53347 | 2481 |
THEN bspec[where x=y], THEN bspec[where x=z]] |
2482 |
by auto |
|
2483 |
then have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> e" |
|
2484 |
using convex_bound_le[OF yz uv] by auto |
|
2485 |
} |
|
2486 |
then show ?thesis by (auto simp add: convex_def Ball_def) |
|
| 33175 | 2487 |
qed |
2488 |
||
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
2489 |
lemma connected_ball [iff]: |
| 33175 | 2490 |
fixes x :: "'a::real_normed_vector" |
2491 |
shows "connected (ball x e)" |
|
2492 |
using convex_connected convex_ball by auto |
|
2493 |
||
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
2494 |
lemma connected_cball [iff]: |
| 33175 | 2495 |
fixes x :: "'a::real_normed_vector" |
| 53302 | 2496 |
shows "connected (cball x e)" |
| 33175 | 2497 |
using convex_connected convex_cball by auto |
2498 |
||
| 50804 | 2499 |
|
| 60420 | 2500 |
subsection \<open>Convex hull\<close> |
| 33175 | 2501 |
|
| 60762 | 2502 |
lemma convex_convex_hull [iff]: "convex (convex hull s)" |
| 53302 | 2503 |
unfolding hull_def |
2504 |
using convex_Inter[of "{t. convex t \<and> s \<subseteq> t}"]
|
|
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44142
diff
changeset
|
2505 |
by auto |
| 33175 | 2506 |
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
2507 |
lemma convex_hull_subset: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
2508 |
"s \<subseteq> convex hull t \<Longrightarrow> convex hull s \<subseteq> convex hull t" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
2509 |
by (simp add: convex_convex_hull subset_hull) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
2510 |
|
|
34064
eee04bbbae7e
avoid dependency on implicit dest rule predicate1D in proofs
haftmann
parents:
33758
diff
changeset
|
2511 |
lemma convex_hull_eq: "convex hull s = s \<longleftrightarrow> convex s" |
| 50804 | 2512 |
by (metis convex_convex_hull hull_same) |
| 33175 | 2513 |
|
2514 |
lemma bounded_convex_hull: |
|
2515 |
fixes s :: "'a::real_normed_vector set" |
|
| 53347 | 2516 |
assumes "bounded s" |
2517 |
shows "bounded (convex hull s)" |
|
| 50804 | 2518 |
proof - |
2519 |
from assms obtain B where B: "\<forall>x\<in>s. norm x \<le> B" |
|
2520 |
unfolding bounded_iff by auto |
|
2521 |
show ?thesis |
|
2522 |
apply (rule bounded_subset[OF bounded_cball, of _ 0 B]) |
|
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44142
diff
changeset
|
2523 |
unfolding subset_hull[of convex, OF convex_cball] |
| 53302 | 2524 |
unfolding subset_eq mem_cball dist_norm using B |
2525 |
apply auto |
|
| 50804 | 2526 |
done |
2527 |
qed |
|
| 33175 | 2528 |
|
2529 |
lemma finite_imp_bounded_convex_hull: |
|
2530 |
fixes s :: "'a::real_normed_vector set" |
|
| 53302 | 2531 |
shows "finite s \<Longrightarrow> bounded (convex hull s)" |
2532 |
using bounded_convex_hull finite_imp_bounded |
|
2533 |
by auto |
|
| 33175 | 2534 |
|
| 50804 | 2535 |
|
| 60420 | 2536 |
subsubsection \<open>Convex hull is "preserved" by a linear function\<close> |
| 40377 | 2537 |
|
2538 |
lemma convex_hull_linear_image: |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2539 |
assumes f: "linear f" |
| 40377 | 2540 |
shows "f ` (convex hull s) = convex hull (f ` s)" |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2541 |
proof |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2542 |
show "convex hull (f ` s) \<subseteq> f ` (convex hull s)" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2543 |
by (intro hull_minimal image_mono hull_subset convex_linear_image assms convex_convex_hull) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2544 |
show "f ` (convex hull s) \<subseteq> convex hull (f ` s)" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2545 |
proof (unfold image_subset_iff_subset_vimage, rule hull_minimal) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2546 |
show "s \<subseteq> f -` (convex hull (f ` s))" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2547 |
by (fast intro: hull_inc) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2548 |
show "convex (f -` (convex hull (f ` s)))" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2549 |
by (intro convex_linear_vimage [OF f] convex_convex_hull) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2550 |
qed |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2551 |
qed |
| 40377 | 2552 |
|
2553 |
lemma in_convex_hull_linear_image: |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2554 |
assumes "linear f" |
| 53347 | 2555 |
and "x \<in> convex hull s" |
| 53339 | 2556 |
shows "f x \<in> convex hull (f ` s)" |
| 50804 | 2557 |
using convex_hull_linear_image[OF assms(1)] assms(2) by auto |
2558 |
||
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2559 |
lemma convex_hull_Times: |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2560 |
"convex hull (s \<times> t) = (convex hull s) \<times> (convex hull t)" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2561 |
proof |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2562 |
show "convex hull (s \<times> t) \<subseteq> (convex hull s) \<times> (convex hull t)" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2563 |
by (intro hull_minimal Sigma_mono hull_subset convex_Times convex_convex_hull) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2564 |
have "\<forall>x\<in>convex hull s. \<forall>y\<in>convex hull t. (x, y) \<in> convex hull (s \<times> t)" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2565 |
proof (intro hull_induct) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2566 |
fix x y assume "x \<in> s" and "y \<in> t" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2567 |
then show "(x, y) \<in> convex hull (s \<times> t)" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2568 |
by (simp add: hull_inc) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2569 |
next |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2570 |
fix x let ?S = "((\<lambda>y. (0, y)) -` (\<lambda>p. (- x, 0) + p) ` (convex hull s \<times> t))" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2571 |
have "convex ?S" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2572 |
by (intro convex_linear_vimage convex_translation convex_convex_hull, |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2573 |
simp add: linear_iff) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2574 |
also have "?S = {y. (x, y) \<in> convex hull (s \<times> t)}"
|
| 57865 | 2575 |
by (auto simp add: image_def Bex_def) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2576 |
finally show "convex {y. (x, y) \<in> convex hull (s \<times> t)}" .
|
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2577 |
next |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2578 |
show "convex {x. \<forall>y\<in>convex hull t. (x, y) \<in> convex hull (s \<times> t)}"
|
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2579 |
proof (unfold Collect_ball_eq, rule convex_INT [rule_format]) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2580 |
fix y let ?S = "((\<lambda>x. (x, 0)) -` (\<lambda>p. (0, - y) + p) ` (convex hull s \<times> t))" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2581 |
have "convex ?S" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2582 |
by (intro convex_linear_vimage convex_translation convex_convex_hull, |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2583 |
simp add: linear_iff) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2584 |
also have "?S = {x. (x, y) \<in> convex hull (s \<times> t)}"
|
| 57865 | 2585 |
by (auto simp add: image_def Bex_def) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2586 |
finally show "convex {x. (x, y) \<in> convex hull (s \<times> t)}" .
|
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2587 |
qed |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2588 |
qed |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2589 |
then show "(convex hull s) \<times> (convex hull t) \<subseteq> convex hull (s \<times> t)" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2590 |
unfolding subset_eq split_paired_Ball_Sigma . |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2591 |
qed |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
2592 |
|
| 40377 | 2593 |
|
| 60420 | 2594 |
subsubsection \<open>Stepping theorems for convex hulls of finite sets\<close> |
| 33175 | 2595 |
|
2596 |
lemma convex_hull_empty[simp]: "convex hull {} = {}"
|
|
| 50804 | 2597 |
by (rule hull_unique) auto |
| 33175 | 2598 |
|
2599 |
lemma convex_hull_singleton[simp]: "convex hull {a} = {a}"
|
|
| 50804 | 2600 |
by (rule hull_unique) auto |
| 33175 | 2601 |
|
2602 |
lemma convex_hull_insert: |
|
2603 |
fixes s :: "'a::real_vector set" |
|
2604 |
assumes "s \<noteq> {}"
|
|
| 50804 | 2605 |
shows "convex hull (insert a s) = |
2606 |
{x. \<exists>u\<ge>0. \<exists>v\<ge>0. \<exists>b. (u + v = 1) \<and> b \<in> (convex hull s) \<and> (x = u *\<^sub>R a + v *\<^sub>R b)}"
|
|
| 53347 | 2607 |
(is "_ = ?hull") |
| 50804 | 2608 |
apply (rule, rule hull_minimal, rule) |
2609 |
unfolding insert_iff |
|
2610 |
prefer 3 |
|
2611 |
apply rule |
|
2612 |
proof - |
|
2613 |
fix x |
|
2614 |
assume x: "x = a \<or> x \<in> s" |
|
2615 |
then show "x \<in> ?hull" |
|
2616 |
apply rule |
|
2617 |
unfolding mem_Collect_eq |
|
2618 |
apply (rule_tac x=1 in exI) |
|
2619 |
defer |
|
2620 |
apply (rule_tac x=0 in exI) |
|
2621 |
using assms hull_subset[of s convex] |
|
2622 |
apply auto |
|
2623 |
done |
|
| 33175 | 2624 |
next |
| 50804 | 2625 |
fix x |
2626 |
assume "x \<in> ?hull" |
|
2627 |
then obtain u v b where obt: "u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull s" "x = u *\<^sub>R a + v *\<^sub>R b" |
|
2628 |
by auto |
|
| 53339 | 2629 |
have "a \<in> convex hull insert a s" "b \<in> convex hull insert a s" |
| 50804 | 2630 |
using hull_mono[of s "insert a s" convex] hull_mono[of "{a}" "insert a s" convex] and obt(4)
|
2631 |
by auto |
|
2632 |
then show "x \<in> convex hull insert a s" |
|
| 53676 | 2633 |
unfolding obt(5) using obt(1-3) |
2634 |
by (rule convexD [OF convex_convex_hull]) |
|
| 33175 | 2635 |
next |
| 50804 | 2636 |
show "convex ?hull" |
| 53676 | 2637 |
proof (rule convexI) |
| 50804 | 2638 |
fix x y u v |
2639 |
assume as: "(0::real) \<le> u" "0 \<le> v" "u + v = 1" "x\<in>?hull" "y\<in>?hull" |
|
| 53339 | 2640 |
from as(4) obtain u1 v1 b1 where |
2641 |
obt1: "u1\<ge>0" "v1\<ge>0" "u1 + v1 = 1" "b1 \<in> convex hull s" "x = u1 *\<^sub>R a + v1 *\<^sub>R b1" |
|
2642 |
by auto |
|
2643 |
from as(5) obtain u2 v2 b2 where |
|
2644 |
obt2: "u2\<ge>0" "v2\<ge>0" "u2 + v2 = 1" "b2 \<in> convex hull s" "y = u2 *\<^sub>R a + v2 *\<^sub>R b2" |
|
2645 |
by auto |
|
| 50804 | 2646 |
have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" |
2647 |
by (auto simp add: algebra_simps) |
|
2648 |
have **: "\<exists>b \<in> convex hull s. u *\<^sub>R x + v *\<^sub>R y = |
|
2649 |
(u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)" |
|
2650 |
proof (cases "u * v1 + v * v2 = 0") |
|
2651 |
case True |
|
2652 |
have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" |
|
2653 |
by (auto simp add: algebra_simps) |
|
2654 |
from True have ***: "u * v1 = 0" "v * v2 = 0" |
|
| 60420 | 2655 |
using mult_nonneg_nonneg[OF \<open>u\<ge>0\<close> \<open>v1\<ge>0\<close>] mult_nonneg_nonneg[OF \<open>v\<ge>0\<close> \<open>v2\<ge>0\<close>] |
| 53302 | 2656 |
by arith+ |
| 50804 | 2657 |
then have "u * u1 + v * u2 = 1" |
2658 |
using as(3) obt1(3) obt2(3) by auto |
|
2659 |
then show ?thesis |
|
2660 |
unfolding obt1(5) obt2(5) * |
|
2661 |
using assms hull_subset[of s convex] |
|
2662 |
by (auto simp add: *** scaleR_right_distrib) |
|
| 33175 | 2663 |
next |
| 50804 | 2664 |
case False |
2665 |
have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)" |
|
2666 |
using as(3) obt1(3) obt2(3) by (auto simp add: field_simps) |
|
2667 |
also have "\<dots> = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)" |
|
2668 |
using as(3) obt1(3) obt2(3) by (auto simp add: field_simps) |
|
2669 |
also have "\<dots> = u * v1 + v * v2" |
|
2670 |
by simp |
|
2671 |
finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto |
|
2672 |
have "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2" |
|
| 56536 | 2673 |
using as(1,2) obt1(1,2) obt2(1,2) by auto |
| 50804 | 2674 |
then show ?thesis |
2675 |
unfolding obt1(5) obt2(5) |
|
2676 |
unfolding * and ** |
|
2677 |
using False |
|
| 53339 | 2678 |
apply (rule_tac |
2679 |
x = "((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" in bexI) |
|
| 50804 | 2680 |
defer |
| 53676 | 2681 |
apply (rule convexD [OF convex_convex_hull]) |
| 50804 | 2682 |
using obt1(4) obt2(4) |
| 49530 | 2683 |
unfolding add_divide_distrib[symmetric] and zero_le_divide_iff |
| 50804 | 2684 |
apply (auto simp add: scaleR_left_distrib scaleR_right_distrib) |
2685 |
done |
|
2686 |
qed |
|
2687 |
have u1: "u1 \<le> 1" |
|
2688 |
unfolding obt1(3)[symmetric] and not_le using obt1(2) by auto |
|
2689 |
have u2: "u2 \<le> 1" |
|
2690 |
unfolding obt2(3)[symmetric] and not_le using obt2(2) by auto |
|
| 53339 | 2691 |
have "u1 * u + u2 * v \<le> max u1 u2 * u + max u1 u2 * v" |
| 50804 | 2692 |
apply (rule add_mono) |
2693 |
apply (rule_tac [!] mult_right_mono) |
|
2694 |
using as(1,2) obt1(1,2) obt2(1,2) |
|
2695 |
apply auto |
|
2696 |
done |
|
2697 |
also have "\<dots> \<le> 1" |
|
2698 |
unfolding distrib_left[symmetric] and as(3) using u1 u2 by auto |
|
2699 |
finally show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" |
|
2700 |
unfolding mem_Collect_eq |
|
2701 |
apply (rule_tac x="u * u1 + v * u2" in exI) |
|
2702 |
apply (rule conjI) |
|
2703 |
defer |
|
2704 |
apply (rule_tac x="1 - u * u1 - v * u2" in exI) |
|
2705 |
unfolding Bex_def |
|
2706 |
using as(1,2) obt1(1,2) obt2(1,2) ** |
|
| 56536 | 2707 |
apply (auto simp add: algebra_simps) |
| 50804 | 2708 |
done |
| 33175 | 2709 |
qed |
2710 |
qed |
|
2711 |
||
2712 |
||
| 60420 | 2713 |
subsubsection \<open>Explicit expression for convex hull\<close> |
| 33175 | 2714 |
|
2715 |
lemma convex_hull_indexed: |
|
2716 |
fixes s :: "'a::real_vector set" |
|
| 50804 | 2717 |
shows "convex hull s = |
| 53347 | 2718 |
{y. \<exists>k u x.
|
2719 |
(\<forall>i\<in>{1::nat .. k}. 0 \<le> u i \<and> x i \<in> s) \<and>
|
|
| 64267 | 2720 |
(sum u {1..k} = 1) \<and> (sum (\<lambda>i. u i *\<^sub>R x i) {1..k} = y)}"
|
| 53339 | 2721 |
(is "?xyz = ?hull") |
| 50804 | 2722 |
apply (rule hull_unique) |
2723 |
apply rule |
|
2724 |
defer |
|
| 53676 | 2725 |
apply (rule convexI) |
| 50804 | 2726 |
proof - |
2727 |
fix x |
|
2728 |
assume "x\<in>s" |
|
2729 |
then show "x \<in> ?hull" |
|
2730 |
unfolding mem_Collect_eq |
|
2731 |
apply (rule_tac x=1 in exI, rule_tac x="\<lambda>x. 1" in exI) |
|
2732 |
apply auto |
|
2733 |
done |
|
| 33175 | 2734 |
next |
| 50804 | 2735 |
fix t |
2736 |
assume as: "s \<subseteq> t" "convex t" |
|
2737 |
show "?hull \<subseteq> t" |
|
2738 |
apply rule |
|
2739 |
unfolding mem_Collect_eq |
|
| 53302 | 2740 |
apply (elim exE conjE) |
| 50804 | 2741 |
proof - |
2742 |
fix x k u y |
|
2743 |
assume assm: |
|
2744 |
"\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> s"
|
|
| 64267 | 2745 |
"sum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
|
| 50804 | 2746 |
show "x\<in>t" |
2747 |
unfolding assm(3) [symmetric] |
|
2748 |
apply (rule as(2)[unfolded convex, rule_format]) |
|
2749 |
using assm(1,2) as(1) apply auto |
|
2750 |
done |
|
2751 |
qed |
|
| 33175 | 2752 |
next |
| 50804 | 2753 |
fix x y u v |
| 53347 | 2754 |
assume uv: "0 \<le> u" "0 \<le> v" "u + v = (1::real)" |
2755 |
assume xy: "x \<in> ?hull" "y \<in> ?hull" |
|
| 50804 | 2756 |
from xy obtain k1 u1 x1 where |
| 64267 | 2757 |
x: "\<forall>i\<in>{1::nat..k1}. 0\<le>u1 i \<and> x1 i \<in> s" "sum u1 {Suc 0..k1} = 1" "(\<Sum>i = Suc 0..k1. u1 i *\<^sub>R x1 i) = x"
|
| 50804 | 2758 |
by auto |
2759 |
from xy obtain k2 u2 x2 where |
|
| 64267 | 2760 |
y: "\<forall>i\<in>{1::nat..k2}. 0\<le>u2 i \<and> x2 i \<in> s" "sum u2 {Suc 0..k2} = 1" "(\<Sum>i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y"
|
| 50804 | 2761 |
by auto |
2762 |
have *: "\<And>P (x1::'a) x2 s1 s2 i. |
|
2763 |
(if P i then s1 else s2) *\<^sub>R (if P i then x1 else x2) = (if P i then s1 *\<^sub>R x1 else s2 *\<^sub>R x2)" |
|
| 33175 | 2764 |
"{1..k1 + k2} \<inter> {1..k1} = {1..k1}" "{1..k1 + k2} \<inter> - {1..k1} = (\<lambda>i. i + k1) ` {1..k2}"
|
| 50804 | 2765 |
prefer 3 |
2766 |
apply (rule, rule) |
|
2767 |
unfolding image_iff |
|
2768 |
apply (rule_tac x = "x - k1" in bexI) |
|
2769 |
apply (auto simp add: not_le) |
|
2770 |
done |
|
2771 |
have inj: "inj_on (\<lambda>i. i + k1) {1..k2}"
|
|
2772 |
unfolding inj_on_def by auto |
|
2773 |
show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull" |
|
2774 |
apply rule |
|
2775 |
apply (rule_tac x="k1 + k2" in exI) |
|
2776 |
apply (rule_tac x="\<lambda>i. if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)" in exI)
|
|
2777 |
apply (rule_tac x="\<lambda>i. if i \<in> {1..k1} then x1 i else x2 (i - k1)" in exI)
|
|
2778 |
apply (rule, rule) |
|
2779 |
defer |
|
2780 |
apply rule |
|
| 64267 | 2781 |
unfolding * and sum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and |
2782 |
sum.reindex[OF inj] and o_def Collect_mem_eq |
|
2783 |
unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] sum_distrib_left[symmetric] |
|
| 50804 | 2784 |
proof - |
2785 |
fix i |
|
2786 |
assume i: "i \<in> {1..k1+k2}"
|
|
2787 |
show "0 \<le> (if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)) \<and>
|
|
2788 |
(if i \<in> {1..k1} then x1 i else x2 (i - k1)) \<in> s"
|
|
2789 |
proof (cases "i\<in>{1..k1}")
|
|
2790 |
case True |
|
2791 |
then show ?thesis |
|
| 56536 | 2792 |
using uv(1) x(1)[THEN bspec[where x=i]] by auto |
| 50804 | 2793 |
next |
2794 |
case False |
|
| 63040 | 2795 |
define j where "j = i - k1" |
| 53347 | 2796 |
from i False have "j \<in> {1..k2}"
|
2797 |
unfolding j_def by auto |
|
| 50804 | 2798 |
then show ?thesis |
| 56536 | 2799 |
using False uv(2) y(1)[THEN bspec[where x=j]] |
2800 |
by (auto simp: j_def[symmetric]) |
|
| 50804 | 2801 |
qed |
2802 |
qed (auto simp add: not_le x(2,3) y(2,3) uv(3)) |
|
| 33175 | 2803 |
qed |
2804 |
||
2805 |
lemma convex_hull_finite: |
|
2806 |
fixes s :: "'a::real_vector set" |
|
2807 |
assumes "finite s" |
|
2808 |
shows "convex hull s = {y. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and>
|
|
| 64267 | 2809 |
sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}" |
| 53339 | 2810 |
(is "?HULL = ?set") |
| 50804 | 2811 |
proof (rule hull_unique, auto simp add: convex_def[of ?set]) |
2812 |
fix x |
|
2813 |
assume "x \<in> s" |
|
| 64267 | 2814 |
then show "\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = x" |
| 50804 | 2815 |
apply (rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI) |
2816 |
apply auto |
|
| 64267 | 2817 |
unfolding sum.delta'[OF assms] and sum_delta''[OF assms] |
| 50804 | 2818 |
apply auto |
2819 |
done |
|
| 33175 | 2820 |
next |
| 50804 | 2821 |
fix u v :: real |
2822 |
assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1" |
|
| 64267 | 2823 |
fix ux assume ux: "\<forall>x\<in>s. 0 \<le> ux x" "sum ux s = (1::real)" |
2824 |
fix uy assume uy: "\<forall>x\<in>s. 0 \<le> uy x" "sum uy s = (1::real)" |
|
| 53339 | 2825 |
{
|
2826 |
fix x |
|
| 50804 | 2827 |
assume "x\<in>s" |
2828 |
then have "0 \<le> u * ux x + v * uy x" |
|
2829 |
using ux(1)[THEN bspec[where x=x]] uy(1)[THEN bspec[where x=x]] and uv(1,2) |
|
| 56536 | 2830 |
by auto |
| 50804 | 2831 |
} |
2832 |
moreover |
|
2833 |
have "(\<Sum>x\<in>s. u * ux x + v * uy x) = 1" |
|
| 64267 | 2834 |
unfolding sum.distrib and sum_distrib_left[symmetric] and ux(2) uy(2) |
| 53302 | 2835 |
using uv(3) by auto |
| 50804 | 2836 |
moreover |
2837 |
have "(\<Sum>x\<in>s. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)" |
|
| 64267 | 2838 |
unfolding scaleR_left_distrib and sum.distrib and scaleR_scaleR[symmetric] |
2839 |
and scaleR_right.sum [symmetric] |
|
| 50804 | 2840 |
by auto |
2841 |
ultimately |
|
| 64267 | 2842 |
show "\<exists>uc. (\<forall>x\<in>s. 0 \<le> uc x) \<and> sum uc s = 1 \<and> |
| 50804 | 2843 |
(\<Sum>x\<in>s. uc x *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)" |
2844 |
apply (rule_tac x="\<lambda>x. u * ux x + v * uy x" in exI) |
|
2845 |
apply auto |
|
2846 |
done |
|
| 33175 | 2847 |
next |
| 50804 | 2848 |
fix t |
2849 |
assume t: "s \<subseteq> t" "convex t" |
|
2850 |
fix u |
|
| 64267 | 2851 |
assume u: "\<forall>x\<in>s. 0 \<le> u x" "sum u s = (1::real)" |
| 50804 | 2852 |
then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> t" |
2853 |
using t(2)[unfolded convex_explicit, THEN spec[where x=s], THEN spec[where x=u]] |
|
| 33175 | 2854 |
using assms and t(1) by auto |
2855 |
qed |
|
2856 |
||
| 50804 | 2857 |
|
| 60420 | 2858 |
subsubsection \<open>Another formulation from Lars Schewe\<close> |
| 33175 | 2859 |
|
2860 |
lemma convex_hull_explicit: |
|
2861 |
fixes p :: "'a::real_vector set" |
|
| 53347 | 2862 |
shows "convex hull p = |
| 64267 | 2863 |
{y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) s = y}"
|
| 53339 | 2864 |
(is "?lhs = ?rhs") |
| 50804 | 2865 |
proof - |
| 53302 | 2866 |
{
|
2867 |
fix x |
|
2868 |
assume "x\<in>?lhs" |
|
| 50804 | 2869 |
then obtain k u y where |
| 64267 | 2870 |
obt: "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> p" "sum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
|
| 33175 | 2871 |
unfolding convex_hull_indexed by auto |
2872 |
||
| 50804 | 2873 |
have fin: "finite {1..k}" by auto
|
2874 |
have fin': "\<And>v. finite {i \<in> {1..k}. y i = v}" by auto
|
|
| 53302 | 2875 |
{
|
2876 |
fix j |
|
| 50804 | 2877 |
assume "j\<in>{1..k}"
|
| 64267 | 2878 |
then have "y j \<in> p" "0 \<le> sum u {i. Suc 0 \<le> i \<and> i \<le> k \<and> y i = y j}"
|
| 50804 | 2879 |
using obt(1)[THEN bspec[where x=j]] and obt(2) |
2880 |
apply simp |
|
| 64267 | 2881 |
apply (rule sum_nonneg) |
| 50804 | 2882 |
using obt(1) |
2883 |
apply auto |
|
2884 |
done |
|
2885 |
} |
|
| 33175 | 2886 |
moreover |
| 64267 | 2887 |
have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v}) = 1"
|
2888 |
unfolding sum_image_gen[OF fin, symmetric] using obt(2) by auto |
|
2889 |
moreover have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x"
|
|
2890 |
using sum_image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric] |
|
2891 |
unfolding scaleR_left.sum using obt(3) by auto |
|
| 50804 | 2892 |
ultimately |
| 64267 | 2893 |
have "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x" |
| 50804 | 2894 |
apply (rule_tac x="y ` {1..k}" in exI)
|
| 64267 | 2895 |
apply (rule_tac x="\<lambda>v. sum u {i\<in>{1..k}. y i = v}" in exI)
|
| 50804 | 2896 |
apply auto |
2897 |
done |
|
2898 |
then have "x\<in>?rhs" by auto |
|
2899 |
} |
|
| 33175 | 2900 |
moreover |
| 53302 | 2901 |
{
|
2902 |
fix y |
|
2903 |
assume "y\<in>?rhs" |
|
| 50804 | 2904 |
then obtain s u where |
| 64267 | 2905 |
obt: "finite s" "s \<subseteq> p" "\<forall>x\<in>s. 0 \<le> u x" "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = y" |
| 53339 | 2906 |
by auto |
| 50804 | 2907 |
|
2908 |
obtain f where f: "inj_on f {1..card s}" "f ` {1..card s} = s"
|
|
2909 |
using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto |
|
2910 |
||
| 53302 | 2911 |
{
|
2912 |
fix i :: nat |
|
| 50804 | 2913 |
assume "i\<in>{1..card s}"
|
2914 |
then have "f i \<in> s" |
|
2915 |
apply (subst f(2)[symmetric]) |
|
2916 |
apply auto |
|
2917 |
done |
|
2918 |
then have "0 \<le> u (f i)" "f i \<in> p" using obt(2,3) by auto |
|
2919 |
} |
|
| 53347 | 2920 |
moreover have *: "finite {1..card s}" by auto
|
| 53302 | 2921 |
{
|
2922 |
fix y |
|
| 50804 | 2923 |
assume "y\<in>s" |
| 53302 | 2924 |
then obtain i where "i\<in>{1..card s}" "f i = y"
|
2925 |
using f using image_iff[of y f "{1..card s}"]
|
|
| 50804 | 2926 |
by auto |
2927 |
then have "{x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = {i}"
|
|
2928 |
apply auto |
|
2929 |
using f(1)[unfolded inj_on_def] |
|
2930 |
apply(erule_tac x=x in ballE) |
|
2931 |
apply auto |
|
2932 |
done |
|
2933 |
then have "card {x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = 1" by auto
|
|
2934 |
then have "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x)) = u y"
|
|
2935 |
"(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y"
|
|
| 64267 | 2936 |
by (auto simp add: sum_constant_scaleR) |
| 50804 | 2937 |
} |
2938 |
then have "(\<Sum>x = 1..card s. u (f x)) = 1" "(\<Sum>i = 1..card s. u (f i) *\<^sub>R f i) = y" |
|
| 64267 | 2939 |
unfolding sum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f] |
2940 |
and sum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f] |
|
| 53339 | 2941 |
unfolding f |
| 64267 | 2942 |
using sum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x)" "\<lambda>v. u v *\<^sub>R v"]
|
2943 |
using sum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u]
|
|
| 53302 | 2944 |
unfolding obt(4,5) |
2945 |
by auto |
|
| 50804 | 2946 |
ultimately |
| 64267 | 2947 |
have "\<exists>k u x. (\<forall>i\<in>{1..k}. 0 \<le> u i \<and> x i \<in> p) \<and> sum u {1..k} = 1 \<and>
|
| 50804 | 2948 |
(\<Sum>i::nat = 1..k. u i *\<^sub>R x i) = y" |
2949 |
apply (rule_tac x="card s" in exI) |
|
2950 |
apply (rule_tac x="u \<circ> f" in exI) |
|
2951 |
apply (rule_tac x=f in exI) |
|
2952 |
apply fastforce |
|
2953 |
done |
|
| 53302 | 2954 |
then have "y \<in> ?lhs" |
2955 |
unfolding convex_hull_indexed by auto |
|
| 50804 | 2956 |
} |
| 53302 | 2957 |
ultimately show ?thesis |
2958 |
unfolding set_eq_iff by blast |
|
| 33175 | 2959 |
qed |
2960 |
||
| 50804 | 2961 |
|
| 60420 | 2962 |
subsubsection \<open>A stepping theorem for that expansion\<close> |
| 33175 | 2963 |
|
2964 |
lemma convex_hull_finite_step: |
|
| 50804 | 2965 |
fixes s :: "'a::real_vector set" |
2966 |
assumes "finite s" |
|
| 53302 | 2967 |
shows |
| 64267 | 2968 |
"(\<exists>u. (\<forall>x\<in>insert a s. 0 \<le> u x) \<and> sum u (insert a s) = w \<and> sum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y) |
2969 |
\<longleftrightarrow> (\<exists>v\<ge>0. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = w - v \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" |
|
| 53302 | 2970 |
(is "?lhs = ?rhs") |
| 50804 | 2971 |
proof (rule, case_tac[!] "a\<in>s") |
| 53302 | 2972 |
assume "a \<in> s" |
| 53339 | 2973 |
then have *: "insert a s = s" by auto |
| 50804 | 2974 |
assume ?lhs |
2975 |
then show ?rhs |
|
2976 |
unfolding * |
|
2977 |
apply (rule_tac x=0 in exI) |
|
2978 |
apply auto |
|
2979 |
done |
|
| 33175 | 2980 |
next |
| 50804 | 2981 |
assume ?lhs |
| 53302 | 2982 |
then obtain u where |
| 64267 | 2983 |
u: "\<forall>x\<in>insert a s. 0 \<le> u x" "sum u (insert a s) = w" "(\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y" |
| 50804 | 2984 |
by auto |
2985 |
assume "a \<notin> s" |
|
2986 |
then show ?rhs |
|
2987 |
apply (rule_tac x="u a" in exI) |
|
2988 |
using u(1)[THEN bspec[where x=a]] |
|
2989 |
apply simp |
|
2990 |
apply (rule_tac x=u in exI) |
|
| 64267 | 2991 |
using u[unfolded sum_clauses(2)[OF assms]] and \<open>a\<notin>s\<close> |
| 50804 | 2992 |
apply auto |
2993 |
done |
|
| 33175 | 2994 |
next |
| 50804 | 2995 |
assume "a \<in> s" |
2996 |
then have *: "insert a s = s" by auto |
|
2997 |
have fin: "finite (insert a s)" using assms by auto |
|
2998 |
assume ?rhs |
|
| 64267 | 2999 |
then obtain v u where uv: "v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "sum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a" |
| 50804 | 3000 |
by auto |
3001 |
show ?lhs |
|
3002 |
apply (rule_tac x = "\<lambda>x. (if a = x then v else 0) + u x" in exI) |
|
| 64267 | 3003 |
unfolding scaleR_left_distrib and sum.distrib and sum_delta''[OF fin] and sum.delta'[OF fin] |
3004 |
unfolding sum_clauses(2)[OF assms] |
|
| 60420 | 3005 |
using uv and uv(2)[THEN bspec[where x=a]] and \<open>a\<in>s\<close> |
| 50804 | 3006 |
apply auto |
3007 |
done |
|
| 33175 | 3008 |
next |
| 50804 | 3009 |
assume ?rhs |
| 53339 | 3010 |
then obtain v u where |
| 64267 | 3011 |
uv: "v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "sum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a" |
| 50804 | 3012 |
by auto |
3013 |
moreover |
|
3014 |
assume "a \<notin> s" |
|
3015 |
moreover |
|
| 64267 | 3016 |
have "(\<Sum>x\<in>s. if a = x then v else u x) = sum u s" |
| 53302 | 3017 |
and "(\<Sum>x\<in>s. (if a = x then v else u x) *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)" |
| 64267 | 3018 |
apply (rule_tac sum.cong) apply rule |
| 50804 | 3019 |
defer |
| 64267 | 3020 |
apply (rule_tac sum.cong) apply rule |
| 60420 | 3021 |
using \<open>a \<notin> s\<close> |
| 50804 | 3022 |
apply auto |
3023 |
done |
|
3024 |
ultimately show ?lhs |
|
3025 |
apply (rule_tac x="\<lambda>x. if a = x then v else u x" in exI) |
|
| 64267 | 3026 |
unfolding sum_clauses(2)[OF assms] |
| 50804 | 3027 |
apply auto |
3028 |
done |
|
3029 |
qed |
|
3030 |
||
| 33175 | 3031 |
|
| 60420 | 3032 |
subsubsection \<open>Hence some special cases\<close> |
| 33175 | 3033 |
|
3034 |
lemma convex_hull_2: |
|
3035 |
"convex hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b | u v. 0 \<le> u \<and> 0 \<le> v \<and> u + v = 1}"
|
|
| 53302 | 3036 |
proof - |
3037 |
have *: "\<And>u. (\<forall>x\<in>{a, b}. 0 \<le> u x) \<longleftrightarrow> 0 \<le> u a \<and> 0 \<le> u b"
|
|
3038 |
by auto |
|
3039 |
have **: "finite {b}" by auto
|
|
3040 |
show ?thesis |
|
3041 |
apply (simp add: convex_hull_finite) |
|
3042 |
unfolding convex_hull_finite_step[OF **, of a 1, unfolded * conj_assoc] |
|
3043 |
apply auto |
|
3044 |
apply (rule_tac x=v in exI) |
|
3045 |
apply (rule_tac x="1 - v" in exI) |
|
3046 |
apply simp |
|
3047 |
apply (rule_tac x=u in exI) |
|
3048 |
apply simp |
|
3049 |
apply (rule_tac x="\<lambda>x. v" in exI) |
|
3050 |
apply simp |
|
3051 |
done |
|
3052 |
qed |
|
| 33175 | 3053 |
|
3054 |
lemma convex_hull_2_alt: "convex hull {a,b} = {a + u *\<^sub>R (b - a) | u. 0 \<le> u \<and> u \<le> 1}"
|
|
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44142
diff
changeset
|
3055 |
unfolding convex_hull_2 |
| 53302 | 3056 |
proof (rule Collect_cong) |
3057 |
have *: "\<And>x y ::real. x + y = 1 \<longleftrightarrow> x = 1 - y" |
|
3058 |
by auto |
|
3059 |
fix x |
|
3060 |
show "(\<exists>v u. x = v *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> v \<and> 0 \<le> u \<and> v + u = 1) \<longleftrightarrow> |
|
3061 |
(\<exists>u. x = a + u *\<^sub>R (b - a) \<and> 0 \<le> u \<and> u \<le> 1)" |
|
3062 |
unfolding * |
|
3063 |
apply auto |
|
3064 |
apply (rule_tac[!] x=u in exI) |
|
3065 |
apply (auto simp add: algebra_simps) |
|
3066 |
done |
|
3067 |
qed |
|
| 33175 | 3068 |
|
3069 |
lemma convex_hull_3: |
|
3070 |
"convex hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c | u v w. 0 \<le> u \<and> 0 \<le> v \<and> 0 \<le> w \<and> u + v + w = 1}"
|
|
| 53302 | 3071 |
proof - |
3072 |
have fin: "finite {a,b,c}" "finite {b,c}" "finite {c}"
|
|
3073 |
by auto |
|
3074 |
have *: "\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" |
|
|
44361
75ec83d45303
remove unnecessary euclidean_space class constraints
huffman
parents:
44349
diff
changeset
|
3075 |
by (auto simp add: field_simps) |
| 53302 | 3076 |
show ?thesis |
3077 |
unfolding convex_hull_finite[OF fin(1)] and convex_hull_finite_step[OF fin(2)] and * |
|
3078 |
unfolding convex_hull_finite_step[OF fin(3)] |
|
3079 |
apply (rule Collect_cong) |
|
3080 |
apply simp |
|
3081 |
apply auto |
|
3082 |
apply (rule_tac x=va in exI) |
|
3083 |
apply (rule_tac x="u c" in exI) |
|
3084 |
apply simp |
|
3085 |
apply (rule_tac x="1 - v - w" in exI) |
|
3086 |
apply simp |
|
3087 |
apply (rule_tac x=v in exI) |
|
3088 |
apply simp |
|
3089 |
apply (rule_tac x="\<lambda>x. w" in exI) |
|
3090 |
apply simp |
|
3091 |
done |
|
3092 |
qed |
|
| 33175 | 3093 |
|
3094 |
lemma convex_hull_3_alt: |
|
3095 |
"convex hull {a,b,c} = {a + u *\<^sub>R (b - a) + v *\<^sub>R (c - a) | u v. 0 \<le> u \<and> 0 \<le> v \<and> u + v \<le> 1}"
|
|
| 53302 | 3096 |
proof - |
3097 |
have *: "\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" |
|
3098 |
by auto |
|
3099 |
show ?thesis |
|
3100 |
unfolding convex_hull_3 |
|
3101 |
apply (auto simp add: *) |
|
3102 |
apply (rule_tac x=v in exI) |
|
3103 |
apply (rule_tac x=w in exI) |
|
3104 |
apply (simp add: algebra_simps) |
|
3105 |
apply (rule_tac x=u in exI) |
|
3106 |
apply (rule_tac x=v in exI) |
|
3107 |
apply (simp add: algebra_simps) |
|
3108 |
done |
|
3109 |
qed |
|
3110 |
||
| 33175 | 3111 |
|
| 60420 | 3112 |
subsection \<open>Relations among closure notions and corresponding hulls\<close> |
| 33175 | 3113 |
|
3114 |
lemma affine_imp_convex: "affine s \<Longrightarrow> convex s" |
|
3115 |
unfolding affine_def convex_def by auto |
|
3116 |
||
| 64394 | 3117 |
lemma convex_affine_hull [simp]: "convex (affine hull S)" |
3118 |
by (simp add: affine_imp_convex) |
|
3119 |
||
|
44361
75ec83d45303
remove unnecessary euclidean_space class constraints
huffman
parents:
44349
diff
changeset
|
3120 |
lemma subspace_imp_convex: "subspace s \<Longrightarrow> convex s" |
| 33175 | 3121 |
using subspace_imp_affine affine_imp_convex by auto |
3122 |
||
|
44361
75ec83d45303
remove unnecessary euclidean_space class constraints
huffman
parents:
44349
diff
changeset
|
3123 |
lemma affine_hull_subset_span: "(affine hull s) \<subseteq> (span s)" |
| 53302 | 3124 |
by (metis hull_minimal span_inc subspace_imp_affine subspace_span) |
| 33175 | 3125 |
|
|
44361
75ec83d45303
remove unnecessary euclidean_space class constraints
huffman
parents:
44349
diff
changeset
|
3126 |
lemma convex_hull_subset_span: "(convex hull s) \<subseteq> (span s)" |
| 53302 | 3127 |
by (metis hull_minimal span_inc subspace_imp_convex subspace_span) |
| 33175 | 3128 |
|
3129 |
lemma convex_hull_subset_affine_hull: "(convex hull s) \<subseteq> (affine hull s)" |
|
| 53302 | 3130 |
by (metis affine_affine_hull affine_imp_convex hull_minimal hull_subset) |
3131 |
||
3132 |
lemma affine_dependent_imp_dependent: "affine_dependent s \<Longrightarrow> dependent s" |
|
| 49531 | 3133 |
unfolding affine_dependent_def dependent_def |
| 33175 | 3134 |
using affine_hull_subset_span by auto |
3135 |
||
3136 |
lemma dependent_imp_affine_dependent: |
|
| 53302 | 3137 |
assumes "dependent {x - a| x . x \<in> s}"
|
3138 |
and "a \<notin> s" |
|
| 33175 | 3139 |
shows "affine_dependent (insert a s)" |
| 53302 | 3140 |
proof - |
| 49531 | 3141 |
from assms(1)[unfolded dependent_explicit] obtain S u v |
| 53347 | 3142 |
where obt: "finite S" "S \<subseteq> {x - a |x. x \<in> s}" "v\<in>S" "u v \<noteq> 0" "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0"
|
3143 |
by auto |
|
| 63040 | 3144 |
define t where "t = (\<lambda>x. x + a) ` S" |
| 33175 | 3145 |
|
| 53347 | 3146 |
have inj: "inj_on (\<lambda>x. x + a) S" |
| 53302 | 3147 |
unfolding inj_on_def by auto |
3148 |
have "0 \<notin> S" |
|
3149 |
using obt(2) assms(2) unfolding subset_eq by auto |
|
| 53347 | 3150 |
have fin: "finite t" and "t \<subseteq> s" |
| 53302 | 3151 |
unfolding t_def using obt(1,2) by auto |
3152 |
then have "finite (insert a t)" and "insert a t \<subseteq> insert a s" |
|
3153 |
by auto |
|
3154 |
moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x)) = (\<Sum>x\<in>t. Q x)" |
|
| 64267 | 3155 |
apply (rule sum.cong) |
| 60420 | 3156 |
using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> |
| 53302 | 3157 |
apply auto |
3158 |
done |
|
| 33175 | 3159 |
have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0" |
| 64267 | 3160 |
unfolding sum_clauses(2)[OF fin] |
| 60420 | 3161 |
using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> |
| 53302 | 3162 |
apply auto |
3163 |
unfolding * |
|
3164 |
apply auto |
|
3165 |
done |
|
| 33175 | 3166 |
moreover have "\<exists>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) \<noteq> 0" |
| 53302 | 3167 |
apply (rule_tac x="v + a" in bexI) |
| 60420 | 3168 |
using obt(3,4) and \<open>0\<notin>S\<close> |
| 53302 | 3169 |
unfolding t_def |
3170 |
apply auto |
|
3171 |
done |
|
3172 |
moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x) *\<^sub>R x) = (\<Sum>x\<in>t. Q x *\<^sub>R x)" |
|
| 64267 | 3173 |
apply (rule sum.cong) |
| 60420 | 3174 |
using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> |
| 53302 | 3175 |
apply auto |
3176 |
done |
|
| 49531 | 3177 |
have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)" |
| 64267 | 3178 |
unfolding scaleR_left.sum |
3179 |
unfolding t_def and sum.reindex[OF inj] and o_def |
|
| 53302 | 3180 |
using obt(5) |
| 64267 | 3181 |
by (auto simp add: sum.distrib scaleR_right_distrib) |
| 53302 | 3182 |
then have "(\<Sum>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) *\<^sub>R v) = 0" |
| 64267 | 3183 |
unfolding sum_clauses(2)[OF fin] |
| 60420 | 3184 |
using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> |
| 53302 | 3185 |
by (auto simp add: *) |
3186 |
ultimately show ?thesis |
|
3187 |
unfolding affine_dependent_explicit |
|
3188 |
apply (rule_tac x="insert a t" in exI) |
|
3189 |
apply auto |
|
3190 |
done |
|
| 33175 | 3191 |
qed |
3192 |
||
3193 |
lemma convex_cone: |
|
| 53302 | 3194 |
"convex s \<and> cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. (x + y) \<in> s) \<and> (\<forall>x\<in>s. \<forall>c\<ge>0. (c *\<^sub>R x) \<in> s)" |
3195 |
(is "?lhs = ?rhs") |
|
3196 |
proof - |
|
3197 |
{
|
|
3198 |
fix x y |
|
3199 |
assume "x\<in>s" "y\<in>s" and ?lhs |
|
3200 |
then have "2 *\<^sub>R x \<in>s" "2 *\<^sub>R y \<in> s" |
|
3201 |
unfolding cone_def by auto |
|
3202 |
then have "x + y \<in> s" |
|
| 60420 | 3203 |
using \<open>?lhs\<close>[unfolded convex_def, THEN conjunct1] |
| 53302 | 3204 |
apply (erule_tac x="2*\<^sub>R x" in ballE) |
3205 |
apply (erule_tac x="2*\<^sub>R y" in ballE) |
|
3206 |
apply (erule_tac x="1/2" in allE) |
|
3207 |
apply simp |
|
3208 |
apply (erule_tac x="1/2" in allE) |
|
3209 |
apply auto |
|
3210 |
done |
|
3211 |
} |
|
3212 |
then show ?thesis |
|
3213 |
unfolding convex_def cone_def by blast |
|
3214 |
qed |
|
3215 |
||
3216 |
lemma affine_dependent_biggerset: |
|
| 53347 | 3217 |
fixes s :: "'a::euclidean_space set" |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
3218 |
assumes "finite s" "card s \<ge> DIM('a) + 2"
|
| 33175 | 3219 |
shows "affine_dependent s" |
| 53302 | 3220 |
proof - |
3221 |
have "s \<noteq> {}" using assms by auto
|
|
3222 |
then obtain a where "a\<in>s" by auto |
|
3223 |
have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})"
|
|
3224 |
by auto |
|
3225 |
have "card {x - a |x. x \<in> s - {a}} = card (s - {a})"
|
|
3226 |
unfolding * |
|
3227 |
apply (rule card_image) |
|
3228 |
unfolding inj_on_def |
|
3229 |
apply auto |
|
3230 |
done |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
3231 |
also have "\<dots> > DIM('a)" using assms(2)
|
| 60420 | 3232 |
unfolding card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>] by auto |
| 53302 | 3233 |
finally show ?thesis |
| 60420 | 3234 |
apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric]) |
| 53302 | 3235 |
apply (rule dependent_imp_affine_dependent) |
3236 |
apply (rule dependent_biggerset) |
|
3237 |
apply auto |
|
3238 |
done |
|
3239 |
qed |
|
| 33175 | 3240 |
|
3241 |
lemma affine_dependent_biggerset_general: |
|
| 53347 | 3242 |
assumes "finite (s :: 'a::euclidean_space set)" |
3243 |
and "card s \<ge> dim s + 2" |
|
| 33175 | 3244 |
shows "affine_dependent s" |
| 53302 | 3245 |
proof - |
| 33175 | 3246 |
from assms(2) have "s \<noteq> {}" by auto
|
3247 |
then obtain a where "a\<in>s" by auto |
|
| 53302 | 3248 |
have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})"
|
3249 |
by auto |
|
3250 |
have **: "card {x - a |x. x \<in> s - {a}} = card (s - {a})"
|
|
3251 |
unfolding * |
|
3252 |
apply (rule card_image) |
|
3253 |
unfolding inj_on_def |
|
3254 |
apply auto |
|
3255 |
done |
|
| 33175 | 3256 |
have "dim {x - a |x. x \<in> s - {a}} \<le> dim s"
|
| 53302 | 3257 |
apply (rule subset_le_dim) |
3258 |
unfolding subset_eq |
|
| 60420 | 3259 |
using \<open>a\<in>s\<close> |
| 63938 | 3260 |
apply (auto simp add:span_superset span_diff) |
| 53302 | 3261 |
done |
| 33175 | 3262 |
also have "\<dots> < dim s + 1" by auto |
| 53302 | 3263 |
also have "\<dots> \<le> card (s - {a})"
|
3264 |
using assms |
|
| 60420 | 3265 |
using card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>] |
| 53302 | 3266 |
by auto |
3267 |
finally show ?thesis |
|
| 60420 | 3268 |
apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric]) |
| 53302 | 3269 |
apply (rule dependent_imp_affine_dependent) |
3270 |
apply (rule dependent_biggerset_general) |
|
3271 |
unfolding ** |
|
3272 |
apply auto |
|
3273 |
done |
|
3274 |
qed |
|
3275 |
||
| 33175 | 3276 |
|
| 60420 | 3277 |
subsection \<open>Some Properties of Affine Dependent Sets\<close> |
| 40377 | 3278 |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
3279 |
lemma affine_independent_0: "\<not> affine_dependent {}"
|
| 40377 | 3280 |
by (simp add: affine_dependent_def) |
3281 |
||
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
3282 |
lemma affine_independent_1: "\<not> affine_dependent {a}"
|
| 53302 | 3283 |
by (simp add: affine_dependent_def) |
3284 |
||
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
3285 |
lemma affine_independent_2: "\<not> affine_dependent {a,b}"
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
3286 |
by (simp add: affine_dependent_def insert_Diff_if hull_same) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
3287 |
|
| 53302 | 3288 |
lemma affine_hull_translation: "affine hull ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` (affine hull S)" |
3289 |
proof - |
|
3290 |
have "affine ((\<lambda>x. a + x) ` (affine hull S))" |
|
| 60303 | 3291 |
using affine_translation affine_affine_hull by blast |
| 53347 | 3292 |
moreover have "(\<lambda>x. a + x) ` S \<subseteq> (\<lambda>x. a + x) ` (affine hull S)" |
| 53302 | 3293 |
using hull_subset[of S] by auto |
| 53347 | 3294 |
ultimately have h1: "affine hull ((\<lambda>x. a + x) ` S) \<subseteq> (\<lambda>x. a + x) ` (affine hull S)" |
| 53302 | 3295 |
by (metis hull_minimal) |
3296 |
have "affine((\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) ` S)))" |
|
| 60303 | 3297 |
using affine_translation affine_affine_hull by blast |
| 53347 | 3298 |
moreover have "(\<lambda>x. -a + x) ` (\<lambda>x. a + x) ` S \<subseteq> (\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) ` S))" |
| 53302 | 3299 |
using hull_subset[of "(\<lambda>x. a + x) ` S"] by auto |
| 53347 | 3300 |
moreover have "S = (\<lambda>x. -a + x) ` (\<lambda>x. a + x) ` S" |
| 53302 | 3301 |
using translation_assoc[of "-a" a] by auto |
3302 |
ultimately have "(\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) ` S)) >= (affine hull S)" |
|
3303 |
by (metis hull_minimal) |
|
3304 |
then have "affine hull ((\<lambda>x. a + x) ` S) >= (\<lambda>x. a + x) ` (affine hull S)" |
|
3305 |
by auto |
|
| 54465 | 3306 |
then show ?thesis using h1 by auto |
| 40377 | 3307 |
qed |
3308 |
||
3309 |
lemma affine_dependent_translation: |
|
3310 |
assumes "affine_dependent S" |
|
| 53339 | 3311 |
shows "affine_dependent ((\<lambda>x. a + x) ` S)" |
| 53302 | 3312 |
proof - |
| 54465 | 3313 |
obtain x where x: "x \<in> S \<and> x \<in> affine hull (S - {x})"
|
| 53302 | 3314 |
using assms affine_dependent_def by auto |
3315 |
have "op + a ` (S - {x}) = op + a ` S - {a + x}"
|
|
3316 |
by auto |
|
| 53347 | 3317 |
then have "a + x \<in> affine hull ((\<lambda>x. a + x) ` S - {a + x})"
|
| 54465 | 3318 |
using affine_hull_translation[of a "S - {x}"] x by auto
|
| 53347 | 3319 |
moreover have "a + x \<in> (\<lambda>x. a + x) ` S" |
| 54465 | 3320 |
using x by auto |
| 53302 | 3321 |
ultimately show ?thesis |
3322 |
unfolding affine_dependent_def by auto |
|
| 40377 | 3323 |
qed |
3324 |
||
3325 |
lemma affine_dependent_translation_eq: |
|
| 54465 | 3326 |
"affine_dependent S \<longleftrightarrow> affine_dependent ((\<lambda>x. a + x) ` S)" |
| 53302 | 3327 |
proof - |
3328 |
{
|
|
| 53339 | 3329 |
assume "affine_dependent ((\<lambda>x. a + x) ` S)" |
| 53302 | 3330 |
then have "affine_dependent S" |
| 53339 | 3331 |
using affine_dependent_translation[of "((\<lambda>x. a + x) ` S)" "-a"] translation_assoc[of "-a" a] |
| 53302 | 3332 |
by auto |
3333 |
} |
|
3334 |
then show ?thesis |
|
3335 |
using affine_dependent_translation by auto |
|
| 40377 | 3336 |
qed |
3337 |
||
3338 |
lemma affine_hull_0_dependent: |
|
| 53339 | 3339 |
assumes "0 \<in> affine hull S" |
| 40377 | 3340 |
shows "dependent S" |
| 53302 | 3341 |
proof - |
| 64267 | 3342 |
obtain s u where s_u: "finite s \<and> s \<noteq> {} \<and> s \<subseteq> S \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
|
| 53302 | 3343 |
using assms affine_hull_explicit[of S] by auto |
| 53339 | 3344 |
then have "\<exists>v\<in>s. u v \<noteq> 0" |
| 64267 | 3345 |
using sum_not_0[of "u" "s"] by auto |
| 53339 | 3346 |
then have "finite s \<and> s \<subseteq> S \<and> (\<exists>v\<in>s. u v \<noteq> 0 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0)" |
| 54465 | 3347 |
using s_u by auto |
| 53302 | 3348 |
then show ?thesis |
3349 |
unfolding dependent_explicit[of S] by auto |
|
| 40377 | 3350 |
qed |
3351 |
||
3352 |
lemma affine_dependent_imp_dependent2: |
|
3353 |
assumes "affine_dependent (insert 0 S)" |
|
3354 |
shows "dependent S" |
|
| 53302 | 3355 |
proof - |
| 54465 | 3356 |
obtain x where x: "x \<in> insert 0 S \<and> x \<in> affine hull (insert 0 S - {x})"
|
| 53302 | 3357 |
using affine_dependent_def[of "(insert 0 S)"] assms by blast |
3358 |
then have "x \<in> span (insert 0 S - {x})"
|
|
3359 |
using affine_hull_subset_span by auto |
|
3360 |
moreover have "span (insert 0 S - {x}) = span (S - {x})"
|
|
3361 |
using insert_Diff_if[of "0" S "{x}"] span_insert_0[of "S-{x}"] by auto
|
|
3362 |
ultimately have "x \<in> span (S - {x})" by auto
|
|
3363 |
then have "x \<noteq> 0 \<Longrightarrow> dependent S" |
|
| 54465 | 3364 |
using x dependent_def by auto |
| 53302 | 3365 |
moreover |
3366 |
{
|
|
3367 |
assume "x = 0" |
|
3368 |
then have "0 \<in> affine hull S" |
|
| 54465 | 3369 |
using x hull_mono[of "S - {0}" S] by auto
|
| 53302 | 3370 |
then have "dependent S" |
3371 |
using affine_hull_0_dependent by auto |
|
3372 |
} |
|
3373 |
ultimately show ?thesis by auto |
|
| 40377 | 3374 |
qed |
3375 |
||
3376 |
lemma affine_dependent_iff_dependent: |
|
| 53302 | 3377 |
assumes "a \<notin> S" |
3378 |
shows "affine_dependent (insert a S) \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` S)" |
|
3379 |
proof - |
|
3380 |
have "(op + (- a) ` S) = {x - a| x . x : S}" by auto
|
|
3381 |
then show ?thesis |
|
3382 |
using affine_dependent_translation_eq[of "(insert a S)" "-a"] |
|
| 49531 | 3383 |
affine_dependent_imp_dependent2 assms |
| 53302 | 3384 |
dependent_imp_affine_dependent[of a S] |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53676
diff
changeset
|
3385 |
by (auto simp del: uminus_add_conv_diff) |
| 40377 | 3386 |
qed |
3387 |
||
3388 |
lemma affine_dependent_iff_dependent2: |
|
| 53339 | 3389 |
assumes "a \<in> S" |
3390 |
shows "affine_dependent S \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` (S-{a}))"
|
|
| 53302 | 3391 |
proof - |
| 53339 | 3392 |
have "insert a (S - {a}) = S"
|
| 53302 | 3393 |
using assms by auto |
3394 |
then show ?thesis |
|
3395 |
using assms affine_dependent_iff_dependent[of a "S-{a}"] by auto
|
|
| 40377 | 3396 |
qed |
3397 |
||
3398 |
lemma affine_hull_insert_span_gen: |
|
| 53339 | 3399 |
"affine hull (insert a s) = (\<lambda>x. a + x) ` span ((\<lambda>x. - a + x) ` s)" |
| 53302 | 3400 |
proof - |
| 53339 | 3401 |
have h1: "{x - a |x. x \<in> s} = ((\<lambda>x. -a+x) ` s)"
|
| 53302 | 3402 |
by auto |
3403 |
{
|
|
3404 |
assume "a \<notin> s" |
|
3405 |
then have ?thesis |
|
3406 |
using affine_hull_insert_span[of a s] h1 by auto |
|
3407 |
} |
|
3408 |
moreover |
|
3409 |
{
|
|
3410 |
assume a1: "a \<in> s" |
|
| 53339 | 3411 |
have "\<exists>x. x \<in> s \<and> -a+x=0" |
| 53302 | 3412 |
apply (rule exI[of _ a]) |
3413 |
using a1 |
|
3414 |
apply auto |
|
3415 |
done |
|
| 53339 | 3416 |
then have "insert 0 ((\<lambda>x. -a+x) ` (s - {a})) = (\<lambda>x. -a+x) ` s"
|
| 53302 | 3417 |
by auto |
| 53339 | 3418 |
then have "span ((\<lambda>x. -a+x) ` (s - {a}))=span ((\<lambda>x. -a+x) ` s)"
|
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53676
diff
changeset
|
3419 |
using span_insert_0[of "op + (- a) ` (s - {a})"] by (auto simp del: uminus_add_conv_diff)
|
| 53339 | 3420 |
moreover have "{x - a |x. x \<in> (s - {a})} = ((\<lambda>x. -a+x) ` (s - {a}))"
|
| 53302 | 3421 |
by auto |
| 53339 | 3422 |
moreover have "insert a (s - {a}) = insert a s"
|
| 63092 | 3423 |
by auto |
| 53302 | 3424 |
ultimately have ?thesis |
| 63092 | 3425 |
using affine_hull_insert_span[of "a" "s-{a}"] by auto
|
| 53302 | 3426 |
} |
3427 |
ultimately show ?thesis by auto |
|
| 40377 | 3428 |
qed |
3429 |
||
3430 |
lemma affine_hull_span2: |
|
| 53302 | 3431 |
assumes "a \<in> s" |
3432 |
shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` (s-{a}))"
|
|
3433 |
using affine_hull_insert_span_gen[of a "s - {a}", unfolded insert_Diff[OF assms]]
|
|
3434 |
by auto |
|
| 40377 | 3435 |
|
3436 |
lemma affine_hull_span_gen: |
|
| 53339 | 3437 |
assumes "a \<in> affine hull s" |
3438 |
shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` s)" |
|
| 53302 | 3439 |
proof - |
3440 |
have "affine hull (insert a s) = affine hull s" |
|
3441 |
using hull_redundant[of a affine s] assms by auto |
|
3442 |
then show ?thesis |
|
3443 |
using affine_hull_insert_span_gen[of a "s"] by auto |
|
| 40377 | 3444 |
qed |
3445 |
||
3446 |
lemma affine_hull_span_0: |
|
| 53339 | 3447 |
assumes "0 \<in> affine hull S" |
| 40377 | 3448 |
shows "affine hull S = span S" |
| 53302 | 3449 |
using affine_hull_span_gen[of "0" S] assms by auto |
| 40377 | 3450 |
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3451 |
lemma extend_to_affine_basis_nonempty: |
| 53339 | 3452 |
fixes S V :: "'n::euclidean_space set" |
3453 |
assumes "\<not> affine_dependent S" "S \<subseteq> V" "S \<noteq> {}"
|
|
3454 |
shows "\<exists>T. \<not> affine_dependent T \<and> S \<subseteq> T \<and> T \<subseteq> V \<and> affine hull T = affine hull V" |
|
| 53302 | 3455 |
proof - |
| 54465 | 3456 |
obtain a where a: "a \<in> S" |
| 53302 | 3457 |
using assms by auto |
| 53339 | 3458 |
then have h0: "independent ((\<lambda>x. -a + x) ` (S-{a}))"
|
| 53302 | 3459 |
using affine_dependent_iff_dependent2 assms by auto |
| 54465 | 3460 |
then obtain B where B: |
| 53339 | 3461 |
"(\<lambda>x. -a+x) ` (S - {a}) \<subseteq> B \<and> B \<subseteq> (\<lambda>x. -a+x) ` V \<and> independent B \<and> (\<lambda>x. -a+x) ` V \<subseteq> span B"
|
3462 |
using maximal_independent_subset_extend[of "(\<lambda>x. -a+x) ` (S-{a})" "(\<lambda>x. -a + x) ` V"] assms
|
|
| 53302 | 3463 |
by blast |
| 63040 | 3464 |
define T where "T = (\<lambda>x. a+x) ` insert 0 B" |
| 53339 | 3465 |
then have "T = insert a ((\<lambda>x. a+x) ` B)" |
3466 |
by auto |
|
3467 |
then have "affine hull T = (\<lambda>x. a+x) ` span B" |
|
3468 |
using affine_hull_insert_span_gen[of a "((\<lambda>x. a+x) ` B)"] translation_assoc[of "-a" a B] |
|
| 53302 | 3469 |
by auto |
| 53347 | 3470 |
then have "V \<subseteq> affine hull T" |
| 54465 | 3471 |
using B assms translation_inverse_subset[of a V "span B"] |
| 53302 | 3472 |
by auto |
| 53339 | 3473 |
moreover have "T \<subseteq> V" |
| 54465 | 3474 |
using T_def B a assms by auto |
| 53302 | 3475 |
ultimately have "affine hull T = affine hull V" |
|
44457
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
huffman
parents:
44365
diff
changeset
|
3476 |
by (metis Int_absorb1 Int_absorb2 hull_hull hull_mono) |
| 53347 | 3477 |
moreover have "S \<subseteq> T" |
| 54465 | 3478 |
using T_def B translation_inverse_subset[of a "S-{a}" B]
|
| 53302 | 3479 |
by auto |
3480 |
moreover have "\<not> affine_dependent T" |
|
| 53339 | 3481 |
using T_def affine_dependent_translation_eq[of "insert 0 B"] |
| 54465 | 3482 |
affine_dependent_imp_dependent2 B |
| 53302 | 3483 |
by auto |
| 60420 | 3484 |
ultimately show ?thesis using \<open>T \<subseteq> V\<close> by auto |
| 40377 | 3485 |
qed |
3486 |
||
| 49531 | 3487 |
lemma affine_basis_exists: |
| 53339 | 3488 |
fixes V :: "'n::euclidean_space set" |
3489 |
shows "\<exists>B. B \<subseteq> V \<and> \<not> affine_dependent B \<and> affine hull V = affine hull B" |
|
| 53302 | 3490 |
proof (cases "V = {}")
|
3491 |
case True |
|
3492 |
then show ?thesis |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
3493 |
using affine_independent_0 by auto |
| 53302 | 3494 |
next |
3495 |
case False |
|
3496 |
then obtain x where "x \<in> V" by auto |
|
3497 |
then show ?thesis |
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3498 |
using affine_dependent_def[of "{x}"] extend_to_affine_basis_nonempty[of "{x}" V]
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3499 |
by auto |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3500 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3501 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3502 |
proposition extend_to_affine_basis: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3503 |
fixes S V :: "'n::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3504 |
assumes "\<not> affine_dependent S" "S \<subseteq> V" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3505 |
obtains T where "\<not> affine_dependent T" "S \<subseteq> T" "T \<subseteq> V" "affine hull T = affine hull V" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3506 |
proof (cases "S = {}")
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3507 |
case True then show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3508 |
using affine_basis_exists by (metis empty_subsetI that) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3509 |
next |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3510 |
case False |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3511 |
then show ?thesis by (metis assms extend_to_affine_basis_nonempty that) |
| 53302 | 3512 |
qed |
3513 |
||
| 40377 | 3514 |
|
| 60420 | 3515 |
subsection \<open>Affine Dimension of a Set\<close> |
| 40377 | 3516 |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
3517 |
definition aff_dim :: "('a::euclidean_space) set \<Rightarrow> int"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
3518 |
where "aff_dim V = |
| 53339 | 3519 |
(SOME d :: int. |
3520 |
\<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = d + 1)" |
|
| 40377 | 3521 |
|
3522 |
lemma aff_dim_basis_exists: |
|
| 49531 | 3523 |
fixes V :: "('n::euclidean_space) set"
|
| 53339 | 3524 |
shows "\<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1" |
| 53302 | 3525 |
proof - |
| 53347 | 3526 |
obtain B where "\<not> affine_dependent B \<and> affine hull B = affine hull V" |
| 53302 | 3527 |
using affine_basis_exists[of V] by auto |
3528 |
then show ?thesis |
|
| 53339 | 3529 |
unfolding aff_dim_def |
| 53347 | 3530 |
some_eq_ex[of "\<lambda>d. \<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = d + 1"] |
| 53302 | 3531 |
apply auto |
| 53339 | 3532 |
apply (rule exI[of _ "int (card B) - (1 :: int)"]) |
| 53302 | 3533 |
apply (rule exI[of _ "B"]) |
3534 |
apply auto |
|
3535 |
done |
|
3536 |
qed |
|
3537 |
||
3538 |
lemma affine_hull_nonempty: "S \<noteq> {} \<longleftrightarrow> affine hull S \<noteq> {}"
|
|
3539 |
proof - |
|
3540 |
have "S = {} \<Longrightarrow> affine hull S = {}"
|
|
3541 |
using affine_hull_empty by auto |
|
3542 |
moreover have "affine hull S = {} \<Longrightarrow> S = {}"
|
|
3543 |
unfolding hull_def by auto |
|
3544 |
ultimately show ?thesis by blast |
|
| 40377 | 3545 |
qed |
3546 |
||
3547 |
lemma aff_dim_parallel_subspace_aux: |
|
| 53347 | 3548 |
fixes B :: "'n::euclidean_space set" |
| 53302 | 3549 |
assumes "\<not> affine_dependent B" "a \<in> B" |
| 53339 | 3550 |
shows "finite B \<and> ((card B) - 1 = dim (span ((\<lambda>x. -a+x) ` (B-{a}))))"
|
| 53302 | 3551 |
proof - |
| 53339 | 3552 |
have "independent ((\<lambda>x. -a + x) ` (B-{a}))"
|
| 53302 | 3553 |
using affine_dependent_iff_dependent2 assms by auto |
| 53339 | 3554 |
then have fin: "dim (span ((\<lambda>x. -a+x) ` (B-{a}))) = card ((\<lambda>x. -a + x) ` (B-{a}))"
|
3555 |
"finite ((\<lambda>x. -a + x) ` (B - {a}))"
|
|
| 53347 | 3556 |
using indep_card_eq_dim_span[of "(\<lambda>x. -a+x) ` (B-{a})"] by auto
|
| 53302 | 3557 |
show ?thesis |
| 53339 | 3558 |
proof (cases "(\<lambda>x. -a + x) ` (B - {a}) = {}")
|
| 53302 | 3559 |
case True |
| 53339 | 3560 |
have "B = insert a ((\<lambda>x. a + x) ` (\<lambda>x. -a + x) ` (B - {a}))"
|
| 53302 | 3561 |
using translation_assoc[of "a" "-a" "(B - {a})"] assms by auto
|
| 53339 | 3562 |
then have "B = {a}" using True by auto
|
| 53302 | 3563 |
then show ?thesis using assms fin by auto |
3564 |
next |
|
3565 |
case False |
|
| 53339 | 3566 |
then have "card ((\<lambda>x. -a + x) ` (B - {a})) > 0"
|
| 53302 | 3567 |
using fin by auto |
| 53339 | 3568 |
moreover have h1: "card ((\<lambda>x. -a + x) ` (B-{a})) = card (B-{a})"
|
| 53302 | 3569 |
apply (rule card_image) |
3570 |
using translate_inj_on |
|
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
53676
diff
changeset
|
3571 |
apply (auto simp del: uminus_add_conv_diff) |
| 53302 | 3572 |
done |
| 53339 | 3573 |
ultimately have "card (B-{a}) > 0" by auto
|
3574 |
then have *: "finite (B - {a})"
|
|
| 53302 | 3575 |
using card_gt_0_iff[of "(B - {a})"] by auto
|
| 53339 | 3576 |
then have "card (B - {a}) = card B - 1"
|
| 53302 | 3577 |
using card_Diff_singleton assms by auto |
3578 |
with * show ?thesis using fin h1 by auto |
|
3579 |
qed |
|
| 40377 | 3580 |
qed |
3581 |
||
3582 |
lemma aff_dim_parallel_subspace: |
|
| 53339 | 3583 |
fixes V L :: "'n::euclidean_space set" |
| 53302 | 3584 |
assumes "V \<noteq> {}"
|
| 53339 | 3585 |
and "subspace L" |
3586 |
and "affine_parallel (affine hull V) L" |
|
| 53302 | 3587 |
shows "aff_dim V = int (dim L)" |
3588 |
proof - |
|
| 53339 | 3589 |
obtain B where |
| 54465 | 3590 |
B: "affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> int (card B) = aff_dim V + 1" |
| 53302 | 3591 |
using aff_dim_basis_exists by auto |
3592 |
then have "B \<noteq> {}"
|
|
| 54465 | 3593 |
using assms B affine_hull_nonempty[of V] affine_hull_nonempty[of B] |
| 53302 | 3594 |
by auto |
| 54465 | 3595 |
then obtain a where a: "a \<in> B" by auto |
| 63040 | 3596 |
define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))"
|
| 40377 | 3597 |
moreover have "affine_parallel (affine hull B) Lb" |
| 54465 | 3598 |
using Lb_def B assms affine_hull_span2[of a B] a |
| 53339 | 3599 |
affine_parallel_commut[of "Lb" "(affine hull B)"] |
3600 |
unfolding affine_parallel_def |
|
3601 |
by auto |
|
| 53302 | 3602 |
moreover have "subspace Lb" |
3603 |
using Lb_def subspace_span by auto |
|
3604 |
moreover have "affine hull B \<noteq> {}"
|
|
| 54465 | 3605 |
using assms B affine_hull_nonempty[of V] by auto |
| 53302 | 3606 |
ultimately have "L = Lb" |
| 54465 | 3607 |
using assms affine_parallel_subspace[of "affine hull B"] affine_affine_hull[of B] B |
| 53302 | 3608 |
by auto |
| 53339 | 3609 |
then have "dim L = dim Lb" |
3610 |
by auto |
|
3611 |
moreover have "card B - 1 = dim Lb" and "finite B" |
|
| 54465 | 3612 |
using Lb_def aff_dim_parallel_subspace_aux a B by auto |
| 53302 | 3613 |
ultimately show ?thesis |
| 60420 | 3614 |
using B \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto
|
| 40377 | 3615 |
qed |
3616 |
||
3617 |
lemma aff_independent_finite: |
|
| 53339 | 3618 |
fixes B :: "'n::euclidean_space set" |
3619 |
assumes "\<not> affine_dependent B" |
|
| 53302 | 3620 |
shows "finite B" |
3621 |
proof - |
|
3622 |
{
|
|
3623 |
assume "B \<noteq> {}"
|
|
3624 |
then obtain a where "a \<in> B" by auto |
|
3625 |
then have ?thesis |
|
3626 |
using aff_dim_parallel_subspace_aux assms by auto |
|
3627 |
} |
|
3628 |
then show ?thesis by auto |
|
| 40377 | 3629 |
qed |
3630 |
||
3631 |
lemma independent_finite: |
|
| 53339 | 3632 |
fixes B :: "'n::euclidean_space set" |
| 53302 | 3633 |
assumes "independent B" |
3634 |
shows "finite B" |
|
3635 |
using affine_dependent_imp_dependent[of B] aff_independent_finite[of B] assms |
|
3636 |
by auto |
|
| 40377 | 3637 |
|
3638 |
lemma subspace_dim_equal: |
|
| 53339 | 3639 |
assumes "subspace (S :: ('n::euclidean_space) set)"
|
3640 |
and "subspace T" |
|
3641 |
and "S \<subseteq> T" |
|
3642 |
and "dim S \<ge> dim T" |
|
| 53302 | 3643 |
shows "S = T" |
3644 |
proof - |
|
| 53347 | 3645 |
obtain B where B: "B \<le> S" "independent B \<and> S \<subseteq> span B" "card B = dim S" |
| 53339 | 3646 |
using basis_exists[of S] by auto |
3647 |
then have "span B \<subseteq> S" |
|
3648 |
using span_mono[of B S] span_eq[of S] assms by metis |
|
3649 |
then have "span B = S" |
|
| 53347 | 3650 |
using B by auto |
| 53339 | 3651 |
have "dim S = dim T" |
3652 |
using assms dim_subset[of S T] by auto |
|
3653 |
then have "T \<subseteq> span B" |
|
| 53347 | 3654 |
using card_eq_dim[of B T] B independent_finite assms by auto |
| 53339 | 3655 |
then show ?thesis |
| 60420 | 3656 |
using assms \<open>span B = S\<close> by auto |
| 40377 | 3657 |
qed |
3658 |
||
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3659 |
corollary dim_eq_span: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3660 |
fixes S :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3661 |
shows "\<lbrakk>S \<subseteq> T; dim T \<le> dim S\<rbrakk> \<Longrightarrow> span S = span T" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3662 |
by (simp add: span_mono subspace_dim_equal subspace_span) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3663 |
|
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
3664 |
lemma dim_eq_full: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
3665 |
fixes S :: "'a :: euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
3666 |
shows "dim S = DIM('a) \<longleftrightarrow> span S = UNIV"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
3667 |
apply (rule iffI) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
3668 |
apply (metis dim_eq_span dim_subset_UNIV span_Basis span_span subset_UNIV) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
3669 |
by (metis dim_UNIV dim_span) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
3670 |
|
|
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
|
3671 |
lemma span_substd_basis: |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
3672 |
assumes d: "d \<subseteq> Basis" |
| 53347 | 3673 |
shows "span d = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
|
3674 |
(is "_ = ?B") |
|
| 53339 | 3675 |
proof - |
3676 |
have "d \<subseteq> ?B" |
|
3677 |
using d by (auto simp: inner_Basis) |
|
3678 |
moreover have s: "subspace ?B" |
|
3679 |
using subspace_substandard[of "\<lambda>i. i \<notin> d"] . |
|
3680 |
ultimately have "span d \<subseteq> ?B" |
|
3681 |
using span_mono[of d "?B"] span_eq[of "?B"] by blast |
|
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53348
diff
changeset
|
3682 |
moreover have *: "card d \<le> dim (span d)" |
| 53339 | 3683 |
using independent_card_le_dim[of d "span d"] independent_substdbasis[OF assms] span_inc[of d] |
3684 |
by auto |
|
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53348
diff
changeset
|
3685 |
moreover from * have "dim ?B \<le> dim (span d)" |
| 53339 | 3686 |
using dim_substandard[OF assms] by auto |
3687 |
ultimately show ?thesis |
|
3688 |
using s subspace_dim_equal[of "span d" "?B"] subspace_span[of d] by auto |
|
| 40377 | 3689 |
qed |
3690 |
||
3691 |
lemma basis_to_substdbasis_subspace_isomorphism: |
|
| 53339 | 3692 |
fixes B :: "'a::euclidean_space set" |
3693 |
assumes "independent B" |
|
3694 |
shows "\<exists>f d::'a set. card d = card B \<and> linear f \<and> f ` B = d \<and> |
|
3695 |
f ` span B = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} \<and> inj_on f (span B) \<and> d \<subseteq> Basis"
|
|
3696 |
proof - |
|
3697 |
have B: "card B = dim B" |
|
3698 |
using dim_unique[of B B "card B"] assms span_inc[of B] by auto |
|
3699 |
have "dim B \<le> card (Basis :: 'a set)" |
|
3700 |
using dim_subset_UNIV[of B] by simp |
|
3701 |
from ex_card[OF this] obtain d :: "'a set" where d: "d \<subseteq> Basis" and t: "card d = dim B" |
|
3702 |
by auto |
|
| 53347 | 3703 |
let ?t = "{x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
|
| 53339 | 3704 |
have "\<exists>f. linear f \<and> f ` B = d \<and> f ` span B = ?t \<and> inj_on f (span B)" |
|
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
|
3705 |
apply (rule basis_to_basis_subspace_isomorphism[of "span B" ?t B "d"]) |
| 53339 | 3706 |
apply (rule subspace_span) |
3707 |
apply (rule subspace_substandard) |
|
3708 |
defer |
|
3709 |
apply (rule span_inc) |
|
3710 |
apply (rule assms) |
|
3711 |
defer |
|
3712 |
unfolding dim_span[of B] |
|
3713 |
apply(rule B) |
|
| 54465 | 3714 |
unfolding span_substd_basis[OF d, symmetric] |
| 53339 | 3715 |
apply (rule span_inc) |
3716 |
apply (rule independent_substdbasis[OF d]) |
|
3717 |
apply rule |
|
3718 |
apply assumption |
|
3719 |
unfolding t[symmetric] span_substd_basis[OF d] dim_substandard[OF d] |
|
3720 |
apply auto |
|
3721 |
done |
|
| 60420 | 3722 |
with t \<open>card B = dim B\<close> d show ?thesis by auto |
| 40377 | 3723 |
qed |
3724 |
||
3725 |
lemma aff_dim_empty: |
|
| 53339 | 3726 |
fixes S :: "'n::euclidean_space set" |
3727 |
shows "S = {} \<longleftrightarrow> aff_dim S = -1"
|
|
3728 |
proof - |
|
3729 |
obtain B where *: "affine hull B = affine hull S" |
|
3730 |
and "\<not> affine_dependent B" |
|
3731 |
and "int (card B) = aff_dim S + 1" |
|
3732 |
using aff_dim_basis_exists by auto |
|
3733 |
moreover |
|
3734 |
from * have "S = {} \<longleftrightarrow> B = {}"
|
|
3735 |
using affine_hull_nonempty[of B] affine_hull_nonempty[of S] by auto |
|
3736 |
ultimately show ?thesis |
|
3737 |
using aff_independent_finite[of B] card_gt_0_iff[of B] by auto |
|
3738 |
qed |
|
3739 |
||
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
3740 |
lemma aff_dim_empty_eq [simp]: "aff_dim ({}::'a::euclidean_space set) = -1"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
3741 |
by (simp add: aff_dim_empty [symmetric]) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
3742 |
|
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
3743 |
lemma aff_dim_affine_hull [simp]: "aff_dim (affine hull S) = aff_dim S" |
| 53339 | 3744 |
unfolding aff_dim_def using hull_hull[of _ S] by auto |
| 40377 | 3745 |
|
3746 |
lemma aff_dim_affine_hull2: |
|
| 53339 | 3747 |
assumes "affine hull S = affine hull T" |
3748 |
shows "aff_dim S = aff_dim T" |
|
3749 |
unfolding aff_dim_def using assms by auto |
|
| 40377 | 3750 |
|
| 49531 | 3751 |
lemma aff_dim_unique: |
| 53339 | 3752 |
fixes B V :: "'n::euclidean_space set" |
3753 |
assumes "affine hull B = affine hull V \<and> \<not> affine_dependent B" |
|
3754 |
shows "of_nat (card B) = aff_dim V + 1" |
|
3755 |
proof (cases "B = {}")
|
|
3756 |
case True |
|
3757 |
then have "V = {}"
|
|
3758 |
using affine_hull_nonempty[of V] affine_hull_nonempty[of B] assms |
|
3759 |
by auto |
|
3760 |
then have "aff_dim V = (-1::int)" |
|
3761 |
using aff_dim_empty by auto |
|
3762 |
then show ?thesis |
|
| 60420 | 3763 |
using \<open>B = {}\<close> by auto
|
| 53339 | 3764 |
next |
3765 |
case False |
|
| 54465 | 3766 |
then obtain a where a: "a \<in> B" by auto |
| 63040 | 3767 |
define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))"
|
| 40377 | 3768 |
have "affine_parallel (affine hull B) Lb" |
| 54465 | 3769 |
using Lb_def affine_hull_span2[of a B] a |
| 53339 | 3770 |
affine_parallel_commut[of "Lb" "(affine hull B)"] |
3771 |
unfolding affine_parallel_def by auto |
|
3772 |
moreover have "subspace Lb" |
|
3773 |
using Lb_def subspace_span by auto |
|
3774 |
ultimately have "aff_dim B = int(dim Lb)" |
|
| 60420 | 3775 |
using aff_dim_parallel_subspace[of B Lb] \<open>B \<noteq> {}\<close> by auto
|
| 53339 | 3776 |
moreover have "(card B) - 1 = dim Lb" "finite B" |
| 54465 | 3777 |
using Lb_def aff_dim_parallel_subspace_aux a assms by auto |
| 53339 | 3778 |
ultimately have "of_nat (card B) = aff_dim B + 1" |
| 60420 | 3779 |
using \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto
|
| 53339 | 3780 |
then show ?thesis |
3781 |
using aff_dim_affine_hull2 assms by auto |
|
| 40377 | 3782 |
qed |
3783 |
||
| 49531 | 3784 |
lemma aff_dim_affine_independent: |
| 53339 | 3785 |
fixes B :: "'n::euclidean_space set" |
3786 |
assumes "\<not> affine_dependent B" |
|
3787 |
shows "of_nat (card B) = aff_dim B + 1" |
|
| 40377 | 3788 |
using aff_dim_unique[of B B] assms by auto |
3789 |
||
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
3790 |
lemma affine_independent_iff_card: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
3791 |
fixes s :: "'a::euclidean_space set" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
3792 |
shows "~ affine_dependent s \<longleftrightarrow> finite s \<and> aff_dim s = int(card s) - 1" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
3793 |
apply (rule iffI) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
3794 |
apply (simp add: aff_dim_affine_independent aff_independent_finite) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
3795 |
by (metis affine_basis_exists [of s] aff_dim_unique card_subset_eq diff_add_cancel of_nat_eq_iff) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
3796 |
|
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
3797 |
lemma aff_dim_sing [simp]: |
| 53339 | 3798 |
fixes a :: "'n::euclidean_space" |
3799 |
shows "aff_dim {a} = 0"
|
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
3800 |
using aff_dim_affine_independent[of "{a}"] affine_independent_1 by auto
|
| 40377 | 3801 |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
3802 |
lemma aff_dim_2 [simp]: "aff_dim {a,b} = (if a = b then 0 else 1)"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
3803 |
proof (clarsimp) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
3804 |
assume "a \<noteq> b" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
3805 |
then have "aff_dim{a,b} = card{a,b} - 1"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
3806 |
using affine_independent_2 [of a b] aff_dim_affine_independent by fastforce |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
3807 |
also have "... = 1" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
3808 |
using \<open>a \<noteq> b\<close> by simp |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
3809 |
finally show "aff_dim {a, b} = 1" .
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
3810 |
qed |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
3811 |
|
| 40377 | 3812 |
lemma aff_dim_inner_basis_exists: |
| 49531 | 3813 |
fixes V :: "('n::euclidean_space) set"
|
| 53339 | 3814 |
shows "\<exists>B. B \<subseteq> V \<and> affine hull B = affine hull V \<and> |
3815 |
\<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1" |
|
3816 |
proof - |
|
| 53347 | 3817 |
obtain B where B: "\<not> affine_dependent B" "B \<subseteq> V" "affine hull B = affine hull V" |
| 53339 | 3818 |
using affine_basis_exists[of V] by auto |
3819 |
then have "of_nat(card B) = aff_dim V+1" using aff_dim_unique by auto |
|
| 53347 | 3820 |
with B show ?thesis by auto |
| 40377 | 3821 |
qed |
3822 |
||
3823 |
lemma aff_dim_le_card: |
|
| 53347 | 3824 |
fixes V :: "'n::euclidean_space set" |
| 53339 | 3825 |
assumes "finite V" |
| 53347 | 3826 |
shows "aff_dim V \<le> of_nat (card V) - 1" |
| 53339 | 3827 |
proof - |
| 53347 | 3828 |
obtain B where B: "B \<subseteq> V" "of_nat (card B) = aff_dim V + 1" |
| 53339 | 3829 |
using aff_dim_inner_basis_exists[of V] by auto |
3830 |
then have "card B \<le> card V" |
|
3831 |
using assms card_mono by auto |
|
| 53347 | 3832 |
with B show ?thesis by auto |
| 40377 | 3833 |
qed |
3834 |
||
3835 |
lemma aff_dim_parallel_eq: |
|
| 53339 | 3836 |
fixes S T :: "'n::euclidean_space set" |
3837 |
assumes "affine_parallel (affine hull S) (affine hull T)" |
|
3838 |
shows "aff_dim S = aff_dim T" |
|
3839 |
proof - |
|
3840 |
{
|
|
3841 |
assume "T \<noteq> {}" "S \<noteq> {}"
|
|
| 53347 | 3842 |
then obtain L where L: "subspace L \<and> affine_parallel (affine hull T) L" |
3843 |
using affine_parallel_subspace[of "affine hull T"] |
|
3844 |
affine_affine_hull[of T] affine_hull_nonempty |
|
| 53339 | 3845 |
by auto |
3846 |
then have "aff_dim T = int (dim L)" |
|
| 60420 | 3847 |
using aff_dim_parallel_subspace \<open>T \<noteq> {}\<close> by auto
|
| 53339 | 3848 |
moreover have *: "subspace L \<and> affine_parallel (affine hull S) L" |
| 53347 | 3849 |
using L affine_parallel_assoc[of "affine hull S" "affine hull T" L] assms by auto |
| 53339 | 3850 |
moreover from * have "aff_dim S = int (dim L)" |
| 60420 | 3851 |
using aff_dim_parallel_subspace \<open>S \<noteq> {}\<close> by auto
|
| 53339 | 3852 |
ultimately have ?thesis by auto |
3853 |
} |
|
3854 |
moreover |
|
3855 |
{
|
|
3856 |
assume "S = {}"
|
|
3857 |
then have "S = {}" and "T = {}"
|
|
3858 |
using assms affine_hull_nonempty |
|
3859 |
unfolding affine_parallel_def |
|
3860 |
by auto |
|
3861 |
then have ?thesis using aff_dim_empty by auto |
|
3862 |
} |
|
3863 |
moreover |
|
3864 |
{
|
|
3865 |
assume "T = {}"
|
|
3866 |
then have "S = {}" and "T = {}"
|
|
3867 |
using assms affine_hull_nonempty |
|
3868 |
unfolding affine_parallel_def |
|
3869 |
by auto |
|
3870 |
then have ?thesis |
|
3871 |
using aff_dim_empty by auto |
|
3872 |
} |
|
3873 |
ultimately show ?thesis by blast |
|
| 40377 | 3874 |
qed |
3875 |
||
3876 |
lemma aff_dim_translation_eq: |
|
| 53339 | 3877 |
fixes a :: "'n::euclidean_space" |
3878 |
shows "aff_dim ((\<lambda>x. a + x) ` S) = aff_dim S" |
|
3879 |
proof - |
|
| 53347 | 3880 |
have "affine_parallel (affine hull S) (affine hull ((\<lambda>x. a + x) ` S))" |
| 53339 | 3881 |
unfolding affine_parallel_def |
3882 |
apply (rule exI[of _ "a"]) |
|
3883 |
using affine_hull_translation[of a S] |
|
3884 |
apply auto |
|
3885 |
done |
|
3886 |
then show ?thesis |
|
3887 |
using aff_dim_parallel_eq[of S "(\<lambda>x. a + x) ` S"] by auto |
|
| 40377 | 3888 |
qed |
3889 |
||
3890 |
lemma aff_dim_affine: |
|
| 53339 | 3891 |
fixes S L :: "'n::euclidean_space set" |
3892 |
assumes "S \<noteq> {}"
|
|
3893 |
and "affine S" |
|
3894 |
and "subspace L" |
|
3895 |
and "affine_parallel S L" |
|
3896 |
shows "aff_dim S = int (dim L)" |
|
3897 |
proof - |
|
3898 |
have *: "affine hull S = S" |
|
3899 |
using assms affine_hull_eq[of S] by auto |
|
3900 |
then have "affine_parallel (affine hull S) L" |
|
3901 |
using assms by (simp add: *) |
|
3902 |
then show ?thesis |
|
3903 |
using assms aff_dim_parallel_subspace[of S L] by blast |
|
| 40377 | 3904 |
qed |
3905 |
||
3906 |
lemma dim_affine_hull: |
|
| 53339 | 3907 |
fixes S :: "'n::euclidean_space set" |
3908 |
shows "dim (affine hull S) = dim S" |
|
3909 |
proof - |
|
3910 |
have "dim (affine hull S) \<ge> dim S" |
|
3911 |
using dim_subset by auto |
|
3912 |
moreover have "dim (span S) \<ge> dim (affine hull S)" |
|
| 60303 | 3913 |
using dim_subset affine_hull_subset_span by blast |
| 53339 | 3914 |
moreover have "dim (span S) = dim S" |
3915 |
using dim_span by auto |
|
3916 |
ultimately show ?thesis by auto |
|
| 40377 | 3917 |
qed |
3918 |
||
3919 |
lemma aff_dim_subspace: |
|
| 53339 | 3920 |
fixes S :: "'n::euclidean_space set" |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3921 |
assumes "subspace S" |
| 53339 | 3922 |
shows "aff_dim S = int (dim S)" |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3923 |
proof (cases "S={}")
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3924 |
case True with assms show ?thesis |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3925 |
by (simp add: subspace_affine) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3926 |
next |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3927 |
case False |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3928 |
with aff_dim_affine[of S S] assms subspace_imp_affine[of S] affine_parallel_reflex[of S] subspace_affine |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3929 |
show ?thesis by auto |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3930 |
qed |
| 40377 | 3931 |
|
3932 |
lemma aff_dim_zero: |
|
| 53339 | 3933 |
fixes S :: "'n::euclidean_space set" |
3934 |
assumes "0 \<in> affine hull S" |
|
3935 |
shows "aff_dim S = int (dim S)" |
|
3936 |
proof - |
|
3937 |
have "subspace (affine hull S)" |
|
3938 |
using subspace_affine[of "affine hull S"] affine_affine_hull assms |
|
3939 |
by auto |
|
3940 |
then have "aff_dim (affine hull S) = int (dim (affine hull S))" |
|
3941 |
using assms aff_dim_subspace[of "affine hull S"] by auto |
|
3942 |
then show ?thesis |
|
3943 |
using aff_dim_affine_hull[of S] dim_affine_hull[of S] |
|
3944 |
by auto |
|
| 40377 | 3945 |
qed |
3946 |
||
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3947 |
lemma aff_dim_eq_dim: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3948 |
fixes S :: "'n::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3949 |
assumes "a \<in> affine hull S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3950 |
shows "aff_dim S = int (dim ((\<lambda>x. -a+x) ` S))" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3951 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3952 |
have "0 \<in> affine hull ((\<lambda>x. -a+x) ` S)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3953 |
unfolding Convex_Euclidean_Space.affine_hull_translation |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3954 |
using assms by (simp add: ab_group_add_class.ab_left_minus image_iff) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3955 |
with aff_dim_zero show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3956 |
by (metis aff_dim_translation_eq) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3957 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
3958 |
|
| 63072 | 3959 |
lemma aff_dim_UNIV [simp]: "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))"
|
| 53347 | 3960 |
using aff_dim_subspace[of "(UNIV :: 'n::euclidean_space set)"] |
| 53339 | 3961 |
dim_UNIV[where 'a="'n::euclidean_space"] |
3962 |
by auto |
|
| 40377 | 3963 |
|
3964 |
lemma aff_dim_geq: |
|
| 53339 | 3965 |
fixes V :: "'n::euclidean_space set" |
3966 |
shows "aff_dim V \<ge> -1" |
|
3967 |
proof - |
|
| 53347 | 3968 |
obtain B where "affine hull B = affine hull V" |
3969 |
and "\<not> affine_dependent B" |
|
3970 |
and "int (card B) = aff_dim V + 1" |
|
| 53339 | 3971 |
using aff_dim_basis_exists by auto |
3972 |
then show ?thesis by auto |
|
| 40377 | 3973 |
qed |
3974 |
||
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
3975 |
lemma aff_dim_negative_iff [simp]: |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
3976 |
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
|
3977 |
shows "aff_dim S < 0 \<longleftrightarrow>S = {}"
|
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
3978 |
by (metis aff_dim_empty aff_dim_geq diff_0 eq_iff zle_diff1_eq) |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
3979 |
|
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3980 |
lemma affine_independent_card_dim_diffs: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3981 |
fixes S :: "'a :: euclidean_space set" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3982 |
assumes "~ affine_dependent S" "a \<in> S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3983 |
shows "card S = dim {x - a|x. x \<in> S} + 1"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3984 |
proof - |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3985 |
have 1: "{b - a|b. b \<in> (S - {a})} \<subseteq> {x - a|x. x \<in> S}" by auto
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3986 |
have 2: "x - a \<in> span {b - a |b. b \<in> S - {a}}" if "x \<in> S" for x
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3987 |
proof (cases "x = a") |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3988 |
case True then show ?thesis by simp |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3989 |
next |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3990 |
case False then show ?thesis |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3991 |
using assms by (blast intro: span_superset that) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3992 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3993 |
have "\<not> affine_dependent (insert a S)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3994 |
by (simp add: assms insert_absorb) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3995 |
then have 3: "independent {b - a |b. b \<in> S - {a}}"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3996 |
using dependent_imp_affine_dependent by fastforce |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3997 |
have "{b - a |b. b \<in> S - {a}} = (\<lambda>b. b-a) ` (S - {a})"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3998 |
by blast |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
3999 |
then have "card {b - a |b. b \<in> S - {a}} = card ((\<lambda>b. b-a) ` (S - {a}))"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4000 |
by simp |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4001 |
also have "... = card (S - {a})"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4002 |
by (metis (no_types, lifting) card_image diff_add_cancel inj_onI) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4003 |
also have "... = card S - 1" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4004 |
by (simp add: aff_independent_finite assms) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4005 |
finally have 4: "card {b - a |b. b \<in> S - {a}} = card S - 1" .
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4006 |
have "finite S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4007 |
by (meson assms aff_independent_finite) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4008 |
with \<open>a \<in> S\<close> have "card S \<noteq> 0" by auto |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4009 |
moreover have "dim {x - a |x. x \<in> S} = card S - 1"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4010 |
using 2 by (blast intro: dim_unique [OF 1 _ 3 4]) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4011 |
ultimately show ?thesis |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4012 |
by auto |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4013 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4014 |
|
| 49531 | 4015 |
lemma independent_card_le_aff_dim: |
| 53347 | 4016 |
fixes B :: "'n::euclidean_space set" |
4017 |
assumes "B \<subseteq> V" |
|
| 53339 | 4018 |
assumes "\<not> affine_dependent B" |
4019 |
shows "int (card B) \<le> aff_dim V + 1" |
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
4020 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
4021 |
obtain T where T: "\<not> affine_dependent T \<and> B \<subseteq> T \<and> T \<subseteq> V \<and> affine hull T = affine hull V" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
4022 |
by (metis assms extend_to_affine_basis[of B V]) |
| 53339 | 4023 |
then have "of_nat (card T) = aff_dim V + 1" |
4024 |
using aff_dim_unique by auto |
|
4025 |
then show ?thesis |
|
| 53347 | 4026 |
using T card_mono[of T B] aff_independent_finite[of T] by auto |
| 40377 | 4027 |
qed |
4028 |
||
4029 |
lemma aff_dim_subset: |
|
| 53347 | 4030 |
fixes S T :: "'n::euclidean_space set" |
4031 |
assumes "S \<subseteq> T" |
|
4032 |
shows "aff_dim S \<le> aff_dim T" |
|
| 53339 | 4033 |
proof - |
| 53347 | 4034 |
obtain B where B: "\<not> affine_dependent B" "B \<subseteq> S" "affine hull B = affine hull S" |
4035 |
"of_nat (card B) = aff_dim S + 1" |
|
| 53339 | 4036 |
using aff_dim_inner_basis_exists[of S] by auto |
4037 |
then have "int (card B) \<le> aff_dim T + 1" |
|
4038 |
using assms independent_card_le_aff_dim[of B T] by auto |
|
| 53347 | 4039 |
with B show ?thesis by auto |
| 40377 | 4040 |
qed |
4041 |
||
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4042 |
lemma aff_dim_le_DIM: |
| 53339 | 4043 |
fixes S :: "'n::euclidean_space set" |
4044 |
shows "aff_dim S \<le> int (DIM('n))"
|
|
| 49531 | 4045 |
proof - |
| 53339 | 4046 |
have "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))"
|
| 63072 | 4047 |
using aff_dim_UNIV by auto |
| 53339 | 4048 |
then show "aff_dim (S:: 'n::euclidean_space set) \<le> int(DIM('n))"
|
| 63092 | 4049 |
using aff_dim_subset[of S "(UNIV :: ('n::euclidean_space) set)"] subset_UNIV by auto
|
| 40377 | 4050 |
qed |
4051 |
||
4052 |
lemma affine_dim_equal: |
|
| 53347 | 4053 |
fixes S :: "'n::euclidean_space set" |
4054 |
assumes "affine S" "affine T" "S \<noteq> {}" "S \<subseteq> T" "aff_dim S = aff_dim T"
|
|
4055 |
shows "S = T" |
|
4056 |
proof - |
|
4057 |
obtain a where "a \<in> S" using assms by auto |
|
4058 |
then have "a \<in> T" using assms by auto |
|
| 63040 | 4059 |
define LS where "LS = {y. \<exists>x \<in> S. (-a) + x = y}"
|
| 53347 | 4060 |
then have ls: "subspace LS" "affine_parallel S LS" |
| 60420 | 4061 |
using assms parallel_subspace_explicit[of S a LS] \<open>a \<in> S\<close> by auto |
| 53347 | 4062 |
then have h1: "int(dim LS) = aff_dim S" |
4063 |
using assms aff_dim_affine[of S LS] by auto |
|
4064 |
have "T \<noteq> {}" using assms by auto
|
|
| 63040 | 4065 |
define LT where "LT = {y. \<exists>x \<in> T. (-a) + x = y}"
|
| 53347 | 4066 |
then have lt: "subspace LT \<and> affine_parallel T LT" |
| 60420 | 4067 |
using assms parallel_subspace_explicit[of T a LT] \<open>a \<in> T\<close> by auto |
| 53347 | 4068 |
then have "int(dim LT) = aff_dim T" |
| 60420 | 4069 |
using assms aff_dim_affine[of T LT] \<open>T \<noteq> {}\<close> by auto
|
| 53347 | 4070 |
then have "dim LS = dim LT" |
4071 |
using h1 assms by auto |
|
4072 |
moreover have "LS \<le> LT" |
|
4073 |
using LS_def LT_def assms by auto |
|
4074 |
ultimately have "LS = LT" |
|
4075 |
using subspace_dim_equal[of LS LT] ls lt by auto |
|
4076 |
moreover have "S = {x. \<exists>y \<in> LS. a+y=x}"
|
|
4077 |
using LS_def by auto |
|
4078 |
moreover have "T = {x. \<exists>y \<in> LT. a+y=x}"
|
|
4079 |
using LT_def by auto |
|
4080 |
ultimately show ?thesis by auto |
|
| 40377 | 4081 |
qed |
4082 |
||
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4083 |
lemma aff_dim_eq_0: |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4084 |
fixes S :: "'a::euclidean_space set" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4085 |
shows "aff_dim S = 0 \<longleftrightarrow> (\<exists>a. S = {a})"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4086 |
proof (cases "S = {}")
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4087 |
case True |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4088 |
then show ?thesis |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4089 |
by auto |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4090 |
next |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4091 |
case False |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4092 |
then obtain a where "a \<in> S" by auto |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4093 |
show ?thesis |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4094 |
proof safe |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4095 |
assume 0: "aff_dim S = 0" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4096 |
have "~ {a,b} \<subseteq> S" if "b \<noteq> a" for b
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4097 |
by (metis "0" aff_dim_2 aff_dim_subset not_one_le_zero that) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4098 |
then show "\<exists>a. S = {a}"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4099 |
using \<open>a \<in> S\<close> by blast |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4100 |
qed auto |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4101 |
qed |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4102 |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4103 |
lemma affine_hull_UNIV: |
| 53347 | 4104 |
fixes S :: "'n::euclidean_space set" |
4105 |
assumes "aff_dim S = int(DIM('n))"
|
|
4106 |
shows "affine hull S = (UNIV :: ('n::euclidean_space) set)"
|
|
4107 |
proof - |
|
4108 |
have "S \<noteq> {}"
|
|
4109 |
using assms aff_dim_empty[of S] by auto |
|
4110 |
have h0: "S \<subseteq> affine hull S" |
|
4111 |
using hull_subset[of S _] by auto |
|
4112 |
have h1: "aff_dim (UNIV :: ('n::euclidean_space) set) = aff_dim S"
|
|
| 63072 | 4113 |
using aff_dim_UNIV assms by auto |
| 53347 | 4114 |
then have h2: "aff_dim (affine hull S) \<le> aff_dim (UNIV :: ('n::euclidean_space) set)"
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4115 |
using aff_dim_le_DIM[of "affine hull S"] assms h0 by auto |
| 53347 | 4116 |
have h3: "aff_dim S \<le> aff_dim (affine hull S)" |
4117 |
using h0 aff_dim_subset[of S "affine hull S"] assms by auto |
|
4118 |
then have h4: "aff_dim (affine hull S) = aff_dim (UNIV :: ('n::euclidean_space) set)"
|
|
4119 |
using h0 h1 h2 by auto |
|
4120 |
then show ?thesis |
|
4121 |
using affine_dim_equal[of "affine hull S" "(UNIV :: ('n::euclidean_space) set)"]
|
|
| 60420 | 4122 |
affine_affine_hull[of S] affine_UNIV assms h4 h0 \<open>S \<noteq> {}\<close>
|
| 53347 | 4123 |
by auto |
| 40377 | 4124 |
qed |
4125 |
||
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4126 |
lemma disjoint_affine_hull: |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4127 |
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
|
4128 |
assumes "~ affine_dependent s" "t \<subseteq> s" "u \<subseteq> s" "t \<inter> u = {}"
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4129 |
shows "(affine hull t) \<inter> (affine hull u) = {}"
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4130 |
proof - |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4131 |
have "finite s" using assms by (simp add: aff_independent_finite) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4132 |
then have "finite t" "finite u" using assms finite_subset by blast+ |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4133 |
{ fix y
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4134 |
assume yt: "y \<in> affine hull t" and yu: "y \<in> affine hull u" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4135 |
then obtain a b |
| 64267 | 4136 |
where a1 [simp]: "sum a t = 1" and [simp]: "sum (\<lambda>v. a v *\<^sub>R v) t = y" |
4137 |
and [simp]: "sum b u = 1" "sum (\<lambda>v. b v *\<^sub>R v) u = y" |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4138 |
by (auto simp: affine_hull_finite \<open>finite t\<close> \<open>finite u\<close>) |
| 63040 | 4139 |
define c where "c x = (if x \<in> t then a x else if x \<in> u then -(b x) else 0)" for x |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4140 |
have [simp]: "s \<inter> t = t" "s \<inter> - t \<inter> u = u" using assms by auto |
| 64267 | 4141 |
have "sum c s = 0" |
4142 |
by (simp add: c_def comm_monoid_add_class.sum.If_cases \<open>finite s\<close> sum_negf) |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4143 |
moreover have "~ (\<forall>v\<in>s. c v = 0)" |
| 64267 | 4144 |
by (metis (no_types) IntD1 \<open>s \<inter> t = t\<close> a1 c_def sum_not_0 zero_neq_one) |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4145 |
moreover have "(\<Sum>v\<in>s. c v *\<^sub>R v) = 0" |
| 64267 | 4146 |
by (simp add: c_def if_smult sum_negf |
4147 |
comm_monoid_add_class.sum.If_cases \<open>finite s\<close>) |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4148 |
ultimately have False |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4149 |
using assms \<open>finite s\<close> by (auto simp: affine_dependent_explicit) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4150 |
} |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4151 |
then show ?thesis by blast |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4152 |
qed |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4153 |
|
| 40377 | 4154 |
lemma aff_dim_convex_hull: |
| 53347 | 4155 |
fixes S :: "'n::euclidean_space set" |
4156 |
shows "aff_dim (convex hull S) = aff_dim S" |
|
| 49531 | 4157 |
using aff_dim_affine_hull[of S] convex_hull_subset_affine_hull[of S] |
| 53347 | 4158 |
hull_subset[of S "convex"] aff_dim_subset[of S "convex hull S"] |
4159 |
aff_dim_subset[of "convex hull S" "affine hull S"] |
|
4160 |
by auto |
|
| 40377 | 4161 |
|
4162 |
lemma aff_dim_cball: |
|
| 53347 | 4163 |
fixes a :: "'n::euclidean_space" |
4164 |
assumes "e > 0" |
|
4165 |
shows "aff_dim (cball a e) = int (DIM('n))"
|
|
4166 |
proof - |
|
4167 |
have "(\<lambda>x. a + x) ` (cball 0 e) \<subseteq> cball a e" |
|
4168 |
unfolding cball_def dist_norm by auto |
|
4169 |
then have "aff_dim (cball (0 :: 'n::euclidean_space) e) \<le> aff_dim (cball a e)" |
|
4170 |
using aff_dim_translation_eq[of a "cball 0 e"] |
|
4171 |
aff_dim_subset[of "op + a ` cball 0 e" "cball a e"] |
|
4172 |
by auto |
|
4173 |
moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))"
|
|
4174 |
using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"] |
|
4175 |
centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms |
|
4176 |
by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"]) |
|
4177 |
ultimately show ?thesis |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4178 |
using aff_dim_le_DIM[of "cball a e"] by auto |
| 40377 | 4179 |
qed |
4180 |
||
4181 |
lemma aff_dim_open: |
|
| 53347 | 4182 |
fixes S :: "'n::euclidean_space set" |
4183 |
assumes "open S" |
|
4184 |
and "S \<noteq> {}"
|
|
4185 |
shows "aff_dim S = int (DIM('n))"
|
|
4186 |
proof - |
|
4187 |
obtain x where "x \<in> S" |
|
4188 |
using assms by auto |
|
4189 |
then obtain e where e: "e > 0" "cball x e \<subseteq> S" |
|
4190 |
using open_contains_cball[of S] assms by auto |
|
4191 |
then have "aff_dim (cball x e) \<le> aff_dim S" |
|
4192 |
using aff_dim_subset by auto |
|
4193 |
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
|
4194 |
using aff_dim_cball[of e x] aff_dim_le_DIM[of S] by auto |
| 40377 | 4195 |
qed |
4196 |
||
4197 |
lemma low_dim_interior: |
|
| 53347 | 4198 |
fixes S :: "'n::euclidean_space set" |
4199 |
assumes "\<not> aff_dim S = int (DIM('n))"
|
|
4200 |
shows "interior S = {}"
|
|
4201 |
proof - |
|
4202 |
have "aff_dim(interior S) \<le> aff_dim S" |
|
4203 |
using interior_subset aff_dim_subset[of "interior S" S] by auto |
|
4204 |
then show ?thesis |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4205 |
using aff_dim_open[of "interior S"] aff_dim_le_DIM[of S] assms by auto |
| 40377 | 4206 |
qed |
4207 |
||
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
4208 |
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
|
4209 |
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
|
4210 |
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
|
4211 |
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
|
4212 |
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
4213 |
corollary aff_dim_nonempty_interior: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
4214 |
fixes S :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
4215 |
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
|
4216 |
by (metis low_dim_interior) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
4217 |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
4218 |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4219 |
subsection \<open>Caratheodory's theorem.\<close> |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4220 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4221 |
lemma convex_hull_caratheodory_aff_dim: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4222 |
fixes p :: "('a::euclidean_space) set"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4223 |
shows "convex hull p = |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4224 |
{y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and>
|
| 64267 | 4225 |
(\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) s = y}" |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4226 |
unfolding convex_hull_explicit set_eq_iff mem_Collect_eq |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4227 |
proof (intro allI iffI) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4228 |
fix y |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4229 |
let ?P = "\<lambda>n. \<exists>s u. finite s \<and> card s = n \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> |
| 64267 | 4230 |
sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y" |
4231 |
assume "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y" |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4232 |
then obtain N where "?P N" by auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4233 |
then have "\<exists>n\<le>N. (\<forall>k<n. \<not> ?P k) \<and> ?P n" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4234 |
apply (rule_tac ex_least_nat_le) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4235 |
apply auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4236 |
done |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4237 |
then obtain n where "?P n" and smallest: "\<forall>k<n. \<not> ?P k" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4238 |
by blast |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4239 |
then obtain s u where obt: "finite s" "card s = n" "s\<subseteq>p" "\<forall>x\<in>s. 0 \<le> u x" |
| 64267 | 4240 |
"sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = y" by auto |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4241 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4242 |
have "card s \<le> aff_dim p + 1" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4243 |
proof (rule ccontr, simp only: not_le) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4244 |
assume "aff_dim p + 1 < card s" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4245 |
then have "affine_dependent s" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4246 |
using affine_dependent_biggerset[OF obt(1)] independent_card_le_aff_dim not_less obt(3) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4247 |
by blast |
| 64267 | 4248 |
then obtain w v where wv: "sum w s = 0" "v\<in>s" "w v \<noteq> 0" "(\<Sum>v\<in>s. w v *\<^sub>R v) = 0" |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4249 |
using affine_dependent_explicit_finite[OF obt(1)] by auto |
| 63040 | 4250 |
define i where "i = (\<lambda>v. (u v) / (- w v)) ` {v\<in>s. w v < 0}"
|
4251 |
define t where "t = Min i" |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4252 |
have "\<exists>x\<in>s. w x < 0" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4253 |
proof (rule ccontr, simp add: not_less) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4254 |
assume as:"\<forall>x\<in>s. 0 \<le> w x" |
| 64267 | 4255 |
then have "sum w (s - {v}) \<ge> 0"
|
4256 |
apply (rule_tac sum_nonneg) |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4257 |
apply auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4258 |
done |
| 64267 | 4259 |
then have "sum w s > 0" |
4260 |
unfolding sum.remove[OF obt(1) \<open>v\<in>s\<close>] |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4261 |
using as[THEN bspec[where x=v]] \<open>v\<in>s\<close> \<open>w v \<noteq> 0\<close> by auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4262 |
then show False using wv(1) by auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4263 |
qed |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4264 |
then have "i \<noteq> {}" unfolding i_def by auto
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4265 |
then have "t \<ge> 0" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4266 |
using Min_ge_iff[of i 0 ] and obt(1) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4267 |
unfolding t_def i_def |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4268 |
using obt(4)[unfolded le_less] |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4269 |
by (auto simp: divide_le_0_iff) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4270 |
have t: "\<forall>v\<in>s. u v + t * w v \<ge> 0" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4271 |
proof |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4272 |
fix v |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4273 |
assume "v \<in> s" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4274 |
then have v: "0 \<le> u v" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4275 |
using obt(4)[THEN bspec[where x=v]] by auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4276 |
show "0 \<le> u v + t * w v" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4277 |
proof (cases "w v < 0") |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4278 |
case False |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4279 |
thus ?thesis using v \<open>t\<ge>0\<close> by auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4280 |
next |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4281 |
case True |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4282 |
then have "t \<le> u v / (- w v)" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4283 |
using \<open>v\<in>s\<close> unfolding t_def i_def |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4284 |
apply (rule_tac Min_le) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4285 |
using obt(1) apply auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4286 |
done |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4287 |
then show ?thesis |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4288 |
unfolding real_0_le_add_iff |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4289 |
using pos_le_divide_eq[OF True[unfolded neg_0_less_iff_less[symmetric]]] |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4290 |
by auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4291 |
qed |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4292 |
qed |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4293 |
obtain a where "a \<in> s" and "t = (\<lambda>v. (u v) / (- w v)) a" and "w a < 0" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4294 |
using Min_in[OF _ \<open>i\<noteq>{}\<close>] and obt(1) unfolding i_def t_def by auto
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4295 |
then have a: "a \<in> s" "u a + t * w a = 0" by auto |
| 64267 | 4296 |
have *: "\<And>f. sum f (s - {a}) = sum f s - ((f a)::'b::ab_group_add)"
|
4297 |
unfolding sum.remove[OF obt(1) \<open>a\<in>s\<close>] by auto |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4298 |
have "(\<Sum>v\<in>s. u v + t * w v) = 1" |
| 64267 | 4299 |
unfolding sum.distrib wv(1) sum_distrib_left[symmetric] obt(5) by auto |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4300 |
moreover have "(\<Sum>v\<in>s. u v *\<^sub>R v + (t * w v) *\<^sub>R v) - (u a *\<^sub>R a + (t * w a) *\<^sub>R a) = y" |
| 64267 | 4301 |
unfolding sum.distrib obt(6) scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] wv(4) |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4302 |
using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by simp |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4303 |
ultimately have "?P (n - 1)" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4304 |
apply (rule_tac x="(s - {a})" in exI)
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4305 |
apply (rule_tac x="\<lambda>v. u v + t * w v" in exI) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4306 |
using obt(1-3) and t and a |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4307 |
apply (auto simp add: * scaleR_left_distrib) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4308 |
done |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4309 |
then show False |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4310 |
using smallest[THEN spec[where x="n - 1"]] by auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4311 |
qed |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4312 |
then show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and> |
| 64267 | 4313 |
(\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y" |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4314 |
using obt by auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4315 |
qed auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4316 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4317 |
lemma caratheodory_aff_dim: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4318 |
fixes p :: "('a::euclidean_space) set"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4319 |
shows "convex hull p = {x. \<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and> x \<in> convex hull s}"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4320 |
(is "?lhs = ?rhs") |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4321 |
proof |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4322 |
show "?lhs \<subseteq> ?rhs" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4323 |
apply (subst convex_hull_caratheodory_aff_dim) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4324 |
apply clarify |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4325 |
apply (rule_tac x="s" in exI) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4326 |
apply (simp add: hull_subset convex_explicit [THEN iffD1, OF convex_convex_hull]) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4327 |
done |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4328 |
next |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4329 |
show "?rhs \<subseteq> ?lhs" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4330 |
using hull_mono by blast |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4331 |
qed |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4332 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4333 |
lemma convex_hull_caratheodory: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4334 |
fixes p :: "('a::euclidean_space) set"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4335 |
shows "convex hull p = |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4336 |
{y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 1 \<and>
|
| 64267 | 4337 |
(\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) s = y}" |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4338 |
(is "?lhs = ?rhs") |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4339 |
proof (intro set_eqI iffI) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4340 |
fix x |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4341 |
assume "x \<in> ?lhs" then show "x \<in> ?rhs" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4342 |
apply (simp only: convex_hull_caratheodory_aff_dim Set.mem_Collect_eq) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4343 |
apply (erule ex_forward)+ |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4344 |
using aff_dim_le_DIM [of p] |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4345 |
apply simp |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4346 |
done |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4347 |
next |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4348 |
fix x |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4349 |
assume "x \<in> ?rhs" then show "x \<in> ?lhs" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4350 |
by (auto simp add: convex_hull_explicit) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4351 |
qed |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4352 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4353 |
theorem caratheodory: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4354 |
"convex hull p = |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4355 |
{x::'a::euclidean_space. \<exists>s. finite s \<and> s \<subseteq> p \<and>
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4356 |
card s \<le> DIM('a) + 1 \<and> x \<in> convex hull s}"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4357 |
proof safe |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4358 |
fix x |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4359 |
assume "x \<in> convex hull p" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4360 |
then obtain s u where "finite s" "s \<subseteq> p" "card s \<le> DIM('a) + 1"
|
| 64267 | 4361 |
"\<forall>x\<in>s. 0 \<le> u x" "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x" |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4362 |
unfolding convex_hull_caratheodory by auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4363 |
then show "\<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 1 \<and> x \<in> convex hull s"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4364 |
apply (rule_tac x=s in exI) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4365 |
using hull_subset[of s convex] |
| 63170 | 4366 |
using convex_convex_hull[simplified convex_explicit, of s, |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4367 |
THEN spec[where x=s], THEN spec[where x=u]] |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4368 |
apply auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4369 |
done |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4370 |
next |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4371 |
fix x s |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4372 |
assume "finite s" "s \<subseteq> p" "card s \<le> DIM('a) + 1" "x \<in> convex hull s"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4373 |
then show "x \<in> convex hull p" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4374 |
using hull_mono[OF \<open>s\<subseteq>p\<close>] by auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4375 |
qed |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4376 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
4377 |
|
| 60420 | 4378 |
subsection \<open>Relative interior of a set\<close> |
| 40377 | 4379 |
|
| 53347 | 4380 |
definition "rel_interior S = |
4381 |
{x. \<exists>T. openin (subtopology euclidean (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}"
|
|
4382 |
||
| 64287 | 4383 |
lemma rel_interior_mono: |
4384 |
"\<lbrakk>S \<subseteq> T; affine hull S = affine hull T\<rbrakk> |
|
4385 |
\<Longrightarrow> (rel_interior S) \<subseteq> (rel_interior T)" |
|
4386 |
by (auto simp: rel_interior_def) |
|
4387 |
||
4388 |
lemma rel_interior_maximal: |
|
4389 |
"\<lbrakk>T \<subseteq> S; openin(subtopology euclidean (affine hull S)) T\<rbrakk> \<Longrightarrow> T \<subseteq> (rel_interior S)" |
|
4390 |
by (auto simp: rel_interior_def) |
|
4391 |
||
| 53347 | 4392 |
lemma rel_interior: |
4393 |
"rel_interior S = {x \<in> S. \<exists>T. open T \<and> x \<in> T \<and> T \<inter> affine hull S \<subseteq> S}"
|
|
4394 |
unfolding rel_interior_def[of S] openin_open[of "affine hull S"] |
|
4395 |
apply auto |
|
4396 |
proof - |
|
4397 |
fix x T |
|
4398 |
assume *: "x \<in> S" "open T" "x \<in> T" "T \<inter> affine hull S \<subseteq> S" |
|
4399 |
then have **: "x \<in> T \<inter> affine hull S" |
|
4400 |
using hull_inc by auto |
|
| 54465 | 4401 |
show "\<exists>Tb. (\<exists>Ta. open Ta \<and> Tb = affine hull S \<inter> Ta) \<and> x \<in> Tb \<and> Tb \<subseteq> S" |
4402 |
apply (rule_tac x = "T \<inter> (affine hull S)" in exI) |
|
| 53347 | 4403 |
using * ** |
4404 |
apply auto |
|
4405 |
done |
|
4406 |
qed |
|
4407 |
||
4408 |
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)" |
|
4409 |
by (auto simp add: rel_interior) |
|
4410 |
||
4411 |
lemma mem_rel_interior_ball: |
|
4412 |
"x \<in> rel_interior S \<longleftrightarrow> x \<in> S \<and> (\<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S)" |
|
| 40377 | 4413 |
apply (simp add: rel_interior, safe) |
4414 |
apply (force simp add: open_contains_ball) |
|
| 53347 | 4415 |
apply (rule_tac x = "ball x e" in exI) |
|
44457
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
huffman
parents:
44365
diff
changeset
|
4416 |
apply simp |
| 40377 | 4417 |
done |
4418 |
||
| 49531 | 4419 |
lemma rel_interior_ball: |
| 53347 | 4420 |
"rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S}"
|
4421 |
using mem_rel_interior_ball [of _ S] by auto |
|
4422 |
||
4423 |
lemma mem_rel_interior_cball: |
|
4424 |
"x \<in> rel_interior S \<longleftrightarrow> x \<in> S \<and> (\<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S)" |
|
| 49531 | 4425 |
apply (simp add: rel_interior, safe) |
| 40377 | 4426 |
apply (force simp add: open_contains_cball) |
| 53347 | 4427 |
apply (rule_tac x = "ball x e" in exI) |
|
44457
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
huffman
parents:
44365
diff
changeset
|
4428 |
apply (simp add: subset_trans [OF ball_subset_cball]) |
| 40377 | 4429 |
apply auto |
4430 |
done |
|
4431 |
||
| 53347 | 4432 |
lemma rel_interior_cball: |
4433 |
"rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S}"
|
|
4434 |
using mem_rel_interior_cball [of _ S] by auto |
|
| 40377 | 4435 |
|
| 60303 | 4436 |
lemma rel_interior_empty [simp]: "rel_interior {} = {}"
|
| 49531 | 4437 |
by (auto simp add: rel_interior_def) |
| 40377 | 4438 |
|
| 60303 | 4439 |
lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}"
|
| 53347 | 4440 |
by (metis affine_hull_eq affine_sing) |
| 40377 | 4441 |
|
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4442 |
lemma rel_interior_sing [simp]: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4443 |
fixes a :: "'n::euclidean_space" shows "rel_interior {a} = {a}"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4444 |
apply (auto simp: rel_interior_ball) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4445 |
apply (rule_tac x=1 in exI) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4446 |
apply force |
| 53347 | 4447 |
done |
| 40377 | 4448 |
|
4449 |
lemma subset_rel_interior: |
|
| 53347 | 4450 |
fixes S T :: "'n::euclidean_space set" |
4451 |
assumes "S \<subseteq> T" |
|
4452 |
and "affine hull S = affine hull T" |
|
4453 |
shows "rel_interior S \<subseteq> rel_interior T" |
|
| 49531 | 4454 |
using assms by (auto simp add: rel_interior_def) |
4455 |
||
| 53347 | 4456 |
lemma rel_interior_subset: "rel_interior S \<subseteq> S" |
4457 |
by (auto simp add: rel_interior_def) |
|
4458 |
||
4459 |
lemma rel_interior_subset_closure: "rel_interior S \<subseteq> closure S" |
|
4460 |
using rel_interior_subset by (auto simp add: closure_def) |
|
4461 |
||
4462 |
lemma interior_subset_rel_interior: "interior S \<subseteq> rel_interior S" |
|
4463 |
by (auto simp add: rel_interior interior_def) |
|
| 40377 | 4464 |
|
4465 |
lemma interior_rel_interior: |
|
| 53347 | 4466 |
fixes S :: "'n::euclidean_space set" |
4467 |
assumes "aff_dim S = int(DIM('n))"
|
|
4468 |
shows "rel_interior S = interior S" |
|
| 40377 | 4469 |
proof - |
| 53347 | 4470 |
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
|
4471 |
using assms affine_hull_UNIV[of S] by auto |
| 53347 | 4472 |
then show ?thesis |
4473 |
unfolding rel_interior interior_def by auto |
|
| 40377 | 4474 |
qed |
4475 |
||
| 60303 | 4476 |
lemma rel_interior_interior: |
4477 |
fixes S :: "'n::euclidean_space set" |
|
4478 |
assumes "affine hull S = UNIV" |
|
4479 |
shows "rel_interior S = interior S" |
|
4480 |
using assms unfolding rel_interior interior_def by auto |
|
4481 |
||
| 40377 | 4482 |
lemma rel_interior_open: |
| 53347 | 4483 |
fixes S :: "'n::euclidean_space set" |
4484 |
assumes "open S" |
|
4485 |
shows "rel_interior S = S" |
|
4486 |
by (metis assms interior_eq interior_subset_rel_interior rel_interior_subset set_eq_subset) |
|
| 40377 | 4487 |
|
|
60800
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
4488 |
lemma interior_ball [simp]: "interior (ball x e) = ball x e" |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
4489 |
by (simp add: interior_open) |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
4490 |
|
| 40377 | 4491 |
lemma interior_rel_interior_gen: |
| 53347 | 4492 |
fixes S :: "'n::euclidean_space set" |
4493 |
shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})"
|
|
4494 |
by (metis interior_rel_interior low_dim_interior) |
|
| 40377 | 4495 |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4496 |
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
|
4497 |
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
|
4498 |
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
|
4499 |
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
|
4500 |
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4501 |
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
|
4502 |
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
|
4503 |
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
|
4504 |
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
|
4505 |
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4506 |
lemma rel_interior_affine_hull [simp]: |
| 53347 | 4507 |
fixes S :: "'n::euclidean_space set" |
4508 |
shows "rel_interior (affine hull S) = affine hull S" |
|
4509 |
proof - |
|
4510 |
have *: "rel_interior (affine hull S) \<subseteq> affine hull S" |
|
4511 |
using rel_interior_subset by auto |
|
4512 |
{
|
|
4513 |
fix x |
|
4514 |
assume x: "x \<in> affine hull S" |
|
| 63040 | 4515 |
define e :: real where "e = 1" |
| 53347 | 4516 |
then have "e > 0" "ball x e \<inter> affine hull (affine hull S) \<subseteq> affine hull S" |
4517 |
using hull_hull[of _ S] by auto |
|
4518 |
then have "x \<in> rel_interior (affine hull S)" |
|
4519 |
using x rel_interior_ball[of "affine hull S"] by auto |
|
4520 |
} |
|
4521 |
then show ?thesis using * by auto |
|
| 40377 | 4522 |
qed |
4523 |
||
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
4524 |
lemma rel_interior_UNIV [simp]: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV"
|
| 53347 | 4525 |
by (metis open_UNIV rel_interior_open) |
| 40377 | 4526 |
|
4527 |
lemma rel_interior_convex_shrink: |
|
| 53347 | 4528 |
fixes S :: "'a::euclidean_space set" |
4529 |
assumes "convex S" |
|
4530 |
and "c \<in> rel_interior S" |
|
4531 |
and "x \<in> S" |
|
4532 |
and "0 < e" |
|
4533 |
and "e \<le> 1" |
|
4534 |
shows "x - e *\<^sub>R (x - c) \<in> rel_interior S" |
|
4535 |
proof - |
|
| 54465 | 4536 |
obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S" |
| 53347 | 4537 |
using assms(2) unfolding mem_rel_interior_ball by auto |
4538 |
{
|
|
4539 |
fix y |
|
4540 |
assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \<in> affine hull S" |
|
4541 |
have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" |
|
| 60420 | 4542 |
using \<open>e > 0\<close> by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) |
| 53347 | 4543 |
have "x \<in> affine hull S" |
4544 |
using assms hull_subset[of S] by auto |
|
| 49531 | 4545 |
moreover have "1 / e + - ((1 - e) / e) = 1" |
| 60420 | 4546 |
using \<open>e > 0\<close> left_diff_distrib[of "1" "(1-e)" "1/e"] by auto |
| 53347 | 4547 |
ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \<in> affine hull S" |
4548 |
using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"] |
|
4549 |
by (simp add: algebra_simps) |
|
| 61945 | 4550 |
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)" |
| 53347 | 4551 |
unfolding dist_norm norm_scaleR[symmetric] |
4552 |
apply (rule arg_cong[where f=norm]) |
|
| 60420 | 4553 |
using \<open>e > 0\<close> |
| 53347 | 4554 |
apply (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) |
4555 |
done |
|
| 61945 | 4556 |
also have "\<dots> = \<bar>1/e\<bar> * norm (x - e *\<^sub>R (x - c) - y)" |
| 53347 | 4557 |
by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) |
4558 |
also have "\<dots> < d" |
|
| 60420 | 4559 |
using as[unfolded dist_norm] and \<open>e > 0\<close> |
4560 |
by (auto simp add:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute) |
|
| 53347 | 4561 |
finally have "y \<in> S" |
4562 |
apply (subst *) |
|
4563 |
apply (rule assms(1)[unfolded convex_alt,rule_format]) |
|
4564 |
apply (rule d[unfolded subset_eq,rule_format]) |
|
4565 |
unfolding mem_ball |
|
4566 |
using assms(3-5) ** |
|
4567 |
apply auto |
|
4568 |
done |
|
4569 |
} |
|
4570 |
then have "ball (x - e *\<^sub>R (x - c)) (e*d) \<inter> affine hull S \<subseteq> S" |
|
4571 |
by auto |
|
4572 |
moreover have "e * d > 0" |
|
| 60420 | 4573 |
using \<open>e > 0\<close> \<open>d > 0\<close> by simp |
| 53347 | 4574 |
moreover have c: "c \<in> S" |
4575 |
using assms rel_interior_subset by auto |
|
4576 |
moreover from c have "x - e *\<^sub>R (x - c) \<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
|
4577 |
using convexD_alt[of S x c e] |
| 53347 | 4578 |
apply (simp add: algebra_simps) |
4579 |
using assms |
|
4580 |
apply auto |
|
4581 |
done |
|
4582 |
ultimately show ?thesis |
|
| 60420 | 4583 |
using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] \<open>e > 0\<close> by auto |
| 40377 | 4584 |
qed |
4585 |
||
4586 |
lemma interior_real_semiline: |
|
| 53347 | 4587 |
fixes a :: real |
4588 |
shows "interior {a..} = {a<..}"
|
|
4589 |
proof - |
|
4590 |
{
|
|
4591 |
fix y |
|
4592 |
assume "a < y" |
|
4593 |
then have "y \<in> interior {a..}"
|
|
4594 |
apply (simp add: mem_interior) |
|
4595 |
apply (rule_tac x="(y-a)" in exI) |
|
4596 |
apply (auto simp add: dist_norm) |
|
4597 |
done |
|
4598 |
} |
|
4599 |
moreover |
|
4600 |
{
|
|
4601 |
fix y |
|
4602 |
assume "y \<in> interior {a..}"
|
|
4603 |
then obtain e where e: "e > 0" "cball y e \<subseteq> {a..}"
|
|
4604 |
using mem_interior_cball[of y "{a..}"] by auto
|
|
4605 |
moreover from e have "y - e \<in> cball y e" |
|
4606 |
by (auto simp add: 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
|
4607 |
ultimately have "a \<le> y - e" by blast |
| 53347 | 4608 |
then have "a < y" using e by auto |
4609 |
} |
|
4610 |
ultimately show ?thesis by auto |
|
| 40377 | 4611 |
qed |
4612 |
||
|
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4613 |
lemma continuous_ge_on_Ioo: |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4614 |
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
|
4615 |
shows "g (x::real) \<ge> (a::real)" |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4616 |
proof- |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4617 |
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
|
4618 |
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
|
4619 |
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
|
4620 |
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
|
4621 |
by (auto simp: continuous_on_closed_vimage) |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4622 |
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
|
4623 |
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
|
4624 |
qed |
|
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4625 |
|
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4626 |
lemma interior_real_semiline': |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4627 |
fixes a :: real |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4628 |
shows "interior {..a} = {..<a}"
|
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4629 |
proof - |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4630 |
{
|
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4631 |
fix y |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4632 |
assume "a > y" |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4633 |
then have "y \<in> interior {..a}"
|
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4634 |
apply (simp add: mem_interior) |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4635 |
apply (rule_tac x="(a-y)" in exI) |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4636 |
apply (auto simp add: dist_norm) |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4637 |
done |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4638 |
} |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4639 |
moreover |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4640 |
{
|
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4641 |
fix y |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4642 |
assume "y \<in> interior {..a}"
|
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4643 |
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
|
4644 |
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
|
4645 |
moreover from e have "y + e \<in> cball y e" |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4646 |
by (auto simp add: cball_def dist_norm) |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4647 |
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
|
4648 |
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
|
4649 |
} |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4650 |
ultimately show ?thesis by auto |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4651 |
qed |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4652 |
|
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
4653 |
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
|
4654 |
proof- |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4655 |
have "{a..b} = {a..} \<inter> {..b}" by auto
|
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
61952
diff
changeset
|
4656 |
also have "interior ... = {a<..} \<inter> {..<b}"
|
|
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4657 |
by (simp add: interior_real_semiline interior_real_semiline') |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4658 |
also have "... = {a<..<b}" by auto
|
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4659 |
finally show ?thesis . |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4660 |
qed |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4661 |
|
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
4662 |
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
|
4663 |
by (metis interior_atLeastAtMost_real interior_interior) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
4664 |
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
4665 |
lemma frontier_real_Iic [simp]: |
|
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4666 |
fixes a :: real |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4667 |
shows "frontier {..a} = {a}"
|
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4668 |
unfolding frontier_def by (auto simp add: interior_real_semiline') |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61848
diff
changeset
|
4669 |
|
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
4670 |
lemma rel_interior_real_box [simp]: |
| 53347 | 4671 |
fixes a b :: real |
4672 |
assumes "a < b" |
|
| 56188 | 4673 |
shows "rel_interior {a .. b} = {a <..< b}"
|
| 53347 | 4674 |
proof - |
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54465
diff
changeset
|
4675 |
have "box a b \<noteq> {}"
|
| 53347 | 4676 |
using assms |
4677 |
unfolding set_eq_iff |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
4678 |
by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def) |
| 40377 | 4679 |
then show ?thesis |
| 56188 | 4680 |
using interior_rel_interior_gen[of "cbox a b", symmetric] |
| 62390 | 4681 |
by (simp split: if_split_asm del: box_real add: box_real[symmetric] interior_cbox) |
| 40377 | 4682 |
qed |
4683 |
||
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
4684 |
lemma rel_interior_real_semiline [simp]: |
| 53347 | 4685 |
fixes a :: real |
4686 |
shows "rel_interior {a..} = {a<..}"
|
|
4687 |
proof - |
|
4688 |
have *: "{a<..} \<noteq> {}"
|
|
4689 |
unfolding set_eq_iff by (auto intro!: exI[of _ "a + 1"]) |
|
4690 |
then show ?thesis using interior_real_semiline interior_rel_interior_gen[of "{a..}"]
|
|
| 62390 | 4691 |
by (auto split: if_split_asm) |
| 40377 | 4692 |
qed |
4693 |
||
| 60420 | 4694 |
subsubsection \<open>Relative open sets\<close> |
| 40377 | 4695 |
|
| 53347 | 4696 |
definition "rel_open S \<longleftrightarrow> rel_interior S = S" |
4697 |
||
4698 |
lemma rel_open: "rel_open S \<longleftrightarrow> openin (subtopology euclidean (affine hull S)) S" |
|
4699 |
unfolding rel_open_def rel_interior_def |
|
4700 |
apply auto |
|
4701 |
using openin_subopen[of "subtopology euclidean (affine hull S)" S] |
|
4702 |
apply auto |
|
4703 |
done |
|
4704 |
||
| 63072 | 4705 |
lemma openin_rel_interior: "openin (subtopology euclidean (affine hull S)) (rel_interior S)" |
| 40377 | 4706 |
apply (simp add: rel_interior_def) |
| 53347 | 4707 |
apply (subst openin_subopen) |
4708 |
apply blast |
|
4709 |
done |
|
| 40377 | 4710 |
|
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4711 |
lemma openin_set_rel_interior: |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4712 |
"openin (subtopology euclidean S) (rel_interior S)" |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4713 |
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
|
4714 |
|
| 49531 | 4715 |
lemma affine_rel_open: |
| 53347 | 4716 |
fixes S :: "'n::euclidean_space set" |
4717 |
assumes "affine S" |
|
4718 |
shows "rel_open S" |
|
4719 |
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
|
4720 |
using assms rel_interior_affine_hull[of S] affine_hull_eq[of S] |
| 53347 | 4721 |
by metis |
| 40377 | 4722 |
|
| 49531 | 4723 |
lemma affine_closed: |
| 53347 | 4724 |
fixes S :: "'n::euclidean_space set" |
4725 |
assumes "affine S" |
|
4726 |
shows "closed S" |
|
4727 |
proof - |
|
4728 |
{
|
|
4729 |
assume "S \<noteq> {}"
|
|
4730 |
then obtain L where L: "subspace L" "affine_parallel S L" |
|
4731 |
using assms affine_parallel_subspace[of S] by auto |
|
4732 |
then obtain a where a: "S = (op + a ` L)" |
|
4733 |
using affine_parallel_def[of L S] affine_parallel_commut by auto |
|
4734 |
from L have "closed L" using closed_subspace by auto |
|
4735 |
then have "closed S" |
|
4736 |
using closed_translation a by auto |
|
4737 |
} |
|
4738 |
then show ?thesis by auto |
|
| 40377 | 4739 |
qed |
4740 |
||
4741 |
lemma closure_affine_hull: |
|
| 53347 | 4742 |
fixes S :: "'n::euclidean_space set" |
4743 |
shows "closure S \<subseteq> affine hull S" |
|
| 44524 | 4744 |
by (intro closure_minimal hull_subset affine_closed affine_affine_hull) |
| 40377 | 4745 |
|
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
4746 |
lemma closure_same_affine_hull [simp]: |
| 53347 | 4747 |
fixes S :: "'n::euclidean_space set" |
| 40377 | 4748 |
shows "affine hull (closure S) = affine hull S" |
| 53347 | 4749 |
proof - |
4750 |
have "affine hull (closure S) \<subseteq> affine hull S" |
|
4751 |
using hull_mono[of "closure S" "affine hull S" "affine"] |
|
4752 |
closure_affine_hull[of S] hull_hull[of "affine" S] |
|
4753 |
by auto |
|
4754 |
moreover have "affine hull (closure S) \<supseteq> affine hull S" |
|
4755 |
using hull_mono[of "S" "closure S" "affine"] closure_subset by auto |
|
4756 |
ultimately show ?thesis by auto |
|
| 49531 | 4757 |
qed |
4758 |
||
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
4759 |
lemma closure_aff_dim [simp]: |
| 53347 | 4760 |
fixes S :: "'n::euclidean_space set" |
| 40377 | 4761 |
shows "aff_dim (closure S) = aff_dim S" |
| 53347 | 4762 |
proof - |
4763 |
have "aff_dim S \<le> aff_dim (closure S)" |
|
4764 |
using aff_dim_subset closure_subset by auto |
|
4765 |
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
|
4766 |
using aff_dim_subset closure_affine_hull by blast |
| 53347 | 4767 |
moreover have "aff_dim (affine hull S) = aff_dim S" |
4768 |
using aff_dim_affine_hull by auto |
|
4769 |
ultimately show ?thesis by auto |
|
| 40377 | 4770 |
qed |
4771 |
||
4772 |
lemma rel_interior_closure_convex_shrink: |
|
| 53347 | 4773 |
fixes S :: "_::euclidean_space set" |
4774 |
assumes "convex S" |
|
4775 |
and "c \<in> rel_interior S" |
|
4776 |
and "x \<in> closure S" |
|
4777 |
and "e > 0" |
|
4778 |
and "e \<le> 1" |
|
4779 |
shows "x - e *\<^sub>R (x - c) \<in> rel_interior S" |
|
4780 |
proof - |
|
4781 |
obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S" |
|
4782 |
using assms(2) unfolding mem_rel_interior_ball by auto |
|
4783 |
have "\<exists>y \<in> S. norm (y - x) * (1 - e) < e * d" |
|
4784 |
proof (cases "x \<in> S") |
|
4785 |
case True |
|
| 60420 | 4786 |
then show ?thesis using \<open>e > 0\<close> \<open>d > 0\<close> |
| 53347 | 4787 |
apply (rule_tac bexI[where x=x]) |
| 56544 | 4788 |
apply (auto) |
| 53347 | 4789 |
done |
4790 |
next |
|
4791 |
case False |
|
4792 |
then have x: "x islimpt S" |
|
4793 |
using assms(3)[unfolded closure_def] by auto |
|
4794 |
show ?thesis |
|
4795 |
proof (cases "e = 1") |
|
4796 |
case True |
|
4797 |
obtain y where "y \<in> S" "y \<noteq> x" "dist y x < 1" |
|
| 40377 | 4798 |
using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto |
| 53347 | 4799 |
then show ?thesis |
4800 |
apply (rule_tac x=y in bexI) |
|
4801 |
unfolding True |
|
| 60420 | 4802 |
using \<open>d > 0\<close> |
| 53347 | 4803 |
apply auto |
4804 |
done |
|
4805 |
next |
|
4806 |
case False |
|
4807 |
then have "0 < e * d / (1 - e)" and *: "1 - e > 0" |
|
| 60420 | 4808 |
using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by (auto) |
| 53347 | 4809 |
then obtain y where "y \<in> S" "y \<noteq> x" "dist y x < e * d / (1 - e)" |
| 40377 | 4810 |
using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto |
| 53347 | 4811 |
then show ?thesis |
4812 |
apply (rule_tac x=y in bexI) |
|
4813 |
unfolding dist_norm |
|
4814 |
using pos_less_divide_eq[OF *] |
|
4815 |
apply auto |
|
4816 |
done |
|
4817 |
qed |
|
4818 |
qed |
|
4819 |
then obtain y where "y \<in> S" and y: "norm (y - x) * (1 - e) < e * d" |
|
4820 |
by auto |
|
| 63040 | 4821 |
define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" |
| 53347 | 4822 |
have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" |
| 60420 | 4823 |
unfolding z_def using \<open>e > 0\<close> |
| 53347 | 4824 |
by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) |
4825 |
have zball: "z \<in> ball c d" |
|
4826 |
using mem_ball z_def dist_norm[of c] |
|
4827 |
using y and assms(4,5) |
|
4828 |
by (auto simp add:field_simps norm_minus_commute) |
|
4829 |
have "x \<in> affine hull S" |
|
4830 |
using closure_affine_hull assms by auto |
|
4831 |
moreover have "y \<in> affine hull S" |
|
| 60420 | 4832 |
using \<open>y \<in> S\<close> hull_subset[of S] by auto |
| 53347 | 4833 |
moreover have "c \<in> affine hull S" |
4834 |
using assms rel_interior_subset hull_subset[of S] by auto |
|
4835 |
ultimately have "z \<in> affine hull S" |
|
| 49531 | 4836 |
using z_def affine_affine_hull[of S] |
| 53347 | 4837 |
mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"] |
4838 |
assms |
|
4839 |
by (auto simp add: field_simps) |
|
4840 |
then have "z \<in> S" using d zball by auto |
|
4841 |
obtain d1 where "d1 > 0" and d1: "ball z d1 \<le> ball c d" |
|
| 40377 | 4842 |
using zball open_ball[of c d] openE[of "ball c d" z] by auto |
| 53347 | 4843 |
then have "ball z d1 \<inter> affine hull S \<subseteq> ball c d \<inter> affine hull S" |
4844 |
by auto |
|
4845 |
then have "ball z d1 \<inter> affine hull S \<subseteq> S" |
|
4846 |
using d by auto |
|
4847 |
then have "z \<in> rel_interior S" |
|
| 60420 | 4848 |
using mem_rel_interior_ball using \<open>d1 > 0\<close> \<open>z \<in> S\<close> by auto |
| 53347 | 4849 |
then have "y - e *\<^sub>R (y - z) \<in> rel_interior S" |
| 60420 | 4850 |
using rel_interior_convex_shrink[of S z y e] assms \<open>y \<in> S\<close> by auto |
| 53347 | 4851 |
then show ?thesis using * by auto |
4852 |
qed |
|
4853 |
||
|
62620
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
4854 |
lemma rel_interior_eq: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
4855 |
"rel_interior s = s \<longleftrightarrow> openin(subtopology euclidean (affine hull s)) s" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
4856 |
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
|
4857 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
4858 |
lemma rel_interior_openin: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
4859 |
"openin(subtopology euclidean (affine hull s)) s \<Longrightarrow> rel_interior s = s" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
4860 |
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
|
4861 |
|
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4862 |
lemma rel_interior_affine: |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4863 |
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
|
4864 |
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
|
4865 |
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
|
4866 |
|
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4867 |
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
|
4868 |
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
|
4869 |
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
|
4870 |
proof (cases "S = {}")
|
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4871 |
case True |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4872 |
then show ?thesis |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4873 |
by auto |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4874 |
next |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4875 |
case False show ?thesis |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4876 |
proof |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4877 |
assume eq: "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
|
4878 |
have "S = {} \<or> S = affine hull S"
|
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4879 |
apply (rule connected_clopen [THEN iffD1, rule_format]) |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4880 |
apply (simp add: affine_imp_convex convex_connected) |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4881 |
apply (rule conjI) |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4882 |
apply (metis eq closure_subset openin_rel_interior rel_interior_subset subset_antisym) |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4883 |
apply (metis closed_subset closure_subset_eq eq hull_subset rel_interior_subset) |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4884 |
done |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4885 |
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
|
4886 |
by auto |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4887 |
then show "affine S" |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4888 |
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
|
4889 |
next |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4890 |
assume "affine S" |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4891 |
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
|
4892 |
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
|
4893 |
qed |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4894 |
qed |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
4895 |
|
| 40377 | 4896 |
|
| 60420 | 4897 |
subsubsection\<open>Relative interior preserves under linear transformations\<close> |
| 40377 | 4898 |
|
4899 |
lemma rel_interior_translation_aux: |
|
| 53347 | 4900 |
fixes a :: "'n::euclidean_space" |
4901 |
shows "((\<lambda>x. a + x) ` rel_interior S) \<subseteq> rel_interior ((\<lambda>x. a + x) ` S)" |
|
4902 |
proof - |
|
4903 |
{
|
|
4904 |
fix x |
|
4905 |
assume x: "x \<in> rel_interior S" |
|
4906 |
then obtain T where "open T" "x \<in> T \<inter> S" "T \<inter> affine hull S \<subseteq> S" |
|
4907 |
using mem_rel_interior[of x S] by auto |
|
4908 |
then have "open ((\<lambda>x. a + x) ` T)" |
|
4909 |
and "a + x \<in> ((\<lambda>x. a + x) ` T) \<inter> ((\<lambda>x. a + x) ` S)" |
|
4910 |
and "((\<lambda>x. a + x) ` T) \<inter> affine hull ((\<lambda>x. a + x) ` S) \<subseteq> (\<lambda>x. a + x) ` S" |
|
4911 |
using affine_hull_translation[of a S] open_translation[of T a] x by auto |
|
4912 |
then have "a + x \<in> rel_interior ((\<lambda>x. a + x) ` S)" |
|
4913 |
using mem_rel_interior[of "a+x" "((\<lambda>x. a + x) ` S)"] by auto |
|
4914 |
} |
|
4915 |
then show ?thesis by auto |
|
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
4916 |
qed |
| 40377 | 4917 |
|
4918 |
lemma rel_interior_translation: |
|
| 53347 | 4919 |
fixes a :: "'n::euclidean_space" |
4920 |
shows "rel_interior ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` rel_interior S" |
|
4921 |
proof - |
|
4922 |
have "(\<lambda>x. (-a) + x) ` rel_interior ((\<lambda>x. a + x) ` S) \<subseteq> rel_interior S" |
|
4923 |
using rel_interior_translation_aux[of "-a" "(\<lambda>x. a + x) ` S"] |
|
4924 |
translation_assoc[of "-a" "a"] |
|
4925 |
by auto |
|
4926 |
then have "((\<lambda>x. a + x) ` rel_interior S) \<supseteq> rel_interior ((\<lambda>x. a + x) ` S)" |
|
4927 |
using translation_inverse_subset[of a "rel_interior (op + a ` S)" "rel_interior S"] |
|
4928 |
by auto |
|
4929 |
then show ?thesis |
|
4930 |
using rel_interior_translation_aux[of a S] by auto |
|
| 40377 | 4931 |
qed |
4932 |
||
4933 |
||
4934 |
lemma affine_hull_linear_image: |
|
| 53347 | 4935 |
assumes "bounded_linear f" |
4936 |
shows "f ` (affine hull s) = affine hull f ` s" |
|
4937 |
apply rule |
|
4938 |
unfolding subset_eq ball_simps |
|
4939 |
apply (rule_tac[!] hull_induct, rule hull_inc) |
|
4940 |
prefer 3 |
|
4941 |
apply (erule imageE) |
|
4942 |
apply (rule_tac x=xa in image_eqI) |
|
4943 |
apply assumption |
|
4944 |
apply (rule hull_subset[unfolded subset_eq, rule_format]) |
|
4945 |
apply assumption |
|
4946 |
proof - |
|
| 40377 | 4947 |
interpret f: bounded_linear f by fact |
| 53347 | 4948 |
show "affine {x. f x \<in> affine hull f ` s}"
|
4949 |
unfolding affine_def |
|
4950 |
by (auto simp add: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format]) |
|
4951 |
show "affine {x. x \<in> f ` (affine hull s)}"
|
|
4952 |
using affine_affine_hull[unfolded affine_def, of s] |
|
| 40377 | 4953 |
unfolding affine_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric]) |
4954 |
qed auto |
|
4955 |
||
4956 |
||
4957 |
lemma rel_interior_injective_on_span_linear_image: |
|
| 53347 | 4958 |
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" |
4959 |
and S :: "'m::euclidean_space set" |
|
4960 |
assumes "bounded_linear f" |
|
4961 |
and "inj_on f (span S)" |
|
4962 |
shows "rel_interior (f ` S) = f ` (rel_interior S)" |
|
4963 |
proof - |
|
4964 |
{
|
|
4965 |
fix z |
|
4966 |
assume z: "z \<in> rel_interior (f ` S)" |
|
4967 |
then have "z \<in> f ` S" |
|
4968 |
using rel_interior_subset[of "f ` S"] by auto |
|
4969 |
then obtain x where x: "x \<in> S" "f x = z" by auto |
|
4970 |
obtain e2 where e2: "e2 > 0" "cball z e2 \<inter> affine hull (f ` S) \<subseteq> (f ` S)" |
|
4971 |
using z rel_interior_cball[of "f ` S"] by auto |
|
4972 |
obtain K where K: "K > 0" "\<And>x. norm (f x) \<le> norm x * K" |
|
4973 |
using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto |
|
| 63040 | 4974 |
define e1 where "e1 = 1 / K" |
| 53347 | 4975 |
then have e1: "e1 > 0" "\<And>x. e1 * norm (f x) \<le> norm x" |
4976 |
using K pos_le_divide_eq[of e1] by auto |
|
| 63040 | 4977 |
define e where "e = e1 * e2" |
| 56544 | 4978 |
then have "e > 0" using e1 e2 by auto |
| 53347 | 4979 |
{
|
4980 |
fix y |
|
4981 |
assume y: "y \<in> cball x e \<inter> affine hull S" |
|
4982 |
then have h1: "f y \<in> affine hull (f ` S)" |
|
4983 |
using affine_hull_linear_image[of f S] assms by auto |
|
4984 |
from y have "norm (x-y) \<le> e1 * e2" |
|
4985 |
using cball_def[of x e] dist_norm[of x y] e_def by auto |
|
4986 |
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
|
4987 |
using assms linear_diff[of f x y] linear_conv_bounded_linear[of f] by auto |
| 53347 | 4988 |
moreover have "e1 * norm (f (x-y)) \<le> norm (x - y)" |
4989 |
using e1 by auto |
|
4990 |
ultimately have "e1 * norm ((f x)-(f y)) \<le> e1 * e2" |
|
4991 |
by auto |
|
4992 |
then have "f y \<in> cball z e2" |
|
4993 |
using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1 x by auto |
|
4994 |
then have "f y \<in> f ` S" |
|
4995 |
using y e2 h1 by auto |
|
4996 |
then have "y \<in> S" |
|
4997 |
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
|
4998 |
inj_on_image_mem_iff [OF \<open>inj_on f (span S)\<close>] |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
4999 |
by (metis Int_iff span_inc subsetCE) |
| 53347 | 5000 |
} |
5001 |
then have "z \<in> f ` (rel_interior S)" |
|
| 60420 | 5002 |
using mem_rel_interior_cball[of x S] \<open>e > 0\<close> x by auto |
| 49531 | 5003 |
} |
| 53347 | 5004 |
moreover |
5005 |
{
|
|
5006 |
fix x |
|
5007 |
assume x: "x \<in> rel_interior S" |
|
| 54465 | 5008 |
then obtain e2 where e2: "e2 > 0" "cball x e2 \<inter> affine hull S \<subseteq> S" |
| 53347 | 5009 |
using rel_interior_cball[of S] by auto |
5010 |
have "x \<in> S" using x rel_interior_subset by auto |
|
5011 |
then have *: "f x \<in> f ` S" by auto |
|
5012 |
have "\<forall>x\<in>span S. f x = 0 \<longrightarrow> x = 0" |
|
5013 |
using assms subspace_span linear_conv_bounded_linear[of f] |
|
5014 |
linear_injective_on_subspace_0[of f "span S"] |
|
5015 |
by auto |
|
5016 |
then obtain e1 where e1: "e1 > 0" "\<forall>x \<in> span S. e1 * norm x \<le> norm (f x)" |
|
5017 |
using assms injective_imp_isometric[of "span S" f] |
|
5018 |
subspace_span[of S] closed_subspace[of "span S"] |
|
5019 |
by auto |
|
| 63040 | 5020 |
define e where "e = e1 * e2" |
| 56544 | 5021 |
hence "e > 0" using e1 e2 by auto |
| 53347 | 5022 |
{
|
5023 |
fix y |
|
5024 |
assume y: "y \<in> cball (f x) e \<inter> affine hull (f ` S)" |
|
5025 |
then have "y \<in> f ` (affine hull S)" |
|
5026 |
using affine_hull_linear_image[of f S] assms by auto |
|
5027 |
then obtain xy where xy: "xy \<in> affine hull S" "f xy = y" by auto |
|
5028 |
with y have "norm (f x - f xy) \<le> e1 * e2" |
|
5029 |
using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto |
|
5030 |
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
|
5031 |
using assms linear_diff[of f x xy] linear_conv_bounded_linear[of f] by auto |
| 53347 | 5032 |
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
|
5033 |
using subspace_diff[of "span S" x xy] subspace_span \<open>x \<in> S\<close> xy |
| 53347 | 5034 |
affine_hull_subset_span[of S] span_inc |
5035 |
by auto |
|
5036 |
moreover from * have "e1 * norm (x - xy) \<le> norm (f (x - xy))" |
|
5037 |
using e1 by auto |
|
5038 |
ultimately have "e1 * norm (x - xy) \<le> e1 * e2" |
|
5039 |
by auto |
|
5040 |
then have "xy \<in> cball x e2" |
|
5041 |
using cball_def[of x e2] dist_norm[of x xy] e1 by auto |
|
5042 |
then have "y \<in> f ` S" |
|
5043 |
using xy e2 by auto |
|
5044 |
} |
|
5045 |
then have "f x \<in> rel_interior (f ` S)" |
|
| 60420 | 5046 |
using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * \<open>e > 0\<close> by auto |
| 49531 | 5047 |
} |
| 53347 | 5048 |
ultimately show ?thesis by auto |
| 40377 | 5049 |
qed |
5050 |
||
5051 |
lemma rel_interior_injective_linear_image: |
|
| 53347 | 5052 |
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" |
5053 |
assumes "bounded_linear f" |
|
5054 |
and "inj f" |
|
5055 |
shows "rel_interior (f ` S) = f ` (rel_interior S)" |
|
5056 |
using assms rel_interior_injective_on_span_linear_image[of f S] |
|
5057 |
subset_inj_on[of f "UNIV" "span S"] |
|
5058 |
by auto |
|
5059 |
||
| 40377 | 5060 |
|
| 60420 | 5061 |
subsection\<open>Some Properties of subset of standard basis\<close> |
| 40377 | 5062 |
|
| 53347 | 5063 |
lemma affine_hull_substd_basis: |
5064 |
assumes "d \<subseteq> Basis" |
|
5065 |
shows "affine hull (insert 0 d) = {x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
|
|
5066 |
(is "affine hull (insert 0 ?A) = ?B") |
|
5067 |
proof - |
|
| 61076 | 5068 |
have *: "\<And>A. op + (0::'a) ` A = A" "\<And>A. op + (- (0::'a)) ` A = A" |
| 53347 | 5069 |
by auto |
5070 |
show ?thesis |
|
5071 |
unfolding affine_hull_insert_span_gen span_substd_basis[OF assms,symmetric] * .. |
|
| 40377 | 5072 |
qed |
5073 |
||
| 60303 | 5074 |
lemma affine_hull_convex_hull [simp]: "affine hull (convex hull S) = affine hull S" |
| 53347 | 5075 |
by (metis Int_absorb1 Int_absorb2 convex_hull_subset_affine_hull hull_hull hull_mono hull_subset) |
5076 |
||
| 40377 | 5077 |
|
| 60420 | 5078 |
subsection \<open>Openness and compactness are preserved by convex hull operation.\<close> |
| 33175 | 5079 |
|
| 34964 | 5080 |
lemma open_convex_hull[intro]: |
| 33175 | 5081 |
fixes s :: "'a::real_normed_vector set" |
5082 |
assumes "open s" |
|
| 53347 | 5083 |
shows "open (convex hull s)" |
5084 |
unfolding open_contains_cball convex_hull_explicit |
|
5085 |
unfolding mem_Collect_eq ball_simps(8) |
|
5086 |
proof (rule, rule) |
|
5087 |
fix a |
|
| 64267 | 5088 |
assume "\<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = a" |
5089 |
then obtain t u where obt: "finite t" "t\<subseteq>s" "\<forall>x\<in>t. 0 \<le> u x" "sum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = a" |
|
| 53347 | 5090 |
by auto |
5091 |
||
5092 |
from assms[unfolded open_contains_cball] obtain b |
|
5093 |
where b: "\<forall>x\<in>s. 0 < b x \<and> cball x (b x) \<subseteq> s" |
|
5094 |
using bchoice[of s "\<lambda>x e. e > 0 \<and> cball x e \<subseteq> s"] by auto |
|
5095 |
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
|
5096 |
using obt by auto |
| 63040 | 5097 |
define i where "i = b ` t" |
| 53347 | 5098 |
|
5099 |
show "\<exists>e > 0. |
|
| 64267 | 5100 |
cball a e \<subseteq> {y. \<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y}"
|
| 53347 | 5101 |
apply (rule_tac x = "Min i" in exI) |
5102 |
unfolding subset_eq |
|
5103 |
apply rule |
|
5104 |
defer |
|
5105 |
apply rule |
|
5106 |
unfolding mem_Collect_eq |
|
5107 |
proof - |
|
5108 |
show "0 < Min i" |
|
| 60420 | 5109 |
unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \<open>b ` t\<noteq>{}\<close>]
|
| 53347 | 5110 |
using b |
5111 |
apply simp |
|
5112 |
apply rule |
|
5113 |
apply (erule_tac x=x in ballE) |
|
| 60420 | 5114 |
using \<open>t\<subseteq>s\<close> |
| 53347 | 5115 |
apply auto |
5116 |
done |
|
5117 |
next |
|
5118 |
fix y |
|
5119 |
assume "y \<in> cball a (Min i)" |
|
5120 |
then have y: "norm (a - y) \<le> Min i" |
|
5121 |
unfolding dist_norm[symmetric] by auto |
|
5122 |
{
|
|
5123 |
fix x |
|
5124 |
assume "x \<in> t" |
|
5125 |
then have "Min i \<le> b x" |
|
5126 |
unfolding i_def |
|
5127 |
apply (rule_tac Min_le) |
|
5128 |
using obt(1) |
|
5129 |
apply auto |
|
5130 |
done |
|
5131 |
then have "x + (y - a) \<in> cball x (b x)" |
|
5132 |
using y unfolding mem_cball dist_norm by auto |
|
| 60420 | 5133 |
moreover from \<open>x\<in>t\<close> have "x \<in> s" |
| 53347 | 5134 |
using obt(2) by auto |
5135 |
ultimately have "x + (y - a) \<in> s" |
|
| 54465 | 5136 |
using y and b[THEN bspec[where x=x]] unfolding subset_eq by fast |
| 53347 | 5137 |
} |
| 33175 | 5138 |
moreover |
| 53347 | 5139 |
have *: "inj_on (\<lambda>v. v + (y - a)) t" |
5140 |
unfolding inj_on_def by auto |
|
| 33175 | 5141 |
have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a))) = 1" |
| 64267 | 5142 |
unfolding sum.reindex[OF *] o_def using obt(4) by auto |
| 33175 | 5143 |
moreover have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R v) = y" |
| 64267 | 5144 |
unfolding sum.reindex[OF *] o_def using obt(4,5) |
5145 |
by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib) |
|
| 53347 | 5146 |
ultimately |
| 64267 | 5147 |
show "\<exists>sa u. finite sa \<and> (\<forall>x\<in>sa. x \<in> s) \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y" |
| 53347 | 5148 |
apply (rule_tac x="(\<lambda>v. v + (y - a)) ` t" in exI) |
5149 |
apply (rule_tac x="\<lambda>v. u (v - (y - a))" in exI) |
|
5150 |
using obt(1, 3) |
|
5151 |
apply auto |
|
5152 |
done |
|
| 33175 | 5153 |
qed |
5154 |
qed |
|
5155 |
||
5156 |
lemma compact_convex_combinations: |
|
5157 |
fixes s t :: "'a::real_normed_vector set" |
|
5158 |
assumes "compact s" "compact t" |
|
5159 |
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 | 5160 |
proof - |
| 33175 | 5161 |
let ?X = "{0..1} \<times> s \<times> t"
|
5162 |
let ?h = "(\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" |
|
| 53347 | 5163 |
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"
|
5164 |
apply (rule set_eqI) |
|
5165 |
unfolding image_iff mem_Collect_eq |
|
5166 |
apply rule |
|
5167 |
apply auto |
|
5168 |
apply (rule_tac x=u in rev_bexI) |
|
5169 |
apply simp |
|
5170 |
apply (erule rev_bexI) |
|
5171 |
apply (erule rev_bexI) |
|
5172 |
apply simp |
|
5173 |
apply auto |
|
5174 |
done |
|
| 56188 | 5175 |
have "continuous_on ?X (\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" |
| 33175 | 5176 |
unfolding continuous_on by (rule ballI) (intro tendsto_intros) |
| 53347 | 5177 |
then show ?thesis |
5178 |
unfolding * |
|
| 33175 | 5179 |
apply (rule compact_continuous_image) |
| 56188 | 5180 |
apply (intro compact_Times compact_Icc assms) |
| 33175 | 5181 |
done |
5182 |
qed |
|
5183 |
||
|
44525
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5184 |
lemma finite_imp_compact_convex_hull: |
| 53347 | 5185 |
fixes s :: "'a::real_normed_vector set" |
5186 |
assumes "finite s" |
|
5187 |
shows "compact (convex hull s)" |
|
|
44525
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5188 |
proof (cases "s = {}")
|
| 53347 | 5189 |
case True |
5190 |
then show ?thesis by simp |
|
|
44525
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5191 |
next |
| 53347 | 5192 |
case False |
5193 |
with assms show ?thesis |
|
|
44525
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5194 |
proof (induct rule: finite_ne_induct) |
| 53347 | 5195 |
case (singleton x) |
5196 |
show ?case by simp |
|
|
44525
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5197 |
next |
|
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5198 |
case (insert x A) |
|
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5199 |
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
|
5200 |
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
|
5201 |
have "continuous_on ?T ?f" |
|
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5202 |
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
|
5203 |
moreover have "compact ?T" |
| 56188 | 5204 |
by (intro compact_Times compact_Icc insert) |
|
44525
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5205 |
ultimately have "compact (?f ` ?T)" |
|
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5206 |
by (rule compact_continuous_image) |
|
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5207 |
also have "?f ` ?T = convex hull (insert x A)" |
| 60420 | 5208 |
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
|
5209 |
apply safe |
|
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5210 |
apply (rule_tac x=a in exI, simp) |
|
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5211 |
apply (rule_tac x="1 - a" in exI, simp) |
|
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5212 |
apply fast |
|
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5213 |
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
|
5214 |
done |
|
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5215 |
finally show "compact (convex hull (insert x A))" . |
|
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5216 |
qed |
|
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5217 |
qed |
|
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5218 |
|
| 53347 | 5219 |
lemma compact_convex_hull: |
5220 |
fixes s :: "'a::euclidean_space set" |
|
5221 |
assumes "compact s" |
|
5222 |
shows "compact (convex hull s)" |
|
5223 |
proof (cases "s = {}")
|
|
5224 |
case True |
|
5225 |
then show ?thesis using compact_empty by simp |
|
| 33175 | 5226 |
next |
| 53347 | 5227 |
case False |
5228 |
then obtain w where "w \<in> s" by auto |
|
5229 |
show ?thesis |
|
5230 |
unfolding caratheodory[of s] |
|
5231 |
proof (induct ("DIM('a) + 1"))
|
|
5232 |
case 0 |
|
5233 |
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
|
5234 |
using compact_empty by auto |
| 53347 | 5235 |
from 0 show ?case unfolding * by simp |
| 33175 | 5236 |
next |
5237 |
case (Suc n) |
|
| 53347 | 5238 |
show ?case |
5239 |
proof (cases "n = 0") |
|
5240 |
case True |
|
5241 |
have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} = s"
|
|
5242 |
unfolding set_eq_iff and mem_Collect_eq |
|
5243 |
proof (rule, rule) |
|
5244 |
fix x |
|
5245 |
assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" |
|
5246 |
then obtain t where t: "finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" |
|
5247 |
by auto |
|
5248 |
show "x \<in> s" |
|
5249 |
proof (cases "card t = 0") |
|
5250 |
case True |
|
5251 |
then show ?thesis |
|
5252 |
using t(4) unfolding card_0_eq[OF t(1)] by simp |
|
| 33175 | 5253 |
next |
| 53347 | 5254 |
case False |
| 60420 | 5255 |
then have "card t = Suc 0" using t(3) \<open>n=0\<close> by auto |
| 33175 | 5256 |
then obtain a where "t = {a}" unfolding card_Suc_eq by auto
|
| 53347 | 5257 |
then show ?thesis using t(2,4) by simp |
| 33175 | 5258 |
qed |
5259 |
next |
|
5260 |
fix x assume "x\<in>s" |
|
| 53347 | 5261 |
then show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" |
5262 |
apply (rule_tac x="{x}" in exI)
|
|
5263 |
unfolding convex_hull_singleton |
|
5264 |
apply auto |
|
5265 |
done |
|
5266 |
qed |
|
5267 |
then show ?thesis using assms by simp |
|
| 33175 | 5268 |
next |
| 53347 | 5269 |
case False |
5270 |
have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} =
|
|
5271 |
{(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u.
|
|
5272 |
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}}"
|
|
5273 |
unfolding set_eq_iff and mem_Collect_eq |
|
5274 |
proof (rule, rule) |
|
5275 |
fix x |
|
5276 |
assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and> |
|
| 33175 | 5277 |
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)" |
| 53347 | 5278 |
then obtain u v c t where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v" |
5279 |
"0 \<le> c \<and> c \<le> 1" "u \<in> s" "finite t" "t \<subseteq> s" "card t \<le> n" "v \<in> convex hull t" |
|
5280 |
by auto |
|
| 33175 | 5281 |
moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \<in> convex hull insert u t" |
|
61426
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61222
diff
changeset
|
5282 |
apply (rule convexD_alt) |
| 53347 | 5283 |
using obt(2) and convex_convex_hull and hull_subset[of "insert u t" convex] |
5284 |
using obt(7) and hull_mono[of t "insert u t"] |
|
5285 |
apply auto |
|
5286 |
done |
|
| 33175 | 5287 |
ultimately show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" |
| 53347 | 5288 |
apply (rule_tac x="insert u t" in exI) |
5289 |
apply (auto simp add: card_insert_if) |
|
5290 |
done |
|
| 33175 | 5291 |
next |
| 53347 | 5292 |
fix x |
5293 |
assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t" |
|
5294 |
then obtain t where t: "finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" |
|
5295 |
by auto |
|
5296 |
show "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and> |
|
| 33175 | 5297 |
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)" |
| 53347 | 5298 |
proof (cases "card t = Suc n") |
5299 |
case False |
|
5300 |
then have "card t \<le> n" using t(3) by auto |
|
5301 |
then show ?thesis |
|
5302 |
apply (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) |
|
| 60420 | 5303 |
using \<open>w\<in>s\<close> and t |
| 53347 | 5304 |
apply (auto intro!: exI[where x=t]) |
5305 |
done |
|
| 33175 | 5306 |
next |
| 53347 | 5307 |
case True |
5308 |
then obtain a u where au: "t = insert a u" "a\<notin>u" |
|
5309 |
apply (drule_tac card_eq_SucD) |
|
5310 |
apply auto |
|
5311 |
done |
|
5312 |
show ?thesis |
|
5313 |
proof (cases "u = {}")
|
|
5314 |
case True |
|
5315 |
then have "x = a" using t(4)[unfolded au] by auto |
|
| 60420 | 5316 |
show ?thesis unfolding \<open>x = a\<close> |
| 53347 | 5317 |
apply (rule_tac x=a in exI) |
5318 |
apply (rule_tac x=a in exI) |
|
5319 |
apply (rule_tac x=1 in exI) |
|
| 60420 | 5320 |
using t and \<open>n \<noteq> 0\<close> |
| 53347 | 5321 |
unfolding au |
5322 |
apply (auto intro!: exI[where x="{a}"])
|
|
5323 |
done |
|
| 33175 | 5324 |
next |
| 53347 | 5325 |
case False |
5326 |
obtain ux vx b where obt: "ux\<ge>0" "vx\<ge>0" "ux + vx = 1" |
|
5327 |
"b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b" |
|
5328 |
using t(4)[unfolded au convex_hull_insert[OF False]] |
|
5329 |
by auto |
|
5330 |
have *: "1 - vx = ux" using obt(3) by auto |
|
5331 |
show ?thesis |
|
5332 |
apply (rule_tac x=a in exI) |
|
5333 |
apply (rule_tac x=b in exI) |
|
5334 |
apply (rule_tac x=vx in exI) |
|
5335 |
using obt and t(1-3) |
|
5336 |
unfolding au and * using card_insert_disjoint[OF _ au(2)] |
|
5337 |
apply (auto intro!: exI[where x=u]) |
|
5338 |
done |
|
| 33175 | 5339 |
qed |
5340 |
qed |
|
5341 |
qed |
|
| 53347 | 5342 |
then show ?thesis |
5343 |
using compact_convex_combinations[OF assms Suc] by simp |
|
| 33175 | 5344 |
qed |
|
36362
06475a1547cb
fix lots of looping simp calls and other warnings
huffman
parents:
36341
diff
changeset
|
5345 |
qed |
| 33175 | 5346 |
qed |
5347 |
||
| 53347 | 5348 |
|
| 60420 | 5349 |
subsection \<open>Extremal points of a simplex are some vertices.\<close> |
| 33175 | 5350 |
|
5351 |
lemma dist_increases_online: |
|
5352 |
fixes a b d :: "'a::real_inner" |
|
5353 |
assumes "d \<noteq> 0" |
|
5354 |
shows "dist a (b + d) > dist a b \<or> dist a (b - d) > dist a b" |
|
| 53347 | 5355 |
proof (cases "inner a d - inner b d > 0") |
5356 |
case True |
|
5357 |
then have "0 < inner d d + (inner a d * 2 - inner b d * 2)" |
|
5358 |
apply (rule_tac add_pos_pos) |
|
5359 |
using assms |
|
5360 |
apply auto |
|
5361 |
done |
|
5362 |
then show ?thesis |
|
5363 |
apply (rule_tac disjI2) |
|
5364 |
unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff |
|
5365 |
apply (simp add: algebra_simps inner_commute) |
|
5366 |
done |
|
| 33175 | 5367 |
next |
| 53347 | 5368 |
case False |
5369 |
then have "0 < inner d d + (inner b d * 2 - inner a d * 2)" |
|
5370 |
apply (rule_tac add_pos_nonneg) |
|
5371 |
using assms |
|
5372 |
apply auto |
|
5373 |
done |
|
5374 |
then show ?thesis |
|
5375 |
apply (rule_tac disjI1) |
|
5376 |
unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff |
|
5377 |
apply (simp add: algebra_simps inner_commute) |
|
5378 |
done |
|
| 33175 | 5379 |
qed |
5380 |
||
5381 |
lemma norm_increases_online: |
|
5382 |
fixes d :: "'a::real_inner" |
|
| 53347 | 5383 |
shows "d \<noteq> 0 \<Longrightarrow> norm (a + d) > norm a \<or> norm(a - d) > norm a" |
| 33175 | 5384 |
using dist_increases_online[of d a 0] unfolding dist_norm by auto |
5385 |
||
5386 |
lemma simplex_furthest_lt: |
|
| 53347 | 5387 |
fixes s :: "'a::real_inner set" |
5388 |
assumes "finite s" |
|
5389 |
shows "\<forall>x \<in> convex hull s. x \<notin> s \<longrightarrow> (\<exists>y \<in> convex hull s. norm (x - a) < norm(y - a))" |
|
5390 |
using assms |
|
5391 |
proof induct |
|
5392 |
fix x s |
|
5393 |
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))" |
|
5394 |
show "\<forall>xa\<in>convex hull insert x s. xa \<notin> insert x s \<longrightarrow> |
|
5395 |
(\<exists>y\<in>convex hull insert x s. norm (xa - a) < norm (y - a))" |
|
5396 |
proof (rule, rule, cases "s = {}")
|
|
5397 |
case False |
|
5398 |
fix y |
|
5399 |
assume y: "y \<in> convex hull insert x s" "y \<notin> insert x s" |
|
5400 |
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 | 5401 |
using y(1)[unfolded convex_hull_insert[OF False]] by auto |
5402 |
show "\<exists>z\<in>convex hull insert x s. norm (y - a) < norm (z - a)" |
|
| 53347 | 5403 |
proof (cases "y \<in> convex hull s") |
5404 |
case True |
|
5405 |
then obtain z where "z \<in> convex hull s" "norm (y - a) < norm (z - a)" |
|
| 33175 | 5406 |
using as(3)[THEN bspec[where x=y]] and y(2) by auto |
| 53347 | 5407 |
then show ?thesis |
5408 |
apply (rule_tac x=z in bexI) |
|
5409 |
unfolding convex_hull_insert[OF False] |
|
5410 |
apply auto |
|
5411 |
done |
|
| 33175 | 5412 |
next |
| 53347 | 5413 |
case False |
5414 |
show ?thesis |
|
5415 |
using obt(3) |
|
5416 |
proof (cases "u = 0", case_tac[!] "v = 0") |
|
5417 |
assume "u = 0" "v \<noteq> 0" |
|
5418 |
then have "y = b" using obt by auto |
|
5419 |
then show ?thesis using False and obt(4) by auto |
|
| 33175 | 5420 |
next |
| 53347 | 5421 |
assume "u \<noteq> 0" "v = 0" |
5422 |
then have "y = x" using obt by auto |
|
5423 |
then show ?thesis using y(2) by auto |
|
5424 |
next |
|
5425 |
assume "u \<noteq> 0" "v \<noteq> 0" |
|
5426 |
then obtain w where w: "w>0" "w<u" "w<v" |
|
5427 |
using real_lbound_gt_zero[of u v] and obt(1,2) by auto |
|
5428 |
have "x \<noteq> b" |
|
5429 |
proof |
|
5430 |
assume "x = b" |
|
5431 |
then have "y = b" unfolding obt(5) |
|
5432 |
using obt(3) by (auto simp add: scaleR_left_distrib[symmetric]) |
|
5433 |
then show False using obt(4) and False by simp |
|
5434 |
qed |
|
5435 |
then have *: "w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto |
|
5436 |
show ?thesis |
|
5437 |
using dist_increases_online[OF *, of a y] |
|
5438 |
proof (elim disjE) |
|
| 33175 | 5439 |
assume "dist a y < dist a (y + w *\<^sub>R (x - b))" |
| 53347 | 5440 |
then have "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)" |
5441 |
unfolding dist_commute[of a] |
|
5442 |
unfolding dist_norm obt(5) |
|
5443 |
by (simp add: algebra_simps) |
|
| 33175 | 5444 |
moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x s" |
| 60420 | 5445 |
unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq
|
| 53347 | 5446 |
apply (rule_tac x="u + w" in exI) |
5447 |
apply rule |
|
5448 |
defer |
|
5449 |
apply (rule_tac x="v - w" in exI) |
|
| 60420 | 5450 |
using \<open>u \<ge> 0\<close> and w and obt(3,4) |
| 53347 | 5451 |
apply auto |
5452 |
done |
|
| 33175 | 5453 |
ultimately show ?thesis by auto |
5454 |
next |
|
5455 |
assume "dist a y < dist a (y - w *\<^sub>R (x - b))" |
|
| 53347 | 5456 |
then have "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)" |
5457 |
unfolding dist_commute[of a] |
|
5458 |
unfolding dist_norm obt(5) |
|
5459 |
by (simp add: algebra_simps) |
|
| 33175 | 5460 |
moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x s" |
| 60420 | 5461 |
unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq
|
| 53347 | 5462 |
apply (rule_tac x="u - w" in exI) |
5463 |
apply rule |
|
5464 |
defer |
|
5465 |
apply (rule_tac x="v + w" in exI) |
|
| 60420 | 5466 |
using \<open>u \<ge> 0\<close> and w and obt(3,4) |
| 53347 | 5467 |
apply auto |
5468 |
done |
|
| 33175 | 5469 |
ultimately show ?thesis by auto |
5470 |
qed |
|
5471 |
qed auto |
|
5472 |
qed |
|
5473 |
qed auto |
|
5474 |
qed (auto simp add: assms) |
|
5475 |
||
5476 |
lemma simplex_furthest_le: |
|
| 53347 | 5477 |
fixes s :: "'a::real_inner set" |
5478 |
assumes "finite s" |
|
5479 |
and "s \<noteq> {}"
|
|
5480 |
shows "\<exists>y\<in>s. \<forall>x\<in> convex hull s. norm (x - a) \<le> norm (y - a)" |
|
5481 |
proof - |
|
5482 |
have "convex hull s \<noteq> {}"
|
|
5483 |
using hull_subset[of s convex] and assms(2) by auto |
|
5484 |
then obtain x where x: "x \<in> convex hull s" "\<forall>y\<in>convex hull s. norm (y - a) \<le> norm (x - a)" |
|
| 33175 | 5485 |
using distance_attains_sup[OF finite_imp_compact_convex_hull[OF assms(1)], of a] |
| 53347 | 5486 |
unfolding dist_commute[of a] |
5487 |
unfolding dist_norm |
|
5488 |
by auto |
|
5489 |
show ?thesis |
|
5490 |
proof (cases "x \<in> s") |
|
5491 |
case False |
|
5492 |
then obtain y where "y \<in> convex hull s" "norm (x - a) < norm (y - a)" |
|
5493 |
using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) |
|
5494 |
by auto |
|
5495 |
then show ?thesis |
|
5496 |
using x(2)[THEN bspec[where x=y]] by auto |
|
5497 |
next |
|
5498 |
case True |
|
5499 |
with x show ?thesis by auto |
|
5500 |
qed |
|
| 33175 | 5501 |
qed |
5502 |
||
5503 |
lemma simplex_furthest_le_exists: |
|
|
44525
fbb777aec0d4
generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents:
44524
diff
changeset
|
5504 |
fixes s :: "('a::real_inner) set"
|
| 53347 | 5505 |
shows "finite s \<Longrightarrow> \<forall>x\<in>(convex hull s). \<exists>y\<in>s. norm (x - a) \<le> norm (y - a)" |
5506 |
using simplex_furthest_le[of s] by (cases "s = {}") auto
|
|
| 33175 | 5507 |
|
5508 |
lemma simplex_extremal_le: |
|
| 53347 | 5509 |
fixes s :: "'a::real_inner set" |
5510 |
assumes "finite s" |
|
5511 |
and "s \<noteq> {}"
|
|
5512 |
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)" |
|
5513 |
proof - |
|
5514 |
have "convex hull s \<noteq> {}"
|
|
5515 |
using hull_subset[of s convex] and assms(2) by auto |
|
5516 |
then obtain u v where obt: "u \<in> convex hull s" "v \<in> convex hull s" |
|
| 33175 | 5517 |
"\<forall>x\<in>convex hull s. \<forall>y\<in>convex hull s. norm (x - y) \<le> norm (u - v)" |
| 53347 | 5518 |
using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] |
5519 |
by (auto simp: dist_norm) |
|
5520 |
then show ?thesis |
|
5521 |
proof (cases "u\<notin>s \<or> v\<notin>s", elim disjE) |
|
5522 |
assume "u \<notin> s" |
|
5523 |
then obtain y where "y \<in> convex hull s" "norm (u - v) < norm (y - v)" |
|
5524 |
using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) |
|
5525 |
by auto |
|
5526 |
then show ?thesis |
|
5527 |
using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) |
|
5528 |
by auto |
|
| 33175 | 5529 |
next |
| 53347 | 5530 |
assume "v \<notin> s" |
5531 |
then obtain y where "y \<in> convex hull s" "norm (v - u) < norm (y - u)" |
|
5532 |
using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) |
|
5533 |
by auto |
|
5534 |
then show ?thesis |
|
5535 |
using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1) |
|
| 33175 | 5536 |
by (auto simp add: norm_minus_commute) |
5537 |
qed auto |
|
| 49531 | 5538 |
qed |
| 33175 | 5539 |
|
5540 |
lemma simplex_extremal_le_exists: |
|
| 53347 | 5541 |
fixes s :: "'a::real_inner set" |
5542 |
shows "finite s \<Longrightarrow> x \<in> convex hull s \<Longrightarrow> y \<in> convex hull s \<Longrightarrow> |
|
5543 |
\<exists>u\<in>s. \<exists>v\<in>s. norm (x - y) \<le> norm (u - v)" |
|
5544 |
using convex_hull_empty simplex_extremal_le[of s] |
|
5545 |
by(cases "s = {}") auto
|
|
5546 |
||
| 33175 | 5547 |
|
| 60420 | 5548 |
subsection \<open>Closest point of a convex set is unique, with a continuous projection.\<close> |
| 33175 | 5549 |
|
| 53347 | 5550 |
definition closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a"
|
5551 |
where "closest_point s a = (SOME x. x \<in> s \<and> (\<forall>y\<in>s. dist a x \<le> dist a y))" |
|
| 33175 | 5552 |
|
5553 |
lemma closest_point_exists: |
|
| 53347 | 5554 |
assumes "closed s" |
5555 |
and "s \<noteq> {}"
|
|
5556 |
shows "closest_point s a \<in> s" |
|
5557 |
and "\<forall>y\<in>s. dist a (closest_point s a) \<le> dist a y" |
|
5558 |
unfolding closest_point_def |
|
5559 |
apply(rule_tac[!] someI2_ex) |
|
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
5560 |
apply (auto intro: distance_attains_inf[OF assms(1,2), of a]) |
| 53347 | 5561 |
done |
5562 |
||
5563 |
lemma closest_point_in_set: "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> closest_point s a \<in> s"
|
|
5564 |
by (meson closest_point_exists) |
|
5565 |
||
5566 |
lemma closest_point_le: "closed s \<Longrightarrow> x \<in> s \<Longrightarrow> dist a (closest_point s a) \<le> dist a x" |
|
| 33175 | 5567 |
using closest_point_exists[of s] by auto |
5568 |
||
5569 |
lemma closest_point_self: |
|
| 53347 | 5570 |
assumes "x \<in> s" |
5571 |
shows "closest_point s x = x" |
|
5572 |
unfolding closest_point_def |
|
5573 |
apply (rule some1_equality, rule ex1I[of _ x]) |
|
5574 |
using assms |
|
5575 |
apply auto |
|
5576 |
done |
|
5577 |
||
5578 |
lemma closest_point_refl: "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> closest_point s x = x \<longleftrightarrow> x \<in> s"
|
|
5579 |
using closest_point_in_set[of s x] closest_point_self[of x s] |
|
5580 |
by auto |
|
| 33175 | 5581 |
|
| 36337 | 5582 |
lemma closer_points_lemma: |
| 33175 | 5583 |
assumes "inner y z > 0" |
5584 |
shows "\<exists>u>0. \<forall>v>0. v \<le> u \<longrightarrow> norm(v *\<^sub>R z - y) < norm y" |
|
| 53347 | 5585 |
proof - |
5586 |
have z: "inner z z > 0" |
|
5587 |
unfolding inner_gt_zero_iff using assms by auto |
|
5588 |
then show ?thesis |
|
5589 |
using assms |
|
5590 |
apply (rule_tac x = "inner y z / inner z z" in exI) |
|
5591 |
apply rule |
|
5592 |
defer |
|
5593 |
proof rule+ |
|
5594 |
fix v |
|
5595 |
assume "0 < v" and "v \<le> inner y z / inner z z" |
|
5596 |
then show "norm (v *\<^sub>R z - y) < norm y" |
|
5597 |
unfolding norm_lt using z and assms |
|
| 60420 | 5598 |
by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \<open>0<v\<close>]) |
| 56541 | 5599 |
qed auto |
| 53347 | 5600 |
qed |
| 33175 | 5601 |
|
5602 |
lemma closer_point_lemma: |
|
5603 |
assumes "inner (y - x) (z - x) > 0" |
|
5604 |
shows "\<exists>u>0. u \<le> 1 \<and> dist (x + u *\<^sub>R (z - x)) y < dist x y" |
|
| 53347 | 5605 |
proof - |
5606 |
obtain u where "u > 0" |
|
5607 |
and u: "\<forall>v>0. v \<le> u \<longrightarrow> norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)" |
|
| 33175 | 5608 |
using closer_points_lemma[OF assms] by auto |
| 53347 | 5609 |
show ?thesis |
5610 |
apply (rule_tac x="min u 1" in exI) |
|
| 60420 | 5611 |
using u[THEN spec[where x="min u 1"]] and \<open>u > 0\<close> |
| 53347 | 5612 |
unfolding dist_norm by (auto simp add: norm_minus_commute field_simps) |
5613 |
qed |
|
| 33175 | 5614 |
|
5615 |
lemma any_closest_point_dot: |
|
5616 |
assumes "convex s" "closed s" "x \<in> s" "y \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z" |
|
5617 |
shows "inner (a - x) (y - x) \<le> 0" |
|
| 53347 | 5618 |
proof (rule ccontr) |
5619 |
assume "\<not> ?thesis" |
|
5620 |
then obtain u where u: "u>0" "u\<le>1" "dist (x + u *\<^sub>R (y - x)) a < dist x a" |
|
5621 |
using closer_point_lemma[of a x y] by auto |
|
5622 |
let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y" |
|
5623 |
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
|
5624 |
using convexD_alt[OF assms(1,3,4), of u] using u by auto |
| 53347 | 5625 |
then show False |
5626 |
using assms(5)[THEN bspec[where x="?z"]] and u(3) |
|
5627 |
by (auto simp add: dist_commute algebra_simps) |
|
5628 |
qed |
|
| 33175 | 5629 |
|
5630 |
lemma any_closest_point_unique: |
|
| 36337 | 5631 |
fixes x :: "'a::real_inner" |
| 33175 | 5632 |
assumes "convex s" "closed s" "x \<in> s" "y \<in> s" |
| 53347 | 5633 |
"\<forall>z\<in>s. dist a x \<le> dist a z" "\<forall>z\<in>s. dist a y \<le> dist a z" |
5634 |
shows "x = y" |
|
5635 |
using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)] |
|
| 33175 | 5636 |
unfolding norm_pths(1) and norm_le_square |
5637 |
by (auto simp add: algebra_simps) |
|
5638 |
||
5639 |
lemma closest_point_unique: |
|
5640 |
assumes "convex s" "closed s" "x \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z" |
|
5641 |
shows "x = closest_point s a" |
|
| 49531 | 5642 |
using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point s a"] |
| 33175 | 5643 |
using closest_point_exists[OF assms(2)] and assms(3) by auto |
5644 |
||
5645 |
lemma closest_point_dot: |
|
5646 |
assumes "convex s" "closed s" "x \<in> s" |
|
5647 |
shows "inner (a - closest_point s a) (x - closest_point s a) \<le> 0" |
|
| 53347 | 5648 |
apply (rule any_closest_point_dot[OF assms(1,2) _ assms(3)]) |
5649 |
using closest_point_exists[OF assms(2)] and assms(3) |
|
5650 |
apply auto |
|
5651 |
done |
|
| 33175 | 5652 |
|
5653 |
lemma closest_point_lt: |
|
5654 |
assumes "convex s" "closed s" "x \<in> s" "x \<noteq> closest_point s a" |
|
5655 |
shows "dist a (closest_point s a) < dist a x" |
|
| 53347 | 5656 |
apply (rule ccontr) |
5657 |
apply (rule_tac notE[OF assms(4)]) |
|
5658 |
apply (rule closest_point_unique[OF assms(1-3), of a]) |
|
5659 |
using closest_point_le[OF assms(2), of _ a] |
|
5660 |
apply fastforce |
|
5661 |
done |
|
| 33175 | 5662 |
|
5663 |
lemma closest_point_lipschitz: |
|
| 53347 | 5664 |
assumes "convex s" |
5665 |
and "closed s" "s \<noteq> {}"
|
|
| 33175 | 5666 |
shows "dist (closest_point s x) (closest_point s y) \<le> dist x y" |
| 53347 | 5667 |
proof - |
| 33175 | 5668 |
have "inner (x - closest_point s x) (closest_point s y - closest_point s x) \<le> 0" |
| 53347 | 5669 |
and "inner (y - closest_point s y) (closest_point s x - closest_point s y) \<le> 0" |
5670 |
apply (rule_tac[!] any_closest_point_dot[OF assms(1-2)]) |
|
5671 |
using closest_point_exists[OF assms(2-3)] |
|
5672 |
apply auto |
|
5673 |
done |
|
5674 |
then show ?thesis unfolding dist_norm and norm_le |
|
| 33175 | 5675 |
using inner_ge_zero[of "(x - closest_point s x) - (y - closest_point s y)"] |
| 53347 | 5676 |
by (simp add: inner_add inner_diff inner_commute) |
5677 |
qed |
|
| 33175 | 5678 |
|
5679 |
lemma continuous_at_closest_point: |
|
| 53347 | 5680 |
assumes "convex s" |
5681 |
and "closed s" |
|
5682 |
and "s \<noteq> {}"
|
|
| 33175 | 5683 |
shows "continuous (at x) (closest_point s)" |
| 49531 | 5684 |
unfolding continuous_at_eps_delta |
| 33175 | 5685 |
using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto |
5686 |
||
5687 |
lemma continuous_on_closest_point: |
|
| 53347 | 5688 |
assumes "convex s" |
5689 |
and "closed s" |
|
5690 |
and "s \<noteq> {}"
|
|
| 33175 | 5691 |
shows "continuous_on t (closest_point s)" |
| 53347 | 5692 |
by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms]) |
5693 |
||
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5694 |
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
|
5695 |
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
|
5696 |
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
|
5697 |
proof (cases "x \<in> S") |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5698 |
case True |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5699 |
then show ?thesis |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5700 |
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
|
5701 |
next |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5702 |
case False |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5703 |
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
|
5704 |
proof - |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5705 |
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
|
5706 |
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
|
5707 |
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
|
5708 |
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
|
5709 |
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
|
5710 |
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
|
5711 |
(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
|
5712 |
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
|
5713 |
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
|
5714 |
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
|
5715 |
by simp |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5716 |
also have "... < 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
|
5717 |
using clo_notx \<open>e > 0\<close> |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5718 |
by (auto simp: mult_less_cancel_right2 divide_simps) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5719 |
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
|
5720 |
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
|
5721 |
unfolding y_def |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5722 |
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
|
5723 |
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
|
5724 |
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
|
5725 |
ultimately have "y \<in> S" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5726 |
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
|
5727 |
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
|
5728 |
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
|
5729 |
with no_less show False |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5730 |
by (simp add: dist_norm) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5731 |
qed |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5732 |
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
|
5733 |
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
|
5734 |
ultimately show ?thesis by blast |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5735 |
qed |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5736 |
|
| 33175 | 5737 |
|
| 60420 | 5738 |
subsubsection \<open>Various point-to-set separating/supporting hyperplane theorems.\<close> |
| 33175 | 5739 |
|
5740 |
lemma supporting_hyperplane_closed_point: |
|
| 36337 | 5741 |
fixes z :: "'a::{real_inner,heine_borel}"
|
| 53347 | 5742 |
assumes "convex s" |
5743 |
and "closed s" |
|
5744 |
and "s \<noteq> {}"
|
|
5745 |
and "z \<notin> s" |
|
5746 |
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)" |
|
5747 |
proof - |
|
5748 |
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
|
5749 |
by (metis distance_attains_inf[OF assms(2-3)]) |
| 53347 | 5750 |
show ?thesis |
5751 |
apply (rule_tac x="y - z" in exI) |
|
5752 |
apply (rule_tac x="inner (y - z) y" in exI) |
|
5753 |
apply (rule_tac x=y in bexI) |
|
5754 |
apply rule |
|
5755 |
defer |
|
5756 |
apply rule |
|
5757 |
defer |
|
5758 |
apply rule |
|
5759 |
apply (rule ccontr) |
|
| 60420 | 5760 |
using \<open>y \<in> s\<close> |
| 53347 | 5761 |
proof - |
5762 |
show "inner (y - z) z < inner (y - z) y" |
|
|
61762
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents:
61738
diff
changeset
|
5763 |
apply (subst diff_gt_0_iff_gt [symmetric]) |
| 53347 | 5764 |
unfolding inner_diff_right[symmetric] and inner_gt_zero_iff |
| 60420 | 5765 |
using \<open>y\<in>s\<close> \<open>z\<notin>s\<close> |
| 53347 | 5766 |
apply auto |
5767 |
done |
|
| 33175 | 5768 |
next |
| 53347 | 5769 |
fix x |
5770 |
assume "x \<in> s" |
|
5771 |
have *: "\<forall>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> dist z y \<le> dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)" |
|
| 60420 | 5772 |
using assms(1)[unfolded convex_alt] and y and \<open>x\<in>s\<close> and \<open>y\<in>s\<close> by auto |
| 53347 | 5773 |
assume "\<not> inner (y - z) y \<le> inner (y - z) x" |
5774 |
then obtain v where "v > 0" "v \<le> 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" |
|
5775 |
using closer_point_lemma[of z y x] by (auto simp add: inner_diff) |
|
5776 |
then show False |
|
5777 |
using *[THEN spec[where x=v]] by (auto simp add: dist_commute algebra_simps) |
|
| 33175 | 5778 |
qed auto |
5779 |
qed |
|
5780 |
||
5781 |
lemma separating_hyperplane_closed_point: |
|
| 36337 | 5782 |
fixes z :: "'a::{real_inner,heine_borel}"
|
| 53347 | 5783 |
assumes "convex s" |
5784 |
and "closed s" |
|
5785 |
and "z \<notin> s" |
|
| 33175 | 5786 |
shows "\<exists>a b. inner a z < b \<and> (\<forall>x\<in>s. inner a x > b)" |
| 53347 | 5787 |
proof (cases "s = {}")
|
5788 |
case True |
|
5789 |
then show ?thesis |
|
5790 |
apply (rule_tac x="-z" in exI) |
|
5791 |
apply (rule_tac x=1 in exI) |
|
5792 |
using less_le_trans[OF _ inner_ge_zero[of z]] |
|
5793 |
apply auto |
|
5794 |
done |
|
| 33175 | 5795 |
next |
| 53347 | 5796 |
case False |
5797 |
obtain y where "y \<in> s" and y: "\<forall>x\<in>s. 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
|
5798 |
by (metis distance_attains_inf[OF assms(2) False]) |
| 53347 | 5799 |
show ?thesis |
5800 |
apply (rule_tac x="y - z" in exI) |
|
5801 |
apply (rule_tac x="inner (y - z) z + (norm (y - z))\<^sup>2 / 2" in exI) |
|
5802 |
apply rule |
|
5803 |
defer |
|
5804 |
apply rule |
|
5805 |
proof - |
|
5806 |
fix x |
|
5807 |
assume "x \<in> s" |
|
5808 |
have "\<not> 0 < inner (z - y) (x - y)" |
|
5809 |
apply (rule notI) |
|
5810 |
apply (drule closer_point_lemma) |
|
5811 |
proof - |
|
| 33175 | 5812 |
assume "\<exists>u>0. u \<le> 1 \<and> dist (y + u *\<^sub>R (x - y)) z < dist y z" |
| 53347 | 5813 |
then obtain u where "u > 0" "u \<le> 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" |
5814 |
by auto |
|
5815 |
then show False using y[THEN bspec[where x="y + u *\<^sub>R (x - y)"]] |
|
| 33175 | 5816 |
using assms(1)[unfolded convex_alt, THEN bspec[where x=y]] |
| 60420 | 5817 |
using \<open>x\<in>s\<close> \<open>y\<in>s\<close> by (auto simp add: dist_commute algebra_simps) |
| 53347 | 5818 |
qed |
5819 |
moreover have "0 < (norm (y - z))\<^sup>2" |
|
| 60420 | 5820 |
using \<open>y\<in>s\<close> \<open>z\<notin>s\<close> by auto |
| 53347 | 5821 |
then have "0 < inner (y - z) (y - z)" |
5822 |
unfolding power2_norm_eq_inner by simp |
|
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51524
diff
changeset
|
5823 |
ultimately show "inner (y - z) z + (norm (y - z))\<^sup>2 / 2 < inner (y - z) x" |
| 53347 | 5824 |
unfolding power2_norm_eq_inner and not_less |
5825 |
by (auto simp add: field_simps inner_commute inner_diff) |
|
| 60420 | 5826 |
qed (insert \<open>y\<in>s\<close> \<open>z\<notin>s\<close>, auto) |
| 33175 | 5827 |
qed |
5828 |
||
5829 |
lemma separating_hyperplane_closed_0: |
|
| 53347 | 5830 |
assumes "convex (s::('a::euclidean_space) set)"
|
5831 |
and "closed s" |
|
5832 |
and "0 \<notin> s" |
|
| 33175 | 5833 |
shows "\<exists>a b. a \<noteq> 0 \<and> 0 < b \<and> (\<forall>x\<in>s. inner a x > b)" |
| 53347 | 5834 |
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
|
5835 |
case True |
| 53347 | 5836 |
have "norm ((SOME i. i\<in>Basis)::'a) = 1" "(SOME i. i\<in>Basis) \<noteq> (0::'a)" |
5837 |
defer |
|
5838 |
apply (subst norm_le_zero_iff[symmetric]) |
|
5839 |
apply (auto simp: SOME_Basis) |
|
5840 |
done |
|
5841 |
then show ?thesis |
|
5842 |
apply (rule_tac x="SOME i. i\<in>Basis" in exI) |
|
5843 |
apply (rule_tac x=1 in exI) |
|
5844 |
using True using DIM_positive[where 'a='a] |
|
5845 |
apply auto |
|
5846 |
done |
|
5847 |
next |
|
5848 |
case False |
|
5849 |
then show ?thesis |
|
5850 |
using False using separating_hyperplane_closed_point[OF assms] |
|
5851 |
apply (elim exE) |
|
5852 |
unfolding inner_zero_right |
|
5853 |
apply (rule_tac x=a in exI) |
|
5854 |
apply (rule_tac x=b in exI) |
|
5855 |
apply auto |
|
5856 |
done |
|
5857 |
qed |
|
5858 |
||
| 33175 | 5859 |
|
| 60420 | 5860 |
subsubsection \<open>Now set-to-set for closed/compact sets\<close> |
| 33175 | 5861 |
|
5862 |
lemma separating_hyperplane_closed_compact: |
|
| 53347 | 5863 |
fixes s :: "'a::euclidean_space set" |
5864 |
assumes "convex s" |
|
5865 |
and "closed s" |
|
5866 |
and "convex t" |
|
5867 |
and "compact t" |
|
5868 |
and "t \<noteq> {}"
|
|
5869 |
and "s \<inter> t = {}"
|
|
| 33175 | 5870 |
shows "\<exists>a b. (\<forall>x\<in>s. inner a x < b) \<and> (\<forall>x\<in>t. inner a x > b)" |
| 53347 | 5871 |
proof (cases "s = {}")
|
| 33175 | 5872 |
case True |
| 53347 | 5873 |
obtain b where b: "b > 0" "\<forall>x\<in>t. norm x \<le> b" |
5874 |
using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto |
|
5875 |
obtain z :: 'a where z: "norm z = b + 1" |
|
5876 |
using vector_choose_size[of "b + 1"] and b(1) by auto |
|
5877 |
then have "z \<notin> t" using b(2)[THEN bspec[where x=z]] by auto |
|
5878 |
then obtain a b where ab: "inner a z < b" "\<forall>x\<in>t. b < inner a x" |
|
5879 |
using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z] |
|
5880 |
by auto |
|
5881 |
then show ?thesis |
|
5882 |
using True by auto |
|
| 33175 | 5883 |
next |
| 53347 | 5884 |
case False |
5885 |
then obtain y where "y \<in> s" by auto |
|
| 33175 | 5886 |
obtain a b where "0 < b" "\<forall>x\<in>{x - y |x y. x \<in> s \<and> y \<in> t}. b < inner a x"
|
5887 |
using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0] |
|
| 53347 | 5888 |
using closed_compact_differences[OF assms(2,4)] |
5889 |
using assms(6) by auto blast |
|
5890 |
then have ab: "\<forall>x\<in>s. \<forall>y\<in>t. b + inner a y < inner a x" |
|
5891 |
apply - |
|
5892 |
apply rule |
|
5893 |
apply rule |
|
5894 |
apply (erule_tac x="x - y" in ballE) |
|
5895 |
apply (auto simp add: inner_diff) |
|
5896 |
done |
|
| 63040 | 5897 |
define k where "k = (SUP x:t. a \<bullet> x)" |
| 53347 | 5898 |
show ?thesis |
5899 |
apply (rule_tac x="-a" in exI) |
|
5900 |
apply (rule_tac x="-(k + b / 2)" in exI) |
|
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
5901 |
apply (intro conjI ballI) |
| 53347 | 5902 |
unfolding inner_minus_left and neg_less_iff_less |
5903 |
proof - |
|
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
5904 |
fix x assume "x \<in> t" |
|
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
5905 |
then have "inner a x - b / 2 < k" |
| 53347 | 5906 |
unfolding k_def |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
5907 |
proof (subst less_cSUP_iff) |
|
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
5908 |
show "t \<noteq> {}" by fact
|
|
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
5909 |
show "bdd_above (op \<bullet> a ` t)" |
| 60420 | 5910 |
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
|
5911 |
by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le) |
| 60420 | 5912 |
qed (auto intro!: bexI[of _ x] \<open>0<b\<close>) |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
5913 |
then show "inner a x < k + b / 2" |
|
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
5914 |
by auto |
| 33175 | 5915 |
next |
| 53347 | 5916 |
fix x |
5917 |
assume "x \<in> s" |
|
5918 |
then have "k \<le> inner a x - b" |
|
5919 |
unfolding k_def |
|
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
5920 |
apply (rule_tac cSUP_least) |
| 53347 | 5921 |
using assms(5) |
5922 |
using ab[THEN bspec[where x=x]] |
|
5923 |
apply auto |
|
5924 |
done |
|
5925 |
then show "k + b / 2 < inner a x" |
|
| 60420 | 5926 |
using \<open>0 < b\<close> by auto |
| 33175 | 5927 |
qed |
5928 |
qed |
|
5929 |
||
5930 |
lemma separating_hyperplane_compact_closed: |
|
| 53347 | 5931 |
fixes s :: "'a::euclidean_space set" |
5932 |
assumes "convex s" |
|
5933 |
and "compact s" |
|
5934 |
and "s \<noteq> {}"
|
|
5935 |
and "convex t" |
|
5936 |
and "closed t" |
|
5937 |
and "s \<inter> t = {}"
|
|
| 33175 | 5938 |
shows "\<exists>a b. (\<forall>x\<in>s. inner a x < b) \<and> (\<forall>x\<in>t. inner a x > b)" |
| 53347 | 5939 |
proof - |
5940 |
obtain a b where "(\<forall>x\<in>t. inner a x < b) \<and> (\<forall>x\<in>s. b < inner a x)" |
|
5941 |
using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6) |
|
5942 |
by auto |
|
5943 |
then show ?thesis |
|
5944 |
apply (rule_tac x="-a" in exI) |
|
5945 |
apply (rule_tac x="-b" in exI) |
|
5946 |
apply auto |
|
5947 |
done |
|
5948 |
qed |
|
5949 |
||
| 33175 | 5950 |
|
| 60420 | 5951 |
subsubsection \<open>General case without assuming closure and getting non-strict separation\<close> |
| 33175 | 5952 |
|
5953 |
lemma separating_hyperplane_set_0: |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
5954 |
assumes "convex s" "(0::'a::euclidean_space) \<notin> s" |
| 33175 | 5955 |
shows "\<exists>a. a \<noteq> 0 \<and> (\<forall>x\<in>s. 0 \<le> inner a x)" |
| 53347 | 5956 |
proof - |
5957 |
let ?k = "\<lambda>c. {x::'a. 0 \<le> inner c x}"
|
|
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
5958 |
have *: "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}" if as: "f \<subseteq> ?k ` s" "finite f" for f
|
| 53347 | 5959 |
proof - |
5960 |
obtain c where c: "f = ?k ` c" "c \<subseteq> s" "finite c" |
|
5961 |
using finite_subset_image[OF as(2,1)] by auto |
|
5962 |
then obtain a b where ab: "a \<noteq> 0" "0 < b" "\<forall>x\<in>convex hull c. b < inner a x" |
|
| 33175 | 5963 |
using separating_hyperplane_closed_0[OF convex_convex_hull, of c] |
5964 |
using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2) |
|
| 53347 | 5965 |
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
|
5966 |
by force |
| 53347 | 5967 |
then have "\<exists>x. norm x = 1 \<and> (\<forall>y\<in>c. 0 \<le> inner y x)" |
5968 |
apply (rule_tac x = "inverse(norm a) *\<^sub>R a" in exI) |
|
5969 |
using hull_subset[of c convex] |
|
5970 |
unfolding subset_eq and inner_scaleR |
|
| 56536 | 5971 |
by (auto simp add: inner_commute del: ballE elim!: ballE) |
| 53347 | 5972 |
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
|
5973 |
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
|
5974 |
qed |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
5975 |
have "frontier (cball 0 1) \<inter> (\<Inter>(?k ` s)) \<noteq> {}"
|
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
5976 |
apply (rule compact_imp_fip) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
5977 |
apply (rule compact_frontier[OF compact_cball]) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
5978 |
using * closed_halfspace_ge |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
5979 |
by auto |
| 53347 | 5980 |
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
|
5981 |
unfolding frontier_cball dist_norm sphere_def by auto |
| 53347 | 5982 |
then show ?thesis |
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
5983 |
by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one) |
| 53347 | 5984 |
qed |
| 33175 | 5985 |
|
5986 |
lemma separating_hyperplane_sets: |
|
| 53347 | 5987 |
fixes s t :: "'a::euclidean_space set" |
5988 |
assumes "convex s" |
|
5989 |
and "convex t" |
|
5990 |
and "s \<noteq> {}"
|
|
5991 |
and "t \<noteq> {}"
|
|
5992 |
and "s \<inter> t = {}"
|
|
| 33175 | 5993 |
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 | 5994 |
proof - |
5995 |
from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] |
|
5996 |
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"
|
|
|
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
|
5997 |
using assms(3-5) by fastforce |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
5998 |
then have *: "\<And>x y. x \<in> t \<Longrightarrow> y \<in> s \<Longrightarrow> inner a y \<le> inner a x" |
| 33270 | 5999 |
by (force simp add: inner_diff) |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
6000 |
then have bdd: "bdd_above ((op \<bullet> a)`s)" |
| 60420 | 6001 |
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
|
6002 |
show ?thesis |
| 60420 | 6003 |
using \<open>a\<noteq>0\<close> |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
54258
diff
changeset
|
6004 |
by (intro exI[of _ a] exI[of _ "SUP x:s. a \<bullet> x"]) |
| 60420 | 6005 |
(auto intro!: cSUP_upper bdd cSUP_least \<open>a \<noteq> 0\<close> \<open>s \<noteq> {}\<close> *)
|
6006 |
qed |
|
6007 |
||
6008 |
||
6009 |
subsection \<open>More convexity generalities\<close> |
|
| 33175 | 6010 |
|
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
6011 |
lemma convex_closure [intro,simp]: |
| 33175 | 6012 |
fixes s :: "'a::real_normed_vector set" |
| 53347 | 6013 |
assumes "convex s" |
6014 |
shows "convex (closure s)" |
|
| 53676 | 6015 |
apply (rule convexI) |
6016 |
apply (unfold closure_sequential, elim exE) |
|
6017 |
apply (rule_tac x="\<lambda>n. u *\<^sub>R xa n + v *\<^sub>R xb n" in exI) |
|
| 53347 | 6018 |
apply (rule,rule) |
| 53676 | 6019 |
apply (rule convexD [OF assms]) |
| 53347 | 6020 |
apply (auto del: tendsto_const intro!: tendsto_intros) |
6021 |
done |
|
| 33175 | 6022 |
|
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
6023 |
lemma convex_interior [intro,simp]: |
| 33175 | 6024 |
fixes s :: "'a::real_normed_vector set" |
| 53347 | 6025 |
assumes "convex s" |
6026 |
shows "convex (interior s)" |
|
6027 |
unfolding convex_alt Ball_def mem_interior |
|
6028 |
apply (rule,rule,rule,rule,rule,rule) |
|
6029 |
apply (elim exE conjE) |
|
6030 |
proof - |
|
6031 |
fix x y u |
|
6032 |
assume u: "0 \<le> u" "u \<le> (1::real)" |
|
6033 |
fix e d |
|
6034 |
assume ed: "ball x e \<subseteq> s" "ball y d \<subseteq> s" "0<d" "0<e" |
|
6035 |
show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> s" |
|
6036 |
apply (rule_tac x="min d e" in exI) |
|
6037 |
apply rule |
|
6038 |
unfolding subset_eq |
|
6039 |
defer |
|
6040 |
apply rule |
|
6041 |
proof - |
|
6042 |
fix z |
|
6043 |
assume "z \<in> ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)" |
|
6044 |
then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> s" |
|
6045 |
apply (rule_tac assms[unfolded convex_alt, rule_format]) |
|
6046 |
using ed(1,2) and u |
|
6047 |
unfolding subset_eq mem_ball Ball_def dist_norm |
|
6048 |
apply (auto simp add: algebra_simps) |
|
6049 |
done |
|
6050 |
then show "z \<in> s" |
|
6051 |
using u by (auto simp add: algebra_simps) |
|
6052 |
qed(insert u ed(3-4), auto) |
|
6053 |
qed |
|
| 33175 | 6054 |
|
| 34964 | 6055 |
lemma convex_hull_eq_empty[simp]: "convex hull s = {} \<longleftrightarrow> s = {}"
|
| 33175 | 6056 |
using hull_subset[of s convex] convex_hull_empty by auto |
6057 |
||
| 53347 | 6058 |
|
| 60420 | 6059 |
subsection \<open>Moving and scaling convex hulls.\<close> |
| 33175 | 6060 |
|
| 53676 | 6061 |
lemma convex_hull_set_plus: |
6062 |
"convex hull (s + t) = convex hull s + convex hull t" |
|
6063 |
unfolding set_plus_image |
|
6064 |
apply (subst convex_hull_linear_image [symmetric]) |
|
6065 |
apply (simp add: linear_iff scaleR_right_distrib) |
|
6066 |
apply (simp add: convex_hull_Times) |
|
6067 |
done |
|
6068 |
||
6069 |
lemma translation_eq_singleton_plus: "(\<lambda>x. a + x) ` t = {a} + t"
|
|
6070 |
unfolding set_plus_def by auto |
|
| 33175 | 6071 |
|
6072 |
lemma convex_hull_translation: |
|
6073 |
"convex hull ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (convex hull s)" |
|
| 53676 | 6074 |
unfolding translation_eq_singleton_plus |
6075 |
by (simp only: convex_hull_set_plus convex_hull_singleton) |
|
| 33175 | 6076 |
|
6077 |
lemma convex_hull_scaling: |
|
6078 |
"convex hull ((\<lambda>x. c *\<^sub>R x) ` s) = (\<lambda>x. c *\<^sub>R x) ` (convex hull s)" |
|
| 53676 | 6079 |
using linear_scaleR by (rule convex_hull_linear_image [symmetric]) |
| 33175 | 6080 |
|
6081 |
lemma convex_hull_affinity: |
|
6082 |
"convex hull ((\<lambda>x. a + c *\<^sub>R x) ` s) = (\<lambda>x. a + c *\<^sub>R x) ` (convex hull s)" |
|
| 53347 | 6083 |
by(simp only: image_image[symmetric] convex_hull_scaling convex_hull_translation) |
6084 |
||
| 33175 | 6085 |
|
| 60420 | 6086 |
subsection \<open>Convexity of cone hulls\<close> |
| 40377 | 6087 |
|
6088 |
lemma convex_cone_hull: |
|
| 53347 | 6089 |
assumes "convex S" |
6090 |
shows "convex (cone hull S)" |
|
| 53676 | 6091 |
proof (rule convexI) |
6092 |
fix x y |
|
6093 |
assume xy: "x \<in> cone hull S" "y \<in> cone hull S" |
|
6094 |
then have "S \<noteq> {}"
|
|
6095 |
using cone_hull_empty_iff[of S] by auto |
|
6096 |
fix u v :: real |
|
6097 |
assume uv: "u \<ge> 0" "v \<ge> 0" "u + v = 1" |
|
6098 |
then have *: "u *\<^sub>R x \<in> cone hull S" "v *\<^sub>R y \<in> cone hull S" |
|
6099 |
using cone_cone_hull[of S] xy cone_def[of "cone hull S"] by auto |
|
6100 |
from * obtain cx :: real and xx where x: "u *\<^sub>R x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S" |
|
6101 |
using cone_hull_expl[of S] by auto |
|
6102 |
from * obtain cy :: real and yy where y: "v *\<^sub>R y = cy *\<^sub>R yy" "cy \<ge> 0" "yy \<in> S" |
|
6103 |
using cone_hull_expl[of S] by auto |
|
| 53347 | 6104 |
{
|
| 53676 | 6105 |
assume "cx + cy \<le> 0" |
6106 |
then have "u *\<^sub>R x = 0" and "v *\<^sub>R y = 0" |
|
6107 |
using x y by auto |
|
6108 |
then have "u *\<^sub>R x + v *\<^sub>R y = 0" |
|
6109 |
by auto |
|
6110 |
then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" |
|
| 60420 | 6111 |
using cone_hull_contains_0[of S] \<open>S \<noteq> {}\<close> by auto
|
| 40377 | 6112 |
} |
| 53676 | 6113 |
moreover |
6114 |
{
|
|
6115 |
assume "cx + cy > 0" |
|
6116 |
then have "(cx / (cx + cy)) *\<^sub>R xx + (cy / (cx + cy)) *\<^sub>R yy \<in> S" |
|
6117 |
using assms mem_convex_alt[of S xx yy cx cy] x y by auto |
|
6118 |
then have "cx *\<^sub>R xx + cy *\<^sub>R yy \<in> cone hull S" |
|
| 60420 | 6119 |
using mem_cone_hull[of "(cx/(cx+cy)) *\<^sub>R xx + (cy/(cx+cy)) *\<^sub>R yy" S "cx+cy"] \<open>cx+cy>0\<close> |
| 53676 | 6120 |
by (auto simp add: scaleR_right_distrib) |
6121 |
then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" |
|
6122 |
using x y by auto |
|
6123 |
} |
|
6124 |
moreover have "cx + cy \<le> 0 \<or> cx + cy > 0" by auto |
|
6125 |
ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" by blast |
|
| 40377 | 6126 |
qed |
6127 |
||
6128 |
lemma cone_convex_hull: |
|
| 53347 | 6129 |
assumes "cone S" |
6130 |
shows "cone (convex hull S)" |
|
6131 |
proof (cases "S = {}")
|
|
6132 |
case True |
|
6133 |
then show ?thesis by auto |
|
6134 |
next |
|
6135 |
case False |
|
| 54465 | 6136 |
then have *: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)" |
6137 |
using cone_iff[of S] assms by auto |
|
| 53347 | 6138 |
{
|
6139 |
fix c :: real |
|
6140 |
assume "c > 0" |
|
6141 |
then have "op *\<^sub>R c ` (convex hull S) = convex hull (op *\<^sub>R c ` S)" |
|
6142 |
using convex_hull_scaling[of _ S] by auto |
|
6143 |
also have "\<dots> = convex hull S" |
|
| 60420 | 6144 |
using * \<open>c > 0\<close> by auto |
| 53347 | 6145 |
finally have "op *\<^sub>R c ` (convex hull S) = convex hull S" |
6146 |
by auto |
|
| 40377 | 6147 |
} |
| 53347 | 6148 |
then have "0 \<in> convex hull S" "\<And>c. c > 0 \<Longrightarrow> (op *\<^sub>R c ` (convex hull S)) = (convex hull S)" |
6149 |
using * hull_subset[of S convex] by auto |
|
6150 |
then show ?thesis |
|
| 60420 | 6151 |
using \<open>S \<noteq> {}\<close> cone_iff[of "convex hull S"] by auto
|
6152 |
qed |
|
6153 |
||
6154 |
subsection \<open>Convex set as intersection of halfspaces\<close> |
|
| 33175 | 6155 |
|
6156 |
lemma convex_halfspace_intersection: |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
6157 |
fixes s :: "('a::euclidean_space) set"
|
| 33175 | 6158 |
assumes "closed s" "convex s" |
| 60585 | 6159 |
shows "s = \<Inter>{h. s \<subseteq> h \<and> (\<exists>a b. h = {x. inner a x \<le> b})}"
|
| 53347 | 6160 |
apply (rule set_eqI) |
6161 |
apply rule |
|
6162 |
unfolding Inter_iff Ball_def mem_Collect_eq |
|
6163 |
apply (rule,rule,erule conjE) |
|
6164 |
proof - |
|
| 54465 | 6165 |
fix x |
| 53347 | 6166 |
assume "\<forall>xa. s \<subseteq> xa \<and> (\<exists>a b. xa = {x. inner a x \<le> b}) \<longrightarrow> x \<in> xa"
|
6167 |
then have "\<forall>a b. s \<subseteq> {x. inner a x \<le> b} \<longrightarrow> x \<in> {x. inner a x \<le> b}"
|
|
6168 |
by blast |
|
6169 |
then show "x \<in> s" |
|
6170 |
apply (rule_tac ccontr) |
|
6171 |
apply (drule separating_hyperplane_closed_point[OF assms(2,1)]) |
|
6172 |
apply (erule exE)+ |
|
6173 |
apply (erule_tac x="-a" in allE) |
|
6174 |
apply (erule_tac x="-b" in allE) |
|
6175 |
apply auto |
|
6176 |
done |
|
| 33175 | 6177 |
qed auto |
6178 |
||
| 53347 | 6179 |
|
| 60420 | 6180 |
subsection \<open>Radon's theorem (from Lars Schewe)\<close> |
| 33175 | 6181 |
|
6182 |
lemma radon_ex_lemma: |
|
6183 |
assumes "finite c" "affine_dependent c" |
|
| 64267 | 6184 |
shows "\<exists>u. sum u c = 0 \<and> (\<exists>v\<in>c. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) c = 0" |
| 53347 | 6185 |
proof - |
| 55697 | 6186 |
from assms(2)[unfolded affine_dependent_explicit] |
6187 |
obtain s u where |
|
| 64267 | 6188 |
"finite s" "s \<subseteq> c" "sum u s = 0" "\<exists>v\<in>s. u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" |
| 55697 | 6189 |
by blast |
| 53347 | 6190 |
then show ?thesis |
6191 |
apply (rule_tac x="\<lambda>v. if v\<in>s then u v else 0" in exI) |
|
| 64267 | 6192 |
unfolding if_smult scaleR_zero_left and sum.inter_restrict[OF assms(1), symmetric] |
| 53347 | 6193 |
apply (auto simp add: Int_absorb1) |
6194 |
done |
|
6195 |
qed |
|
| 33175 | 6196 |
|
6197 |
lemma radon_s_lemma: |
|
| 53347 | 6198 |
assumes "finite s" |
| 64267 | 6199 |
and "sum f s = (0::real)" |
6200 |
shows "sum f {x\<in>s. 0 < f x} = - sum f {x\<in>s. f x < 0}"
|
|
| 53347 | 6201 |
proof - |
6202 |
have *: "\<And>x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x" |
|
6203 |
by auto |
|
6204 |
show ?thesis |
|
| 64267 | 6205 |
unfolding add_eq_0_iff[symmetric] and sum.inter_filter[OF assms(1)] |
6206 |
and sum.distrib[symmetric] and * |
|
| 53347 | 6207 |
using assms(2) |
|
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
|
6208 |
by assumption |
| 53347 | 6209 |
qed |
| 33175 | 6210 |
|
6211 |
lemma radon_v_lemma: |
|
| 53347 | 6212 |
assumes "finite s" |
| 64267 | 6213 |
and "sum f s = 0" |
| 53347 | 6214 |
and "\<forall>x. g x = (0::real) \<longrightarrow> f x = (0::'a::euclidean_space)" |
| 64267 | 6215 |
shows "(sum f {x\<in>s. 0 < g x}) = - sum f {x\<in>s. g x < 0}"
|
| 53347 | 6216 |
proof - |
6217 |
have *: "\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" |
|
6218 |
using assms(3) by auto |
|
6219 |
show ?thesis |
|
| 64267 | 6220 |
unfolding eq_neg_iff_add_eq_0 and sum.inter_filter[OF assms(1)] |
6221 |
and sum.distrib[symmetric] and * |
|
| 53347 | 6222 |
using assms(2) |
6223 |
apply assumption |
|
6224 |
done |
|
6225 |
qed |
|
| 33175 | 6226 |
|
6227 |
lemma radon_partition: |
|
6228 |
assumes "finite c" "affine_dependent c" |
|
| 53347 | 6229 |
shows "\<exists>m p. m \<inter> p = {} \<and> m \<union> p = c \<and> (convex hull m) \<inter> (convex hull p) \<noteq> {}"
|
6230 |
proof - |
|
| 64267 | 6231 |
obtain u v where uv: "sum u c = 0" "v\<in>c" "u v \<noteq> 0" "(\<Sum>v\<in>c. u v *\<^sub>R v) = 0" |
| 53347 | 6232 |
using radon_ex_lemma[OF assms] by auto |
6233 |
have fin: "finite {x \<in> c. 0 < u x}" "finite {x \<in> c. 0 > u x}"
|
|
6234 |
using assms(1) by auto |
|
| 64267 | 6235 |
define z where "z = inverse (sum u {x\<in>c. u x > 0}) *\<^sub>R sum (\<lambda>x. u x *\<^sub>R x) {x\<in>c. u x > 0}"
|
6236 |
have "sum u {x \<in> c. 0 < u x} \<noteq> 0"
|
|
| 53347 | 6237 |
proof (cases "u v \<ge> 0") |
6238 |
case False |
|
6239 |
then have "u v < 0" by auto |
|
6240 |
then show ?thesis |
|
6241 |
proof (cases "\<exists>w\<in>{x \<in> c. 0 < u x}. u w > 0")
|
|
6242 |
case True |
|
6243 |
then show ?thesis |
|
| 64267 | 6244 |
using sum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto |
| 33175 | 6245 |
next |
| 53347 | 6246 |
case False |
| 64267 | 6247 |
then have "sum u c \<le> sum (\<lambda>x. if x=v then u v else 0) c" |
6248 |
apply (rule_tac sum_mono) |
|
| 53347 | 6249 |
apply auto |
6250 |
done |
|
6251 |
then show ?thesis |
|
| 64267 | 6252 |
unfolding sum.delta[OF assms(1)] using uv(2) and \<open>u v < 0\<close> and uv(1) by auto |
| 53347 | 6253 |
qed |
| 64267 | 6254 |
qed (insert sum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto) |
6255 |
||
6256 |
then have *: "sum u {x\<in>c. u x > 0} > 0"
|
|
| 53347 | 6257 |
unfolding less_le |
6258 |
apply (rule_tac conjI) |
|
| 64267 | 6259 |
apply (rule_tac sum_nonneg) |
6260 |
apply auto |
|
6261 |
done |
|
6262 |
moreover have "sum u ({x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}) = sum u c"
|
|
| 33175 | 6263 |
"(\<Sum>x\<in>{x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}. u x *\<^sub>R x) = (\<Sum>x\<in>c. u x *\<^sub>R x)"
|
| 53347 | 6264 |
using assms(1) |
| 64267 | 6265 |
apply (rule_tac[!] sum.mono_neutral_left) |
6266 |
apply auto |
|
6267 |
done |
|
6268 |
then have "sum u {x \<in> c. 0 < u x} = - sum u {x \<in> c. 0 > u x}"
|
|
| 53347 | 6269 |
"(\<Sum>x\<in>{x \<in> c. 0 < u x}. u x *\<^sub>R x) = - (\<Sum>x\<in>{x \<in> c. 0 > u x}. u x *\<^sub>R x)"
|
6270 |
unfolding eq_neg_iff_add_eq_0 |
|
6271 |
using uv(1,4) |
|
| 64267 | 6272 |
by (auto simp add: sum.union_inter_neutral[OF fin, symmetric]) |
6273 |
moreover have "\<forall>x\<in>{v \<in> c. u v < 0}. 0 \<le> inverse (sum u {x \<in> c. 0 < u x}) * - u x"
|
|
| 53347 | 6274 |
apply rule |
6275 |
apply (rule mult_nonneg_nonneg) |
|
6276 |
using * |
|
6277 |
apply auto |
|
6278 |
done |
|
6279 |
ultimately have "z \<in> convex hull {v \<in> c. u v \<le> 0}"
|
|
6280 |
unfolding convex_hull_explicit mem_Collect_eq |
|
6281 |
apply (rule_tac x="{v \<in> c. u v < 0}" in exI)
|
|
| 64267 | 6282 |
apply (rule_tac x="\<lambda>y. inverse (sum u {x\<in>c. u x > 0}) * - u y" in exI)
|
6283 |
using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] and z_def |
|
6284 |
apply (auto simp add: sum_negf sum_distrib_left[symmetric]) |
|
6285 |
done |
|
6286 |
moreover have "\<forall>x\<in>{v \<in> c. 0 < u v}. 0 \<le> inverse (sum u {x \<in> c. 0 < u x}) * u x"
|
|
| 53347 | 6287 |
apply rule |
6288 |
apply (rule mult_nonneg_nonneg) |
|
6289 |
using * |
|
6290 |
apply auto |
|
6291 |
done |
|
6292 |
then have "z \<in> convex hull {v \<in> c. u v > 0}"
|
|
6293 |
unfolding convex_hull_explicit mem_Collect_eq |
|
6294 |
apply (rule_tac x="{v \<in> c. 0 < u v}" in exI)
|
|
| 64267 | 6295 |
apply (rule_tac x="\<lambda>y. inverse (sum u {x\<in>c. u x > 0}) * u y" in exI)
|
| 53347 | 6296 |
using assms(1) |
| 64267 | 6297 |
unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] and z_def |
| 53347 | 6298 |
using * |
| 64267 | 6299 |
apply (auto simp add: sum_negf sum_distrib_left[symmetric]) |
| 53347 | 6300 |
done |
6301 |
ultimately show ?thesis |
|
6302 |
apply (rule_tac x="{v\<in>c. u v \<le> 0}" in exI)
|
|
6303 |
apply (rule_tac x="{v\<in>c. u v > 0}" in exI)
|
|
6304 |
apply auto |
|
6305 |
done |
|
6306 |
qed |
|
6307 |
||
6308 |
lemma radon: |
|
6309 |
assumes "affine_dependent c" |
|
6310 |
obtains m p where "m \<subseteq> c" "p \<subseteq> c" "m \<inter> p = {}" "(convex hull m) \<inter> (convex hull p) \<noteq> {}"
|
|
6311 |
proof - |
|
| 55697 | 6312 |
from assms[unfolded affine_dependent_explicit] |
6313 |
obtain s u where |
|
| 64267 | 6314 |
"finite s" "s \<subseteq> c" "sum u s = 0" "\<exists>v\<in>s. u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" |
| 55697 | 6315 |
by blast |
| 53347 | 6316 |
then have *: "finite s" "affine_dependent s" and s: "s \<subseteq> c" |
6317 |
unfolding affine_dependent_explicit by auto |
|
| 55697 | 6318 |
from radon_partition[OF *] |
6319 |
obtain m p where "m \<inter> p = {}" "m \<union> p = s" "convex hull m \<inter> convex hull p \<noteq> {}"
|
|
6320 |
by blast |
|
| 53347 | 6321 |
then show ?thesis |
6322 |
apply (rule_tac that[of p m]) |
|
6323 |
using s |
|
6324 |
apply auto |
|
6325 |
done |
|
6326 |
qed |
|
6327 |
||
| 33175 | 6328 |
|
| 60420 | 6329 |
subsection \<open>Helly's theorem\<close> |
| 33175 | 6330 |
|
| 53347 | 6331 |
lemma helly_induct: |
6332 |
fixes f :: "'a::euclidean_space set set" |
|
6333 |
assumes "card f = n" |
|
6334 |
and "n \<ge> DIM('a) + 1"
|
|
| 60585 | 6335 |
and "\<forall>s\<in>f. convex s" "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}"
|
| 53347 | 6336 |
shows "\<Inter>f \<noteq> {}"
|
6337 |
using assms |
|
6338 |
proof (induct n arbitrary: f) |
|
6339 |
case 0 |
|
6340 |
then show ?case by auto |
|
6341 |
next |
|
6342 |
case (Suc n) |
|
6343 |
have "finite f" |
|
| 60420 | 6344 |
using \<open>card f = Suc n\<close> by (auto intro: card_ge_0_finite) |
| 53347 | 6345 |
show "\<Inter>f \<noteq> {}"
|
6346 |
apply (cases "n = DIM('a)")
|
|
6347 |
apply (rule Suc(5)[rule_format]) |
|
| 60420 | 6348 |
unfolding \<open>card f = Suc n\<close> |
| 53347 | 6349 |
proof - |
6350 |
assume ng: "n \<noteq> DIM('a)"
|
|
6351 |
then have "\<exists>X. \<forall>s\<in>f. X s \<in> \<Inter>(f - {s})"
|
|
6352 |
apply (rule_tac bchoice) |
|
6353 |
unfolding ex_in_conv |
|
6354 |
apply (rule, rule Suc(1)[rule_format]) |
|
| 60420 | 6355 |
unfolding card_Diff_singleton_if[OF \<open>finite f\<close>] \<open>card f = Suc n\<close> |
| 53347 | 6356 |
defer |
6357 |
defer |
|
6358 |
apply (rule Suc(4)[rule_format]) |
|
6359 |
defer |
|
6360 |
apply (rule Suc(5)[rule_format]) |
|
| 60420 | 6361 |
using Suc(3) \<open>finite f\<close> |
| 53347 | 6362 |
apply auto |
6363 |
done |
|
6364 |
then obtain X where X: "\<forall>s\<in>f. X s \<in> \<Inter>(f - {s})" by auto
|
|
6365 |
show ?thesis |
|
6366 |
proof (cases "inj_on X f") |
|
6367 |
case False |
|
6368 |
then obtain s t where st: "s\<noteq>t" "s\<in>f" "t\<in>f" "X s = X t" |
|
6369 |
unfolding inj_on_def by auto |
|
6370 |
then have *: "\<Inter>f = \<Inter>(f - {s}) \<inter> \<Inter>(f - {t})" by auto
|
|
6371 |
show ?thesis |
|
6372 |
unfolding * |
|
6373 |
unfolding ex_in_conv[symmetric] |
|
6374 |
apply (rule_tac x="X s" in exI) |
|
6375 |
apply rule |
|
6376 |
apply (rule X[rule_format]) |
|
6377 |
using X st |
|
6378 |
apply auto |
|
6379 |
done |
|
6380 |
next |
|
6381 |
case True |
|
6382 |
then obtain m p where mp: "m \<inter> p = {}" "m \<union> p = X ` f" "convex hull m \<inter> convex hull p \<noteq> {}"
|
|
6383 |
using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"] |
|
| 60420 | 6384 |
unfolding card_image[OF True] and \<open>card f = Suc n\<close> |
6385 |
using Suc(3) \<open>finite f\<close> and ng |
|
| 53347 | 6386 |
by auto |
6387 |
have "m \<subseteq> X ` f" "p \<subseteq> X ` f" |
|
6388 |
using mp(2) by auto |
|
6389 |
then obtain g h where gh:"m = X ` g" "p = X ` h" "g \<subseteq> f" "h \<subseteq> f" |
|
6390 |
unfolding subset_image_iff by auto |
|
6391 |
then have "f \<union> (g \<union> h) = f" by auto |
|
6392 |
then have f: "f = g \<union> h" |
|
6393 |
using inj_on_Un_image_eq_iff[of X f "g \<union> h"] and True |
|
6394 |
unfolding mp(2)[unfolded image_Un[symmetric] gh] |
|
6395 |
by auto |
|
6396 |
have *: "g \<inter> h = {}"
|
|
6397 |
using mp(1) |
|
6398 |
unfolding gh |
|
6399 |
using inj_on_image_Int[OF True gh(3,4)] |
|
6400 |
by auto |
|
6401 |
have "convex hull (X ` h) \<subseteq> \<Inter>g" "convex hull (X ` g) \<subseteq> \<Inter>h" |
|
6402 |
apply (rule_tac [!] hull_minimal) |
|
6403 |
using Suc gh(3-4) |
|
6404 |
unfolding subset_eq |
|
6405 |
apply (rule_tac [2] convex_Inter, rule_tac [4] convex_Inter) |
|
6406 |
apply rule |
|
6407 |
prefer 3 |
|
6408 |
apply rule |
|
6409 |
proof - |
|
6410 |
fix x |
|
6411 |
assume "x \<in> X ` g" |
|
| 55697 | 6412 |
then obtain y where "y \<in> g" "x = X y" |
6413 |
unfolding image_iff .. |
|
| 53347 | 6414 |
then show "x \<in> \<Inter>h" |
6415 |
using X[THEN bspec[where x=y]] using * f by auto |
|
6416 |
next |
|
6417 |
fix x |
|
6418 |
assume "x \<in> X ` h" |
|
| 55697 | 6419 |
then obtain y where "y \<in> h" "x = X y" |
6420 |
unfolding image_iff .. |
|
| 53347 | 6421 |
then show "x \<in> \<Inter>g" |
6422 |
using X[THEN bspec[where x=y]] using * f by auto |
|
6423 |
qed auto |
|
6424 |
then show ?thesis |
|
6425 |
unfolding f using mp(3)[unfolded gh] by blast |
|
6426 |
qed |
|
6427 |
qed auto |
|
6428 |
qed |
|
6429 |
||
6430 |
lemma helly: |
|
6431 |
fixes f :: "'a::euclidean_space set set" |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
6432 |
assumes "card f \<ge> DIM('a) + 1" "\<forall>s\<in>f. convex s"
|
| 60585 | 6433 |
and "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}"
|
| 53347 | 6434 |
shows "\<Inter>f \<noteq> {}"
|
6435 |
apply (rule helly_induct) |
|
6436 |
using assms |
|
6437 |
apply auto |
|
6438 |
done |
|
6439 |
||
| 33175 | 6440 |
|
| 60420 | 6441 |
subsection \<open>Epigraphs of convex functions\<close> |
| 33175 | 6442 |
|
| 53348 | 6443 |
definition "epigraph s (f :: _ \<Rightarrow> real) = {xy. fst xy \<in> s \<and> f (fst xy) \<le> snd xy}"
|
6444 |
||
6445 |
lemma mem_epigraph: "(x, y) \<in> epigraph s f \<longleftrightarrow> x \<in> s \<and> f x \<le> y" |
|
6446 |
unfolding epigraph_def by auto |
|
6447 |
||
6448 |
lemma convex_epigraph: "convex (epigraph s f) \<longleftrightarrow> convex_on s f \<and> convex s" |
|
| 36338 | 6449 |
unfolding convex_def convex_on_def |
6450 |
unfolding Ball_def split_paired_All epigraph_def |
|
6451 |
unfolding mem_Collect_eq fst_conv snd_conv fst_add snd_add fst_scaleR snd_scaleR Ball_def[symmetric] |
|
| 53348 | 6452 |
apply safe |
6453 |
defer |
|
6454 |
apply (erule_tac x=x in allE) |
|
6455 |
apply (erule_tac x="f x" in allE) |
|
6456 |
apply safe |
|
6457 |
apply (erule_tac x=xa in allE) |
|
6458 |
apply (erule_tac x="f xa" in allE) |
|
6459 |
prefer 3 |
|
6460 |
apply (rule_tac y="u * f a + v * f aa" in order_trans) |
|
6461 |
defer |
|
6462 |
apply (auto intro!:mult_left_mono add_mono) |
|
6463 |
done |
|
6464 |
||
6465 |
lemma convex_epigraphI: "convex_on s f \<Longrightarrow> convex s \<Longrightarrow> convex (epigraph s f)" |
|
6466 |
unfolding convex_epigraph by auto |
|
6467 |
||
6468 |
lemma convex_epigraph_convex: "convex s \<Longrightarrow> convex_on s f \<longleftrightarrow> convex(epigraph s f)" |
|
6469 |
by (simp add: convex_epigraph) |
|
6470 |
||
| 33175 | 6471 |
|
| 60420 | 6472 |
subsubsection \<open>Use this to derive general bound property of convex function\<close> |
| 33175 | 6473 |
|
6474 |
lemma convex_on: |
|
6475 |
assumes "convex s" |
|
| 53348 | 6476 |
shows "convex_on s f \<longleftrightarrow> |
| 64267 | 6477 |
(\<forall>k u x. (\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> x i \<in> s) \<and> sum u {1..k} = 1 \<longrightarrow>
|
6478 |
f (sum (\<lambda>i. u i *\<^sub>R x i) {1..k} ) \<le> sum (\<lambda>i. u i * f(x i)) {1..k})"
|
|
| 33175 | 6479 |
unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq |
| 64267 | 6480 |
unfolding fst_sum snd_sum fst_scaleR snd_scaleR |
| 36338 | 6481 |
apply safe |
6482 |
apply (drule_tac x=k in spec) |
|
6483 |
apply (drule_tac x=u in spec) |
|
6484 |
apply (drule_tac x="\<lambda>i. (x i, f (x i))" in spec) |
|
6485 |
apply simp |
|
| 53348 | 6486 |
using assms[unfolded convex] |
6487 |
apply simp |
|
6488 |
apply (rule_tac y="\<Sum>i = 1..k. u i * f (fst (x i))" in order_trans) |
|
6489 |
defer |
|
| 64267 | 6490 |
apply (rule sum_mono) |
| 53348 | 6491 |
apply (erule_tac x=i in allE) |
6492 |
unfolding real_scaleR_def |
|
6493 |
apply (rule mult_left_mono) |
|
6494 |
using assms[unfolded convex] |
|
6495 |
apply auto |
|
6496 |
done |
|
| 33175 | 6497 |
|
| 36338 | 6498 |
|
| 60420 | 6499 |
subsection \<open>Convexity of general and special intervals\<close> |
| 33175 | 6500 |
|
6501 |
lemma is_interval_convex: |
|
| 53348 | 6502 |
fixes s :: "'a::euclidean_space set" |
6503 |
assumes "is_interval s" |
|
6504 |
shows "convex s" |
|
|
37732
6432bf0d7191
generalize type of is_interval to class euclidean_space
huffman
parents:
37673
diff
changeset
|
6505 |
proof (rule convexI) |
| 53348 | 6506 |
fix x y and u v :: real |
6507 |
assume as: "x \<in> s" "y \<in> s" "0 \<le> u" "0 \<le> v" "u + v = 1" |
|
6508 |
then have *: "u = 1 - v" "1 - v \<ge> 0" and **: "v = 1 - u" "1 - u \<ge> 0" |
|
6509 |
by auto |
|
6510 |
{
|
|
6511 |
fix a b |
|
6512 |
assume "\<not> b \<le> u * a + v * b" |
|
6513 |
then have "u * a < (1 - v) * b" |
|
6514 |
unfolding not_le using as(4) by (auto simp add: field_simps) |
|
6515 |
then have "a < b" |
|
6516 |
unfolding * using as(4) *(2) |
|
6517 |
apply (rule_tac mult_left_less_imp_less[of "1 - v"]) |
|
6518 |
apply (auto simp add: field_simps) |
|
6519 |
done |
|
6520 |
then have "a \<le> u * a + v * b" |
|
6521 |
unfolding * using as(4) |
|
6522 |
by (auto simp add: field_simps intro!:mult_right_mono) |
|
6523 |
} |
|
6524 |
moreover |
|
6525 |
{
|
|
6526 |
fix a b |
|
6527 |
assume "\<not> u * a + v * b \<le> a" |
|
6528 |
then have "v * b > (1 - u) * a" |
|
6529 |
unfolding not_le using as(4) by (auto simp add: field_simps) |
|
6530 |
then have "a < b" |
|
6531 |
unfolding * using as(4) |
|
6532 |
apply (rule_tac mult_left_less_imp_less) |
|
6533 |
apply (auto simp add: field_simps) |
|
6534 |
done |
|
6535 |
then have "u * a + v * b \<le> b" |
|
6536 |
unfolding ** |
|
6537 |
using **(2) as(3) |
|
6538 |
by (auto simp add: field_simps intro!:mult_right_mono) |
|
6539 |
} |
|
6540 |
ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> s" |
|
6541 |
apply - |
|
6542 |
apply (rule assms[unfolded is_interval_def, rule_format, OF as(1,2)]) |
|
6543 |
using as(3-) DIM_positive[where 'a='a] |
|
6544 |
apply (auto simp: inner_simps) |
|
6545 |
done |
|
|
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
|
6546 |
qed |
| 33175 | 6547 |
|
6548 |
lemma is_interval_connected: |
|
| 53348 | 6549 |
fixes s :: "'a::euclidean_space set" |
| 33175 | 6550 |
shows "is_interval s \<Longrightarrow> connected s" |
6551 |
using is_interval_convex convex_connected by auto |
|
6552 |
||
|
62618
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
6553 |
lemma convex_box [simp]: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))" |
| 56188 | 6554 |
apply (rule_tac[!] is_interval_convex)+ |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6555 |
using is_interval_box is_interval_cbox |
| 53348 | 6556 |
apply auto |
6557 |
done |
|
| 33175 | 6558 |
|
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6559 |
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
|
6560 |
lemma connected_imp_perfect: |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6561 |
fixes a :: "'a::metric_space" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6562 |
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
|
6563 |
shows "a islimpt S" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6564 |
proof - |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6565 |
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
|
6566 |
proof - |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6567 |
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
|
6568 |
using \<open>open T\<close> \<open>a \<in> T\<close> by (auto simp: open_contains_cball) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6569 |
have "openin (subtopology euclidean S) {a}"
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6570 |
unfolding openin_open using that \<open>a \<in> S\<close> by blast |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6571 |
moreover have "closedin (subtopology euclidean S) {a}"
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6572 |
by (simp add: assms) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6573 |
ultimately show "False" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6574 |
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
|
6575 |
qed |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6576 |
then show ?thesis |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6577 |
unfolding islimpt_def by blast |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6578 |
qed |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6579 |
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6580 |
lemma connected_imp_perfect_aff_dim: |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
6581 |
"\<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
|
6582 |
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
|
6583 |
|
| 61808 | 6584 |
subsection \<open>On \<open>real\<close>, \<open>is_interval\<close>, \<open>convex\<close> and \<open>connected\<close> are all equivalent.\<close> |
| 33175 | 6585 |
|
6586 |
lemma is_interval_1: |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6587 |
"is_interval (s::real set) \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. \<forall> x. a \<le> x \<and> x \<le> b \<longrightarrow> x \<in> s)" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6588 |
unfolding is_interval_def by auto |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6589 |
|
| 54465 | 6590 |
lemma is_interval_connected_1: |
6591 |
fixes s :: "real set" |
|
6592 |
shows "is_interval s \<longleftrightarrow> connected s" |
|
6593 |
apply rule |
|
6594 |
apply (rule is_interval_connected, assumption) |
|
6595 |
unfolding is_interval_1 |
|
6596 |
apply rule |
|
6597 |
apply rule |
|
6598 |
apply rule |
|
6599 |
apply rule |
|
6600 |
apply (erule conjE) |
|
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
6601 |
apply (rule ccontr) |
| 54465 | 6602 |
proof - |
6603 |
fix a b x |
|
6604 |
assume as: "connected s" "a \<in> s" "b \<in> s" "a \<le> x" "x \<le> b" "x \<notin> s" |
|
6605 |
then have *: "a < x" "x < b" |
|
6606 |
unfolding not_le [symmetric] by auto |
|
6607 |
let ?halfl = "{..<x} "
|
|
6608 |
let ?halfr = "{x<..}"
|
|
6609 |
{
|
|
6610 |
fix y |
|
6611 |
assume "y \<in> s" |
|
| 60420 | 6612 |
with \<open>x \<notin> s\<close> have "x \<noteq> y" by auto |
| 54465 | 6613 |
then have "y \<in> ?halfr \<union> ?halfl" by auto |
6614 |
} |
|
6615 |
moreover have "a \<in> ?halfl" "b \<in> ?halfr" using * by auto |
|
6616 |
then have "?halfl \<inter> s \<noteq> {}" "?halfr \<inter> s \<noteq> {}"
|
|
6617 |
using as(2-3) by auto |
|
6618 |
ultimately show False |
|
6619 |
apply (rule_tac notE[OF as(1)[unfolded connected_def]]) |
|
6620 |
apply (rule_tac x = ?halfl in exI) |
|
6621 |
apply (rule_tac x = ?halfr in exI) |
|
6622 |
apply rule |
|
6623 |
apply (rule open_lessThan) |
|
6624 |
apply rule |
|
6625 |
apply (rule open_greaterThan) |
|
6626 |
apply auto |
|
6627 |
done |
|
6628 |
qed |
|
| 33175 | 6629 |
|
6630 |
lemma is_interval_convex_1: |
|
| 54465 | 6631 |
fixes s :: "real set" |
6632 |
shows "is_interval s \<longleftrightarrow> convex s" |
|
6633 |
by (metis is_interval_convex convex_connected is_interval_connected_1) |
|
| 33175 | 6634 |
|
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
6635 |
lemma connected_compact_interval_1: |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
6636 |
"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
|
6637 |
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
|
6638 |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6639 |
lemma connected_convex_1: |
| 54465 | 6640 |
fixes s :: "real set" |
6641 |
shows "connected s \<longleftrightarrow> convex s" |
|
6642 |
by (metis is_interval_convex convex_connected is_interval_connected_1) |
|
| 53348 | 6643 |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6644 |
lemma connected_convex_1_gen: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6645 |
fixes s :: "'a :: euclidean_space set" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6646 |
assumes "DIM('a) = 1"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6647 |
shows "connected s \<longleftrightarrow> convex s" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6648 |
proof - |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6649 |
obtain f:: "'a \<Rightarrow> real" where linf: "linear f" and "inj f" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6650 |
using subspace_isomorphism [where 'a = 'a and 'b = real] |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6651 |
by (metis DIM_real dim_UNIV subspace_UNIV assms) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6652 |
then have "f -` (f ` s) = s" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6653 |
by (simp add: inj_vimage_image_eq) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6654 |
then show ?thesis |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6655 |
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
|
6656 |
qed |
| 53348 | 6657 |
|
| 60420 | 6658 |
subsection \<open>Another intermediate value theorem formulation\<close> |
| 33175 | 6659 |
|
| 53348 | 6660 |
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
|
6661 |
fixes f :: "real \<Rightarrow> 'a::euclidean_space" |
| 53348 | 6662 |
assumes "a \<le> b" |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6663 |
and "continuous_on {a..b} f"
|
| 53348 | 6664 |
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
|
6665 |
shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
|
| 56188 | 6666 |
proof - |
6667 |
have "f a \<in> f ` cbox a b" "f b \<in> f ` cbox a b" |
|
| 53348 | 6668 |
apply (rule_tac[!] imageI) |
6669 |
using assms(1) |
|
6670 |
apply auto |
|
6671 |
done |
|
6672 |
then show ?thesis |
|
| 56188 | 6673 |
using connected_ivt_component[of "f ` cbox a b" "f a" "f b" k y] |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6674 |
by (simp add: Topology_Euclidean_Space.connected_continuous_image assms) |
| 53348 | 6675 |
qed |
6676 |
||
6677 |
lemma ivt_increasing_component_1: |
|
6678 |
fixes f :: "real \<Rightarrow> 'a::euclidean_space" |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6679 |
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
|
6680 |
f a\<bullet>k \<le> y \<Longrightarrow> y \<le> f b\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
|
| 53348 | 6681 |
by (rule ivt_increasing_component_on_1) (auto simp add: continuous_at_imp_continuous_on) |
6682 |
||
6683 |
lemma ivt_decreasing_component_on_1: |
|
6684 |
fixes f :: "real \<Rightarrow> 'a::euclidean_space" |
|
6685 |
assumes "a \<le> b" |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6686 |
and "continuous_on {a..b} f"
|
| 53348 | 6687 |
and "(f b)\<bullet>k \<le> y" |
6688 |
and "y \<le> (f a)\<bullet>k" |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6689 |
shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
|
| 53348 | 6690 |
apply (subst neg_equal_iff_equal[symmetric]) |
|
44531
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44525
diff
changeset
|
6691 |
using ivt_increasing_component_on_1[of a b "\<lambda>x. - f x" k "- y"] |
| 53348 | 6692 |
using assms using continuous_on_minus |
6693 |
apply auto |
|
6694 |
done |
|
6695 |
||
6696 |
lemma ivt_decreasing_component_1: |
|
6697 |
fixes f :: "real \<Rightarrow> 'a::euclidean_space" |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
6698 |
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
|
6699 |
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 | 6700 |
by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) |
6701 |
||
| 33175 | 6702 |
|
| 60420 | 6703 |
subsection \<open>A bound within a convex hull, and so an interval\<close> |
| 33175 | 6704 |
|
6705 |
lemma convex_on_convex_hull_bound: |
|
| 53348 | 6706 |
assumes "convex_on (convex hull s) f" |
6707 |
and "\<forall>x\<in>s. f x \<le> b" |
|
6708 |
shows "\<forall>x\<in> convex hull s. f x \<le> b" |
|
6709 |
proof |
|
6710 |
fix x |
|
6711 |
assume "x \<in> convex hull s" |
|
6712 |
then obtain k u v where |
|
| 64267 | 6713 |
obt: "\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> v i \<in> s" "sum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R v i) = x"
|
| 33175 | 6714 |
unfolding convex_hull_indexed mem_Collect_eq by auto |
| 53348 | 6715 |
have "(\<Sum>i = 1..k. u i * f (v i)) \<le> b" |
| 64267 | 6716 |
using sum_mono[of "{1..k}" "\<lambda>i. u i * f (v i)" "\<lambda>i. u i * b"]
|
6717 |
unfolding sum_distrib_right[symmetric] obt(2) mult_1 |
|
| 53348 | 6718 |
apply (drule_tac meta_mp) |
6719 |
apply (rule mult_left_mono) |
|
6720 |
using assms(2) obt(1) |
|
6721 |
apply auto |
|
6722 |
done |
|
6723 |
then show "f x \<le> b" |
|
6724 |
using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v] |
|
6725 |
unfolding obt(2-3) |
|
6726 |
using obt(1) and hull_subset[unfolded subset_eq, rule_format, of _ s] |
|
6727 |
by auto |
|
6728 |
qed |
|
6729 |
||
| 64267 | 6730 |
lemma inner_sum_Basis[simp]: "i \<in> Basis \<Longrightarrow> (\<Sum>Basis) \<bullet> i = 1" |
6731 |
by (simp add: inner_sum_left sum.If_cases inner_Basis) |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6732 |
|
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6733 |
lemma convex_set_plus: |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6734 |
assumes "convex s" and "convex t" shows "convex (s + t)" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6735 |
proof - |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6736 |
have "convex {x + y |x y. x \<in> s \<and> y \<in> t}"
|
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6737 |
using assms by (rule convex_sums) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6738 |
moreover have "{x + y |x y. x \<in> s \<and> y \<in> t} = s + t"
|
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6739 |
unfolding set_plus_def by auto |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6740 |
finally show "convex (s + t)" . |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6741 |
qed |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6742 |
|
| 64267 | 6743 |
lemma convex_set_sum: |
|
55929
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents:
55928
diff
changeset
|
6744 |
assumes "\<And>i. i \<in> A \<Longrightarrow> convex (B i)" |
|
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents:
55928
diff
changeset
|
6745 |
shows "convex (\<Sum>i\<in>A. B i)" |
|
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents:
55928
diff
changeset
|
6746 |
proof (cases "finite A") |
|
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents:
55928
diff
changeset
|
6747 |
case True then show ?thesis using assms |
|
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents:
55928
diff
changeset
|
6748 |
by induct (auto simp: convex_set_plus) |
|
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents:
55928
diff
changeset
|
6749 |
qed auto |
|
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents:
55928
diff
changeset
|
6750 |
|
| 64267 | 6751 |
lemma finite_set_sum: |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6752 |
assumes "finite A" and "\<forall>i\<in>A. finite (B i)" shows "finite (\<Sum>i\<in>A. B i)" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6753 |
using assms by (induct set: finite, simp, simp add: finite_set_plus) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6754 |
|
| 64267 | 6755 |
lemma set_sum_eq: |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6756 |
"finite A \<Longrightarrow> (\<Sum>i\<in>A. B i) = {\<Sum>i\<in>A. f i |f. \<forall>i\<in>A. f i \<in> B i}"
|
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6757 |
apply (induct set: finite) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6758 |
apply simp |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6759 |
apply simp |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6760 |
apply (safe elim!: set_plus_elim) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6761 |
apply (rule_tac x="fun_upd f x a" in exI) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6762 |
apply simp |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6763 |
apply (rule_tac f="\<lambda>x. a + x" in arg_cong) |
| 64267 | 6764 |
apply (rule sum.cong [OF refl]) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6765 |
apply clarsimp |
| 57865 | 6766 |
apply fast |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6767 |
done |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6768 |
|
| 64267 | 6769 |
lemma box_eq_set_sum_Basis: |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6770 |
shows "{x. \<forall>i\<in>Basis. x\<bullet>i \<in> B i} = (\<Sum>i\<in>Basis. image (\<lambda>x. x *\<^sub>R i) (B i))"
|
| 64267 | 6771 |
apply (subst set_sum_eq [OF finite_Basis]) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6772 |
apply safe |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6773 |
apply (fast intro: euclidean_representation [symmetric]) |
| 64267 | 6774 |
apply (subst inner_sum_left) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6775 |
apply (subgoal_tac "(\<Sum>x\<in>Basis. f x \<bullet> i) = f i \<bullet> i") |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6776 |
apply (drule (1) bspec) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6777 |
apply clarsimp |
| 64267 | 6778 |
apply (frule sum.remove [OF finite_Basis]) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6779 |
apply (erule trans) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6780 |
apply simp |
| 64267 | 6781 |
apply (rule sum.neutral) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6782 |
apply clarsimp |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6783 |
apply (frule_tac x=i in bspec, assumption) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6784 |
apply (drule_tac x=x in bspec, assumption) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6785 |
apply clarsimp |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6786 |
apply (cut_tac u=x and v=i in inner_Basis, assumption+) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6787 |
apply (rule ccontr) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6788 |
apply simp |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6789 |
done |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6790 |
|
| 64267 | 6791 |
lemma convex_hull_set_sum: |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6792 |
"convex hull (\<Sum>i\<in>A. B i) = (\<Sum>i\<in>A. convex hull (B i))" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6793 |
proof (cases "finite A") |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6794 |
assume "finite A" then show ?thesis |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6795 |
by (induct set: finite, simp, simp add: convex_hull_set_plus) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6796 |
qed simp |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6797 |
|
| 56188 | 6798 |
lemma convex_hull_eq_real_cbox: |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6799 |
fixes x y :: real assumes "x \<le> y" |
| 56188 | 6800 |
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
|
6801 |
proof (rule hull_unique) |
| 60420 | 6802 |
show "{x, y} \<subseteq> cbox x y" using \<open>x \<le> y\<close> by auto
|
| 56188 | 6803 |
show "convex (cbox x y)" |
6804 |
by (rule convex_box) |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6805 |
next |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6806 |
fix s assume "{x, y} \<subseteq> s" and "convex s"
|
| 56188 | 6807 |
then show "cbox x y \<subseteq> s" |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6808 |
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
|
6809 |
by - (clarify, simp (no_asm_use), fast) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6810 |
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
|
6811 |
|
| 33175 | 6812 |
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
|
6813 |
"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
|
6814 |
(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
|
6815 |
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
|
6816 |
have One[simp]: "\<And>i. i \<in> Basis \<Longrightarrow> One \<bullet> i = 1" |
| 64267 | 6817 |
by (simp add: inner_sum_left sum.If_cases inner_Basis) |
| 56188 | 6818 |
have "?int = {x. \<forall>i\<in>Basis. x \<bullet> i \<in> cbox 0 1}"
|
6819 |
by (auto simp: cbox_def) |
|
6820 |
also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` cbox 0 1)" |
|
| 64267 | 6821 |
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
|
6822 |
also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` (convex hull {0, 1}))"
|
| 56188 | 6823 |
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
|
6824 |
also have "\<dots> = (\<Sum>i\<in>Basis. convex hull ((\<lambda>x. x *\<^sub>R i) ` {0, 1}))"
|
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6825 |
by (simp only: convex_hull_linear_image linear_scaleR_left) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6826 |
also have "\<dots> = convex hull (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` {0, 1})"
|
| 64267 | 6827 |
by (simp only: convex_hull_set_sum) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6828 |
also have "\<dots> = convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}}"
|
| 64267 | 6829 |
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
|
6830 |
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
|
6831 |
by simp |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
6832 |
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
|
6833 |
qed |
| 33175 | 6834 |
|
| 60420 | 6835 |
text \<open>And this is a finite set of vertices.\<close> |
| 33175 | 6836 |
|
|
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
|
6837 |
lemma unit_cube_convex_hull: |
| 56188 | 6838 |
obtains s :: "'a::euclidean_space set" |
6839 |
where "finite s" and "cbox 0 (\<Sum>Basis) = convex hull s" |
|
| 53348 | 6840 |
apply (rule that[of "{x::'a. \<forall>i\<in>Basis. x\<bullet>i=0 \<or> x\<bullet>i=1}"])
|
6841 |
apply (rule finite_subset[of _ "(\<lambda>s. (\<Sum>i\<in>Basis. (if i\<in>s then 1 else 0) *\<^sub>R i)::'a) ` Pow Basis"]) |
|
6842 |
prefer 3 |
|
6843 |
apply (rule unit_interval_convex_hull) |
|
6844 |
apply rule |
|
6845 |
unfolding mem_Collect_eq |
|
6846 |
proof - |
|
6847 |
fix x :: 'a |
|
6848 |
assume as: "\<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1" |
|
|
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
|
6849 |
show "x \<in> (\<lambda>s. \<Sum>i\<in>Basis. (if i\<in>s then 1 else 0) *\<^sub>R i) ` Pow Basis" |
| 53348 | 6850 |
apply (rule image_eqI[where x="{i. i\<in>Basis \<and> x\<bullet>i = 1}"])
|
6851 |
using as |
|
6852 |
apply (subst euclidean_eq_iff) |
|
| 57865 | 6853 |
apply auto |
| 53348 | 6854 |
done |
|
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
|
6855 |
qed auto |
| 33175 | 6856 |
|
| 60420 | 6857 |
text \<open>Hence any cube (could do any nonempty interval).\<close> |
| 33175 | 6858 |
|
6859 |
lemma cube_convex_hull: |
|
| 53348 | 6860 |
assumes "d > 0" |
| 56188 | 6861 |
obtains s :: "'a::euclidean_space set" where |
6862 |
"finite s" and "cbox (x - (\<Sum>i\<in>Basis. d*\<^sub>Ri)) (x + (\<Sum>i\<in>Basis. d*\<^sub>Ri)) = convex hull s" |
|
| 53348 | 6863 |
proof - |
|
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
|
6864 |
let ?d = "(\<Sum>i\<in>Basis. d*\<^sub>Ri)::'a" |
| 56188 | 6865 |
have *: "cbox (x - ?d) (x + ?d) = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\<Sum>Basis)" |
| 53348 | 6866 |
apply (rule set_eqI, rule) |
6867 |
unfolding image_iff |
|
6868 |
defer |
|
6869 |
apply (erule bexE) |
|
6870 |
proof - |
|
6871 |
fix y |
|
| 56188 | 6872 |
assume as: "y\<in>cbox (x - ?d) (x + ?d)" |
6873 |
then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> cbox 0 (\<Sum>Basis)" |
|
|
58776
95e58e04e534
use NO_MATCH-simproc for distribution rules in field_simps, otherwise field_simps on '(a / (c + d)) * (e + f)' can be non-terminating
hoelzl
parents:
57865
diff
changeset
|
6874 |
using assms by (simp add: mem_box field_simps inner_simps) |
| 60420 | 6875 |
with \<open>0 < d\<close> show "\<exists>z\<in>cbox 0 (\<Sum>Basis). y = x - ?d + (2 * d) *\<^sub>R z" |
|
58776
95e58e04e534
use NO_MATCH-simproc for distribution rules in field_simps, otherwise field_simps on '(a / (c + d)) * (e + f)' can be non-terminating
hoelzl
parents:
57865
diff
changeset
|
6876 |
by (intro bexI[of _ "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) auto |
| 33175 | 6877 |
next |
| 53348 | 6878 |
fix y z |
| 56188 | 6879 |
assume as: "z\<in>cbox 0 (\<Sum>Basis)" "y = x - ?d + (2*d) *\<^sub>R z" |
|
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
|
6880 |
have "\<And>i. i\<in>Basis \<Longrightarrow> 0 \<le> d * (z \<bullet> i) \<and> d * (z \<bullet> i) \<le> d" |
| 56188 | 6881 |
using assms as(1)[unfolded mem_box] |
| 53348 | 6882 |
apply (erule_tac x=i in ballE) |
6883 |
apply rule |
|
| 56536 | 6884 |
prefer 2 |
| 53348 | 6885 |
apply (rule mult_right_le_one_le) |
6886 |
using assms |
|
6887 |
apply auto |
|
6888 |
done |
|
| 56188 | 6889 |
then show "y \<in> cbox (x - ?d) (x + ?d)" |
6890 |
unfolding as(2) mem_box |
|
| 53348 | 6891 |
apply - |
6892 |
apply rule |
|
| 56188 | 6893 |
using as(1)[unfolded mem_box] |
| 53348 | 6894 |
apply (erule_tac x=i in ballE) |
6895 |
using assms |
|
6896 |
apply (auto simp: inner_simps) |
|
6897 |
done |
|
6898 |
qed |
|
| 56188 | 6899 |
obtain s where "finite s" "cbox 0 (\<Sum>Basis::'a) = convex hull s" |
| 53348 | 6900 |
using unit_cube_convex_hull by auto |
6901 |
then show ?thesis |
|
6902 |
apply (rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` s"]) |
|
6903 |
unfolding * and convex_hull_affinity |
|
6904 |
apply auto |
|
6905 |
done |
|
6906 |
qed |
|
6907 |
||
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6908 |
subsubsection\<open>Representation of any interval as a finite convex hull\<close> |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6909 |
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6910 |
lemma image_stretch_interval: |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6911 |
"(\<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
|
6912 |
(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
|
6913 |
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
|
6914 |
(\<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
|
6915 |
proof cases |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6916 |
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
|
6917 |
show ?thesis |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6918 |
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
|
6919 |
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
|
6920 |
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
|
6921 |
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
|
6922 |
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
|
6923 |
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
|
6924 |
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
|
6925 |
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
|
6926 |
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
|
6927 |
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
|
6928 |
case True |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6929 |
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
|
6930 |
next |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6931 |
case False |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6932 |
then have *: "\<And>a b. a = m i * b \<longleftrightarrow> b = a / m i" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6933 |
by (auto simp add: field_simps) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6934 |
from False have |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6935 |
"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
|
6936 |
"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
|
6937 |
using a_le_b by (auto simp: min_def max_def mult_le_cancel_left) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6938 |
with False show ?thesis using a_le_b |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6939 |
unfolding * by (auto simp add: le_divide_eq divide_le_eq ac_simps) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6940 |
qed |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6941 |
qed |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6942 |
qed simp |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6943 |
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6944 |
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
|
6945 |
"\<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
|
6946 |
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
|
6947 |
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6948 |
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
|
6949 |
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
|
6950 |
using box_ne_empty [of "a+c" "b+c"] box_ne_empty [of a b] |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6951 |
by (auto simp add: inner_left_distrib add.commute) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6952 |
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6953 |
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
|
6954 |
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
|
6955 |
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
|
6956 |
shows "cbox a b = |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6957 |
op + a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` cbox 0 One" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6958 |
using assms |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6959 |
apply (simp add: box_ne_empty image_stretch_interval cbox_translation [symmetric]) |
| 64267 | 6960 |
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
|
6961 |
done |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6962 |
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6963 |
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
|
6964 |
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
|
6965 |
obtains s where "finite s" "cbox a b = convex hull s" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6966 |
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
|
6967 |
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
|
6968 |
by blast |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6969 |
next |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6970 |
case False |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6971 |
obtain s::"'a set" where "finite s" and eq: "cbox 0 One = convex hull s" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6972 |
by (blast intro: unit_cube_convex_hull) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6973 |
have lin: "linear (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k)" |
| 64267 | 6974 |
by (rule linear_compose_sum) (auto simp: algebra_simps linearI) |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6975 |
have "finite (op + a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` s)" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6976 |
by (rule finite_imageI \<open>finite s\<close>)+ |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6977 |
then show ?thesis |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6978 |
apply (rule that) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6979 |
apply (simp add: convex_hull_translation convex_hull_linear_image [OF lin, symmetric]) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6980 |
apply (simp add: eq [symmetric] cbox_image_unit_interval [OF False]) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6981 |
done |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6982 |
qed |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
6983 |
|
| 33175 | 6984 |
|
| 60420 | 6985 |
subsection \<open>Bounded convex function on open set is continuous\<close> |
| 33175 | 6986 |
|
6987 |
lemma convex_on_bounded_continuous: |
|
| 36338 | 6988 |
fixes s :: "('a::real_normed_vector) set"
|
| 53348 | 6989 |
assumes "open s" |
6990 |
and "convex_on s f" |
|
| 61945 | 6991 |
and "\<forall>x\<in>s. \<bar>f x\<bar> \<le> b" |
| 33175 | 6992 |
shows "continuous_on s f" |
| 53348 | 6993 |
apply (rule continuous_at_imp_continuous_on) |
6994 |
unfolding continuous_at_real_range |
|
6995 |
proof (rule,rule,rule) |
|
6996 |
fix x and e :: real |
|
6997 |
assume "x \<in> s" "e > 0" |
|
| 63040 | 6998 |
define B where "B = \<bar>b\<bar> + 1" |
| 61945 | 6999 |
have B: "0 < B" "\<And>x. x\<in>s \<Longrightarrow> \<bar>f x\<bar> \<le> B" |
| 53348 | 7000 |
unfolding B_def |
7001 |
defer |
|
7002 |
apply (drule assms(3)[rule_format]) |
|
7003 |
apply auto |
|
7004 |
done |
|
7005 |
obtain k where "k > 0" and k: "cball x k \<subseteq> s" |
|
7006 |
using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]] |
|
| 60420 | 7007 |
using \<open>x\<in>s\<close> by auto |
| 33175 | 7008 |
show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e" |
| 53348 | 7009 |
apply (rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI) |
7010 |
apply rule |
|
7011 |
defer |
|
7012 |
proof (rule, rule) |
|
7013 |
fix y |
|
7014 |
assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)" |
|
7015 |
show "\<bar>f y - f x\<bar> < e" |
|
7016 |
proof (cases "y = x") |
|
7017 |
case False |
|
| 63040 | 7018 |
define t where "t = k / norm (y - x)" |
| 53348 | 7019 |
have "2 < t" "0<t" |
| 60420 | 7020 |
unfolding t_def using as False and \<open>k>0\<close> |
| 53348 | 7021 |
by (auto simp add:field_simps) |
7022 |
have "y \<in> s" |
|
7023 |
apply (rule k[unfolded subset_eq,rule_format]) |
|
7024 |
unfolding mem_cball dist_norm |
|
7025 |
apply (rule order_trans[of _ "2 * norm (x - y)"]) |
|
7026 |
using as |
|
7027 |
by (auto simp add: field_simps norm_minus_commute) |
|
7028 |
{
|
|
| 63040 | 7029 |
define w where "w = x + t *\<^sub>R (y - x)" |
| 53348 | 7030 |
have "w \<in> s" |
7031 |
unfolding w_def |
|
7032 |
apply (rule k[unfolded subset_eq,rule_format]) |
|
7033 |
unfolding mem_cball dist_norm |
|
7034 |
unfolding t_def |
|
| 60420 | 7035 |
using \<open>k>0\<close> |
| 53348 | 7036 |
apply auto |
7037 |
done |
|
7038 |
have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" |
|
7039 |
by (auto simp add: algebra_simps) |
|
7040 |
also have "\<dots> = 0" |
|
| 60420 | 7041 |
using \<open>t > 0\<close> by (auto simp add:field_simps) |
| 53348 | 7042 |
finally have w: "(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" |
| 60420 | 7043 |
unfolding w_def using False and \<open>t > 0\<close> |
| 53348 | 7044 |
by (auto simp add: algebra_simps) |
7045 |
have "2 * B < e * t" |
|
| 60420 | 7046 |
unfolding t_def using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False |
| 53348 | 7047 |
by (auto simp add:field_simps) |
7048 |
then have "(f w - f x) / t < e" |
|
| 60420 | 7049 |
using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>x\<in>s\<close>] |
7050 |
using \<open>t > 0\<close> by (auto simp add:field_simps) |
|
| 53348 | 7051 |
then have th1: "f y - f x < e" |
7052 |
apply - |
|
7053 |
apply (rule le_less_trans) |
|
7054 |
defer |
|
7055 |
apply assumption |
|
| 33175 | 7056 |
using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w] |
| 60420 | 7057 |
using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>x \<in> s\<close> \<open>w \<in> s\<close> |
| 53348 | 7058 |
by (auto simp add:field_simps) |
7059 |
} |
|
| 49531 | 7060 |
moreover |
| 53348 | 7061 |
{
|
| 63040 | 7062 |
define w where "w = x - t *\<^sub>R (y - x)" |
| 53348 | 7063 |
have "w \<in> s" |
7064 |
unfolding w_def |
|
7065 |
apply (rule k[unfolded subset_eq,rule_format]) |
|
7066 |
unfolding mem_cball dist_norm |
|
7067 |
unfolding t_def |
|
| 60420 | 7068 |
using \<open>k > 0\<close> |
| 53348 | 7069 |
apply auto |
7070 |
done |
|
7071 |
have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" |
|
7072 |
by (auto simp add: algebra_simps) |
|
7073 |
also have "\<dots> = x" |
|
| 60420 | 7074 |
using \<open>t > 0\<close> by (auto simp add:field_simps) |
| 53348 | 7075 |
finally have w: "(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" |
| 60420 | 7076 |
unfolding w_def using False and \<open>t > 0\<close> |
| 53348 | 7077 |
by (auto simp add: algebra_simps) |
7078 |
have "2 * B < e * t" |
|
7079 |
unfolding t_def |
|
| 60420 | 7080 |
using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False |
| 53348 | 7081 |
by (auto simp add:field_simps) |
7082 |
then have *: "(f w - f y) / t < e" |
|
| 60420 | 7083 |
using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>y\<in>s\<close>] |
7084 |
using \<open>t > 0\<close> |
|
| 53348 | 7085 |
by (auto simp add:field_simps) |
| 49531 | 7086 |
have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y" |
| 33175 | 7087 |
using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w] |
| 60420 | 7088 |
using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>y \<in> s\<close> \<open>w \<in> s\<close> |
| 53348 | 7089 |
by (auto simp add:field_simps) |
7090 |
also have "\<dots> = (f w + t * f y) / (1 + t)" |
|
| 60420 | 7091 |
using \<open>t > 0\<close> by (auto simp add: divide_simps) |
| 53348 | 7092 |
also have "\<dots> < e + f y" |
| 60420 | 7093 |
using \<open>t > 0\<close> * \<open>e > 0\<close> by (auto simp add: field_simps) |
| 53348 | 7094 |
finally have "f x - f y < e" by auto |
7095 |
} |
|
| 49531 | 7096 |
ultimately show ?thesis by auto |
| 60420 | 7097 |
qed (insert \<open>0<e\<close>, auto) |
7098 |
qed (insert \<open>0<e\<close> \<open>0<k\<close> \<open>0<B\<close>, auto simp: field_simps) |
|
7099 |
qed |
|
7100 |
||
7101 |
||
7102 |
subsection \<open>Upper bound on a ball implies upper and lower bounds\<close> |
|
| 33175 | 7103 |
|
7104 |
lemma convex_bounds_lemma: |
|
| 36338 | 7105 |
fixes x :: "'a::real_normed_vector" |
| 53348 | 7106 |
assumes "convex_on (cball x e) f" |
7107 |
and "\<forall>y \<in> cball x e. f y \<le> b" |
|
| 61945 | 7108 |
shows "\<forall>y \<in> cball x e. \<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" |
| 53348 | 7109 |
apply rule |
7110 |
proof (cases "0 \<le> e") |
|
7111 |
case True |
|
7112 |
fix y |
|
7113 |
assume y: "y \<in> cball x e" |
|
| 63040 | 7114 |
define z where "z = 2 *\<^sub>R x - y" |
| 53348 | 7115 |
have *: "x - (2 *\<^sub>R x - y) = y - x" |
7116 |
by (simp add: scaleR_2) |
|
7117 |
have z: "z \<in> cball x e" |
|
7118 |
using y unfolding z_def mem_cball dist_norm * by (auto simp add: norm_minus_commute) |
|
7119 |
have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" |
|
7120 |
unfolding z_def by (auto simp add: algebra_simps) |
|
7121 |
then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" |
|
7122 |
using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"] |
|
7123 |
using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z] |
|
7124 |
by (auto simp add:field_simps) |
|
7125 |
next |
|
7126 |
case False |
|
7127 |
fix y |
|
7128 |
assume "y \<in> cball x e" |
|
7129 |
then have "dist x y < 0" |
|
7130 |
using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero) |
|
7131 |
then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" |
|
7132 |
using zero_le_dist[of x y] by auto |
|
7133 |
qed |
|
7134 |
||
| 33175 | 7135 |
|
| 60420 | 7136 |
subsubsection \<open>Hence a convex function on an open set is continuous\<close> |
| 33175 | 7137 |
|
|
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
|
7138 |
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
|
7139 |
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
|
7140 |
|
| 33175 | 7141 |
lemma convex_on_continuous: |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7142 |
assumes "open (s::('a::euclidean_space) set)" "convex_on s f"
|
| 33175 | 7143 |
shows "continuous_on s f" |
| 53348 | 7144 |
unfolding continuous_on_eq_continuous_at[OF assms(1)] |
7145 |
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
|
7146 |
note dimge1 = DIM_positive[where 'a='a] |
| 53348 | 7147 |
fix x |
7148 |
assume "x \<in> s" |
|
7149 |
then obtain e where e: "cball x e \<subseteq> s" "e > 0" |
|
7150 |
using assms(1) unfolding open_contains_cball by auto |
|
| 63040 | 7151 |
define d where "d = e / real DIM('a)"
|
| 53348 | 7152 |
have "0 < d" |
| 60420 | 7153 |
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
|
7154 |
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
|
7155 |
obtain c |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7156 |
where c: "finite c" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7157 |
and c1: "convex hull c \<subseteq> cball x e" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7158 |
and c2: "cball x d \<subseteq> convex hull c" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7159 |
proof |
| 63040 | 7160 |
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
|
7161 |
show "finite c" |
| 64267 | 7162 |
unfolding c_def by (simp add: finite_set_sum) |
| 56188 | 7163 |
have 1: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cbox (x \<bullet> i - d) (x \<bullet> i + d)}"
|
| 64267 | 7164 |
unfolding box_eq_set_sum_Basis |
7165 |
unfolding c_def convex_hull_set_sum |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7166 |
apply (subst convex_hull_linear_image [symmetric]) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7167 |
apply (simp add: linear_iff scaleR_add_left) |
| 64267 | 7168 |
apply (rule sum.cong [OF refl]) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7169 |
apply (rule image_cong [OF _ refl]) |
| 56188 | 7170 |
apply (rule convex_hull_eq_real_cbox) |
| 60420 | 7171 |
apply (cut_tac \<open>0 < d\<close>, simp) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7172 |
done |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7173 |
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
|
7174 |
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
|
7175 |
show "cball x d \<subseteq> convex hull c" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7176 |
unfolding 2 |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7177 |
apply clarsimp |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7178 |
apply (simp only: dist_norm) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7179 |
apply (subst inner_diff_left [symmetric]) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7180 |
apply simp |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7181 |
apply (erule (1) order_trans [OF Basis_le_norm]) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7182 |
done |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7183 |
have e': "e = (\<Sum>(i::'a)\<in>Basis. d)" |
|
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
|
7184 |
by (simp add: d_def DIM_positive) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7185 |
show "convex hull c \<subseteq> cball x e" |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7186 |
unfolding 2 |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7187 |
apply clarsimp |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7188 |
apply (subst euclidean_dist_l2) |
| 64267 | 7189 |
apply (rule order_trans [OF setL2_le_sum]) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7190 |
apply (rule zero_le_dist) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7191 |
unfolding e' |
| 64267 | 7192 |
apply (rule sum_mono) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7193 |
apply simp |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7194 |
done |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7195 |
qed |
| 63040 | 7196 |
define k where "k = Max (f ` c)" |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7197 |
have "convex_on (convex hull c) f" |
|
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
|
7198 |
apply(rule convex_on_subset[OF assms(2)]) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
7199 |
apply(rule subset_trans[OF _ e(1)]) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7200 |
apply(rule c1) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7201 |
done |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7202 |
then have k: "\<forall>y\<in>convex hull c. f y \<le> k" |
| 53348 | 7203 |
apply (rule_tac convex_on_convex_hull_bound) |
7204 |
apply assumption |
|
7205 |
unfolding k_def |
|
7206 |
apply (rule, rule Max_ge) |
|
7207 |
using c(1) |
|
7208 |
apply auto |
|
7209 |
done |
|
|
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
|
7210 |
have "d \<le> e" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
7211 |
unfolding d_def |
| 53348 | 7212 |
apply (rule mult_imp_div_pos_le) |
| 60420 | 7213 |
using \<open>e > 0\<close> |
|
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
|
7214 |
unfolding mult_le_cancel_left1 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
7215 |
apply (auto simp: real_of_nat_ge_one_iff Suc_le_eq DIM_positive) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
7216 |
done |
| 53348 | 7217 |
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
|
7218 |
by (rule subset_cball) |
| 53348 | 7219 |
have conv: "convex_on (cball x d) f" |
7220 |
apply (rule convex_on_subset) |
|
7221 |
apply (rule convex_on_subset[OF assms(2)]) |
|
7222 |
apply (rule e(1)) |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7223 |
apply (rule dsube) |
| 53348 | 7224 |
done |
| 61945 | 7225 |
then have "\<forall>y\<in>cball x d. \<bar>f y\<bar> \<le> k + 2 * \<bar>f x\<bar>" |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7226 |
apply (rule convex_bounds_lemma) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7227 |
apply (rule ballI) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7228 |
apply (rule k [rule_format]) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7229 |
apply (erule rev_subsetD) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7230 |
apply (rule c2) |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
7231 |
done |
| 53348 | 7232 |
then have "continuous_on (ball x d) f" |
7233 |
apply (rule_tac convex_on_bounded_continuous) |
|
7234 |
apply (rule open_ball, rule convex_on_subset[OF conv]) |
|
7235 |
apply (rule ball_subset_cball) |
|
| 33270 | 7236 |
apply force |
7237 |
done |
|
| 53348 | 7238 |
then show "continuous (at x) f" |
7239 |
unfolding continuous_on_eq_continuous_at[OF open_ball] |
|
| 60420 | 7240 |
using \<open>d > 0\<close> by auto |
7241 |
qed |
|
7242 |
||
7243 |
||
7244 |
subsection \<open>Line segments, Starlike Sets, etc.\<close> |
|
| 33270 | 7245 |
|
| 53348 | 7246 |
definition midpoint :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a" |
7247 |
where "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)" |
|
7248 |
||
7249 |
definition closed_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set" |
|
7250 |
where "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \<le> u \<and> u \<le> 1}"
|
|
7251 |
||
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7252 |
definition open_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set" where |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7253 |
"open_segment a b \<equiv> closed_segment a b - {a,b}"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7254 |
|
|
62626
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
7255 |
lemmas segment = open_segment_def closed_segment_def |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
7256 |
|
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
7257 |
lemma in_segment: |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
7258 |
"x \<in> closed_segment a b \<longleftrightarrow> (\<exists>u. 0 \<le> u \<and> u \<le> 1 \<and> x = (1 - u) *\<^sub>R a + u *\<^sub>R b)" |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
7259 |
"x \<in> open_segment a b \<longleftrightarrow> a \<noteq> b \<and> (\<exists>u. 0 < u \<and> u < 1 \<and> x = (1 - u) *\<^sub>R a + u *\<^sub>R b)" |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
7260 |
using less_eq_real_def by (auto simp: segment algebra_simps) |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
7261 |
|
| 63077 | 7262 |
lemma closed_segment_linear_image: |
7263 |
"linear f \<Longrightarrow> closed_segment (f a) (f b) = f ` (closed_segment a b)" |
|
7264 |
by (force simp add: in_segment linear_add_cmul) |
|
7265 |
||
7266 |
lemma open_segment_linear_image: |
|
7267 |
"\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> open_segment (f a) (f b) = f ` (open_segment a b)" |
|
7268 |
by (force simp: open_segment_def closed_segment_linear_image inj_on_def) |
|
7269 |
||
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7270 |
lemma closed_segment_translation: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7271 |
"closed_segment (c + a) (c + b) = image (\<lambda>x. c + x) (closed_segment a b)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7272 |
apply safe |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7273 |
apply (rule_tac x="x-c" in image_eqI) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7274 |
apply (auto simp: in_segment algebra_simps) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7275 |
done |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7276 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7277 |
lemma open_segment_translation: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7278 |
"open_segment (c + a) (c + b) = image (\<lambda>x. c + x) (open_segment a b)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7279 |
by (simp add: open_segment_def closed_segment_translation translation_diff) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7280 |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7281 |
lemma open_segment_PairD: |
| 63077 | 7282 |
"(x, x') \<in> open_segment (a, a') (b, b') |
7283 |
\<Longrightarrow> (x \<in> open_segment a b \<or> a = b) \<and> (x' \<in> open_segment a' b' \<or> a' = b')" |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7284 |
by (auto simp: in_segment) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7285 |
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7286 |
lemma closed_segment_PairD: |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7287 |
"(x, x') \<in> closed_segment (a, a') (b, b') \<Longrightarrow> x \<in> closed_segment a b \<and> x' \<in> closed_segment a' b'" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7288 |
by (auto simp: closed_segment_def) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7289 |
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7290 |
lemma closed_segment_translation_eq [simp]: |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7291 |
"d + x \<in> closed_segment (d + a) (d + b) \<longleftrightarrow> x \<in> closed_segment a b" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7292 |
proof - |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7293 |
have *: "\<And>d x a b. x \<in> closed_segment a b \<Longrightarrow> d + x \<in> closed_segment (d + a) (d + b)" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7294 |
apply (simp add: closed_segment_def) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7295 |
apply (erule ex_forward) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7296 |
apply (simp add: algebra_simps) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7297 |
done |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7298 |
show ?thesis |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7299 |
using * [where d = "-d"] * |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7300 |
by (fastforce simp add:) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7301 |
qed |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7302 |
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7303 |
lemma open_segment_translation_eq [simp]: |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7304 |
"d + x \<in> open_segment (d + a) (d + b) \<longleftrightarrow> x \<in> open_segment a b" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7305 |
by (simp add: open_segment_def) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
7306 |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7307 |
lemma midpoint_idem [simp]: "midpoint x x = x" |
| 53348 | 7308 |
unfolding midpoint_def |
7309 |
unfolding scaleR_right_distrib |
|
7310 |
unfolding scaleR_left_distrib[symmetric] |
|
7311 |
by auto |
|
7312 |
||
7313 |
lemma midpoint_sym: "midpoint a b = midpoint b a" |
|
7314 |
unfolding midpoint_def by (auto simp add: scaleR_right_distrib) |
|
| 33175 | 7315 |
|
| 36338 | 7316 |
lemma midpoint_eq_iff: "midpoint a b = c \<longleftrightarrow> a + b = c + c" |
7317 |
proof - |
|
7318 |
have "midpoint a b = c \<longleftrightarrow> scaleR 2 (midpoint a b) = scaleR 2 c" |
|
7319 |
by simp |
|
| 53348 | 7320 |
then show ?thesis |
| 36338 | 7321 |
unfolding midpoint_def scaleR_2 [symmetric] by simp |
7322 |
qed |
|
7323 |
||
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7324 |
lemma |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7325 |
fixes a::real |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7326 |
assumes "a \<le> b" shows ge_midpoint_1: "a \<le> midpoint a b" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7327 |
and le_midpoint_1: "midpoint a b \<le> b" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7328 |
by (simp_all add: midpoint_def assms) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7329 |
|
| 33175 | 7330 |
lemma dist_midpoint: |
| 36338 | 7331 |
fixes a b :: "'a::real_normed_vector" shows |
| 33175 | 7332 |
"dist a (midpoint a b) = (dist a b) / 2" (is ?t1) |
7333 |
"dist b (midpoint a b) = (dist a b) / 2" (is ?t2) |
|
7334 |
"dist (midpoint a b) a = (dist a b) / 2" (is ?t3) |
|
7335 |
"dist (midpoint a b) b = (dist a b) / 2" (is ?t4) |
|
| 53348 | 7336 |
proof - |
7337 |
have *: "\<And>x y::'a. 2 *\<^sub>R x = - y \<Longrightarrow> norm x = (norm y) / 2" |
|
7338 |
unfolding equation_minus_iff by auto |
|
7339 |
have **: "\<And>x y::'a. 2 *\<^sub>R x = y \<Longrightarrow> norm x = (norm y) / 2" |
|
7340 |
by auto |
|
| 33175 | 7341 |
note scaleR_right_distrib [simp] |
| 53348 | 7342 |
show ?t1 |
7343 |
unfolding midpoint_def dist_norm |
|
7344 |
apply (rule **) |
|
7345 |
apply (simp add: scaleR_right_diff_distrib) |
|
7346 |
apply (simp add: scaleR_2) |
|
7347 |
done |
|
7348 |
show ?t2 |
|
7349 |
unfolding midpoint_def dist_norm |
|
7350 |
apply (rule *) |
|
7351 |
apply (simp add: scaleR_right_diff_distrib) |
|
7352 |
apply (simp add: scaleR_2) |
|
7353 |
done |
|
7354 |
show ?t3 |
|
|
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
|
7355 |
unfolding midpoint_def dist_norm |
| 53348 | 7356 |
apply (rule *) |
7357 |
apply (simp add: scaleR_right_diff_distrib) |
|
7358 |
apply (simp add: scaleR_2) |
|
7359 |
done |
|
7360 |
show ?t4 |
|
7361 |
unfolding midpoint_def dist_norm |
|
7362 |
apply (rule **) |
|
7363 |
apply (simp add: scaleR_right_diff_distrib) |
|
7364 |
apply (simp add: scaleR_2) |
|
7365 |
done |
|
| 36338 | 7366 |
qed |
| 33175 | 7367 |
|
| 63301 | 7368 |
lemma midpoint_eq_endpoint [simp]: |
| 36338 | 7369 |
"midpoint a b = a \<longleftrightarrow> a = b" |
| 33175 | 7370 |
"midpoint a b = b \<longleftrightarrow> a = b" |
| 36338 | 7371 |
unfolding midpoint_eq_iff by auto |
| 33175 | 7372 |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7373 |
lemma midpoint_plus_self [simp]: "midpoint a b + midpoint a b = a + b" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7374 |
using midpoint_eq_iff by metis |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7375 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7376 |
lemma midpoint_linear_image: |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7377 |
"linear f \<Longrightarrow> midpoint(f a)(f b) = f(midpoint a b)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7378 |
by (simp add: linear_iff midpoint_def) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7379 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7380 |
subsection\<open>Starlike sets\<close> |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7381 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7382 |
definition "starlike S \<longleftrightarrow> (\<exists>a\<in>S. \<forall>x\<in>S. closed_segment a x \<subseteq> S)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7383 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7384 |
lemma starlike_UNIV [simp]: "starlike UNIV" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7385 |
by (simp add: starlike_def) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7386 |
|
| 33175 | 7387 |
lemma convex_contains_segment: |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7388 |
"convex S \<longleftrightarrow> (\<forall>a\<in>S. \<forall>b\<in>S. closed_segment a b \<subseteq> S)" |
| 33175 | 7389 |
unfolding convex_alt closed_segment_def by auto |
7390 |
||
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7391 |
lemma closed_segment_subset: "\<lbrakk>x \<in> S; y \<in> S; convex S\<rbrakk> \<Longrightarrow> closed_segment x y \<subseteq> S" |
| 61848 | 7392 |
by (simp add: convex_contains_segment) |
7393 |
||
| 60762 | 7394 |
lemma closed_segment_subset_convex_hull: |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7395 |
"\<lbrakk>x \<in> convex hull S; y \<in> convex hull S\<rbrakk> \<Longrightarrow> closed_segment x y \<subseteq> convex hull S" |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
7396 |
using convex_contains_segment by blast |
| 60762 | 7397 |
|
| 33175 | 7398 |
lemma convex_imp_starlike: |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
7399 |
"convex S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> starlike S"
|
| 33175 | 7400 |
unfolding convex_contains_segment starlike_def by auto |
7401 |
||
7402 |
lemma segment_convex_hull: |
|
| 53348 | 7403 |
"closed_segment a b = convex hull {a,b}"
|
7404 |
proof - |
|
7405 |
have *: "\<And>x. {x} \<noteq> {}" by auto
|
|
7406 |
show ?thesis |
|
7407 |
unfolding segment convex_hull_insert[OF *] convex_hull_singleton |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7408 |
by (safe; rule_tac x="1 - u" in exI; force) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7409 |
qed |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7410 |
|
|
61520
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7411 |
lemma open_closed_segment: "u \<in> open_segment w z \<Longrightarrow> u \<in> closed_segment w z" |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7412 |
by (auto simp add: closed_segment_def open_segment_def) |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7413 |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7414 |
lemma segment_open_subset_closed: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7415 |
"open_segment a b \<subseteq> closed_segment a b" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7416 |
by (auto simp: closed_segment_def open_segment_def) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7417 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7418 |
lemma bounded_closed_segment: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7419 |
fixes a :: "'a::euclidean_space" shows "bounded (closed_segment a b)" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7420 |
by (simp add: segment_convex_hull compact_convex_hull compact_imp_bounded) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7421 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7422 |
lemma bounded_open_segment: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7423 |
fixes a :: "'a::euclidean_space" shows "bounded (open_segment a b)" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7424 |
by (rule bounded_subset [OF bounded_closed_segment segment_open_subset_closed]) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7425 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7426 |
lemmas bounded_segment = bounded_closed_segment open_closed_segment |
|
61426
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61222
diff
changeset
|
7427 |
|
|
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61222
diff
changeset
|
7428 |
lemma ends_in_segment [iff]: "a \<in> closed_segment a b" "b \<in> closed_segment a b" |
| 53348 | 7429 |
unfolding segment_convex_hull |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7430 |
by (auto intro!: hull_subset[unfolded subset_eq, rule_format]) |
| 33175 | 7431 |
|
|
63957
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7432 |
lemma eventually_closed_segment: |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7433 |
fixes x0::"'a::real_normed_vector" |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7434 |
assumes "open X0" "x0 \<in> X0" |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7435 |
shows "\<forall>\<^sub>F x in at x0 within U. closed_segment x0 x \<subseteq> X0" |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7436 |
proof - |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7437 |
from openE[OF assms] |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7438 |
obtain e where e: "0 < e" "ball x0 e \<subseteq> X0" . |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7439 |
then have "\<forall>\<^sub>F x in at x0 within U. x \<in> ball x0 e" |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7440 |
by (auto simp: dist_commute eventually_at) |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7441 |
then show ?thesis |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7442 |
proof eventually_elim |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7443 |
case (elim x) |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7444 |
have "x0 \<in> ball x0 e" using \<open>e > 0\<close> by simp |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7445 |
from convex_ball[unfolded convex_contains_segment, rule_format, OF this elim] |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7446 |
have "closed_segment x0 x \<subseteq> ball x0 e" . |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7447 |
also note \<open>\<dots> \<subseteq> X0\<close> |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7448 |
finally show ?case . |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7449 |
qed |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7450 |
qed |
|
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
7451 |
|
| 33175 | 7452 |
lemma segment_furthest_le: |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36844
diff
changeset
|
7453 |
fixes a b x y :: "'a::euclidean_space" |
| 53348 | 7454 |
assumes "x \<in> closed_segment a b" |
7455 |
shows "norm (y - x) \<le> norm (y - a) \<or> norm (y - x) \<le> norm (y - b)" |
|
7456 |
proof - |
|
7457 |
obtain z where "z \<in> {a, b}" "norm (x - y) \<le> norm (z - y)"
|
|
7458 |
using simplex_furthest_le[of "{a, b}" y]
|
|
7459 |
using assms[unfolded segment_convex_hull] |
|
7460 |
by auto |
|
7461 |
then show ?thesis |
|
7462 |
by (auto simp add:norm_minus_commute) |
|
7463 |
qed |
|
| 33175 | 7464 |
|
| 60176 | 7465 |
lemma closed_segment_commute: "closed_segment a b = closed_segment b a" |
7466 |
proof - |
|
7467 |
have "{a, b} = {b, a}" by auto
|
|
7468 |
thus ?thesis |
|
7469 |
by (simp add: segment_convex_hull) |
|
7470 |
qed |
|
7471 |
||
|
61738
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7472 |
lemma segment_bound1: |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7473 |
assumes "x \<in> closed_segment a b" |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7474 |
shows "norm (x - a) \<le> norm (b - a)" |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7475 |
proof - |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7476 |
obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1" |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7477 |
using assms by (auto simp add: closed_segment_def) |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7478 |
then show "norm (x - a) \<le> norm (b - a)" |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7479 |
apply clarify |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7480 |
apply (auto simp: algebra_simps) |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7481 |
apply (simp add: scaleR_diff_right [symmetric] mult_left_le_one_le) |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7482 |
done |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7483 |
qed |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7484 |
|
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7485 |
lemma segment_bound: |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7486 |
assumes "x \<in> closed_segment a b" |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7487 |
shows "norm (x - a) \<le> norm (b - a)" "norm (x - b) \<le> norm (b - a)" |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7488 |
apply (simp add: assms segment_bound1) |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7489 |
by (metis assms closed_segment_commute dist_commute dist_norm segment_bound1) |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7490 |
|
|
61520
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7491 |
lemma open_segment_commute: "open_segment a b = open_segment b a" |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7492 |
proof - |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7493 |
have "{a, b} = {b, a}" by auto
|
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7494 |
thus ?thesis |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7495 |
by (simp add: closed_segment_commute open_segment_def) |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7496 |
qed |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7497 |
|
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7498 |
lemma closed_segment_idem [simp]: "closed_segment a a = {a}"
|
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7499 |
unfolding segment by (auto simp add: algebra_simps) |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7500 |
|
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7501 |
lemma open_segment_idem [simp]: "open_segment a a = {}"
|
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7502 |
by (simp add: open_segment_def) |
|
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents:
61518
diff
changeset
|
7503 |
|
|
62626
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
7504 |
lemma closed_segment_eq_open: "closed_segment a b = open_segment a b \<union> {a,b}"
|
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
7505 |
using open_segment_def by auto |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
7506 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
7507 |
lemma convex_contains_open_segment: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
7508 |
"convex s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. open_segment a b \<subseteq> s)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
7509 |
by (simp add: convex_contains_segment closed_segment_eq_open) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
7510 |
|
| 60176 | 7511 |
lemma closed_segment_eq_real_ivl: |
7512 |
fixes a b::real |
|
7513 |
shows "closed_segment a b = (if a \<le> b then {a .. b} else {b .. a})"
|
|
7514 |
proof - |
|
7515 |
have "b \<le> a \<Longrightarrow> closed_segment b a = {b .. a}"
|
|
7516 |
and "a \<le> b \<Longrightarrow> closed_segment a b = {a .. b}"
|
|
7517 |
by (auto simp: convex_hull_eq_real_cbox segment_convex_hull) |
|
7518 |
thus ?thesis |
|
7519 |
by (auto simp: closed_segment_commute) |
|
7520 |
qed |
|
7521 |
||
|
64006
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
7522 |
lemma open_segment_eq_real_ivl: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
7523 |
fixes a b::real |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
7524 |
shows "open_segment a b = (if a \<le> b then {a<..<b} else {b<..<a})"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
7525 |
by (auto simp: closed_segment_eq_real_ivl open_segment_def split: if_split_asm) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
7526 |
|
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
7527 |
lemma closed_segment_real_eq: |
|
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
7528 |
fixes u::real shows "closed_segment u v = (\<lambda>x. (v - u) * x + u) ` {0..1}"
|
|
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
7529 |
by (simp add: add.commute [of u] image_affinity_atLeastAtMost [where c=u] closed_segment_eq_real_ivl) |
|
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
7530 |
|
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7531 |
lemma dist_in_closed_segment: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7532 |
fixes a :: "'a :: euclidean_space" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7533 |
assumes "x \<in> closed_segment a b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7534 |
shows "dist x a \<le> dist a b \<and> dist x b \<le> dist a b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7535 |
proof (intro conjI) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7536 |
obtain u where u: "0 \<le> u" "u \<le> 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7537 |
using assms by (force simp: in_segment algebra_simps) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7538 |
have "dist x a = u * dist a b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7539 |
apply (simp add: dist_norm algebra_simps x) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7540 |
by (metis \<open>0 \<le> u\<close> abs_of_nonneg norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7541 |
also have "... \<le> dist a b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7542 |
by (simp add: mult_left_le_one_le u) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7543 |
finally show "dist x a \<le> dist a b" . |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7544 |
have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7545 |
by (simp add: dist_norm algebra_simps x) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7546 |
also have "... = (1-u) * dist a b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7547 |
proof - |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7548 |
have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7549 |
using \<open>u \<le> 1\<close> by force |
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7550 |
then show ?thesis |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7551 |
by (simp add: dist_norm real_vector.scale_right_diff_distrib) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7552 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7553 |
also have "... \<le> dist a b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7554 |
by (simp add: mult_left_le_one_le u) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7555 |
finally show "dist x b \<le> dist a b" . |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7556 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7557 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7558 |
lemma dist_in_open_segment: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7559 |
fixes a :: "'a :: euclidean_space" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7560 |
assumes "x \<in> open_segment a b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7561 |
shows "dist x a < dist a b \<and> dist x b < dist a b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7562 |
proof (intro conjI) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7563 |
obtain u where u: "0 < u" "u < 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7564 |
using assms by (force simp: in_segment algebra_simps) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7565 |
have "dist x a = u * dist a b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7566 |
apply (simp add: dist_norm algebra_simps x) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7567 |
by (metis abs_of_nonneg less_eq_real_def norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib \<open>0 < u\<close>) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7568 |
also have *: "... < dist a b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7569 |
by (metis (no_types) assms dist_eq_0_iff dist_not_less_zero in_segment(2) linorder_neqE_linordered_idom mult.left_neutral real_mult_less_iff1 \<open>u < 1\<close>) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7570 |
finally show "dist x a < dist a b" . |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7571 |
have ab_ne0: "dist a b \<noteq> 0" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7572 |
using * by fastforce |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7573 |
have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7574 |
by (simp add: dist_norm algebra_simps x) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7575 |
also have "... = (1-u) * dist a b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7576 |
proof - |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7577 |
have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7578 |
using \<open>u < 1\<close> by force |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7579 |
then show ?thesis |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7580 |
by (simp add: dist_norm real_vector.scale_right_diff_distrib) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7581 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7582 |
also have "... < dist a b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7583 |
using ab_ne0 \<open>0 < u\<close> by simp |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7584 |
finally show "dist x b < dist a b" . |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7585 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7586 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7587 |
lemma dist_decreases_open_segment_0: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7588 |
fixes x :: "'a :: euclidean_space" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7589 |
assumes "x \<in> open_segment 0 b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7590 |
shows "dist c x < dist c 0 \<or> dist c x < dist c b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7591 |
proof (rule ccontr, clarsimp simp: not_less) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7592 |
obtain u where u: "0 \<noteq> b" "0 < u" "u < 1" and x: "x = u *\<^sub>R b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7593 |
using assms by (auto simp: in_segment) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7594 |
have xb: "x \<bullet> b < b \<bullet> b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7595 |
using u x by auto |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7596 |
assume "norm c \<le> dist c x" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7597 |
then have "c \<bullet> c \<le> (c - x) \<bullet> (c - x)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7598 |
by (simp add: dist_norm norm_le) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7599 |
moreover have "0 < x \<bullet> b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7600 |
using u x by auto |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7601 |
ultimately have less: "c \<bullet> b < x \<bullet> b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7602 |
by (simp add: x algebra_simps inner_commute u) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7603 |
assume "dist c b \<le> dist c x" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7604 |
then have "(c - b) \<bullet> (c - b) \<le> (c - x) \<bullet> (c - x)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7605 |
by (simp add: dist_norm norm_le) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7606 |
then have "(b \<bullet> b) * (1 - u*u) \<le> 2 * (b \<bullet> c) * (1-u)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7607 |
by (simp add: x algebra_simps inner_commute) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7608 |
then have "(1+u) * (b \<bullet> b) * (1-u) \<le> 2 * (b \<bullet> c) * (1-u)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7609 |
by (simp add: algebra_simps) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7610 |
then have "(1+u) * (b \<bullet> b) \<le> 2 * (b \<bullet> c)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7611 |
using \<open>u < 1\<close> by auto |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7612 |
with xb have "c \<bullet> b \<ge> x \<bullet> b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7613 |
by (auto simp: x algebra_simps inner_commute) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7614 |
with less show False by auto |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7615 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7616 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7617 |
proposition dist_decreases_open_segment: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7618 |
fixes a :: "'a :: euclidean_space" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7619 |
assumes "x \<in> open_segment a b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7620 |
shows "dist c x < dist c a \<or> dist c x < dist c b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7621 |
proof - |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7622 |
have *: "x - a \<in> open_segment 0 (b - a)" using assms |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7623 |
by (metis diff_self open_segment_translation_eq uminus_add_conv_diff) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7624 |
show ?thesis |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7625 |
using dist_decreases_open_segment_0 [OF *, of "c-a"] assms |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7626 |
by (simp add: dist_norm) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7627 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7628 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7629 |
lemma dist_decreases_closed_segment: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7630 |
fixes a :: "'a :: euclidean_space" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7631 |
assumes "x \<in> closed_segment a b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7632 |
shows "dist c x \<le> dist c a \<or> dist c x \<le> dist c b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7633 |
apply (cases "x \<in> open_segment a b") |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7634 |
using dist_decreases_open_segment less_eq_real_def apply blast |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7635 |
by (metis DiffI assms empty_iff insertE open_segment_def order_refl) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7636 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7637 |
lemma convex_intermediate_ball: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7638 |
fixes a :: "'a :: euclidean_space" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7639 |
shows "\<lbrakk>ball a r \<subseteq> T; T \<subseteq> cball a r\<rbrakk> \<Longrightarrow> convex T" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7640 |
apply (simp add: convex_contains_open_segment, clarify) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7641 |
by (metis (no_types, hide_lams) less_le_trans mem_ball mem_cball subsetCE dist_decreases_open_segment) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
7642 |
|
| 63301 | 7643 |
lemma csegment_midpoint_subset: "closed_segment (midpoint a b) b \<subseteq> closed_segment a b" |
7644 |
apply (clarsimp simp: midpoint_def in_segment) |
|
7645 |
apply (rule_tac x="(1 + u) / 2" in exI) |
|
7646 |
apply (auto simp: algebra_simps add_divide_distrib diff_divide_distrib) |
|
7647 |
by (metis real_sum_of_halves scaleR_left.add) |
|
7648 |
||
7649 |
lemma notin_segment_midpoint: |
|
7650 |
fixes a :: "'a :: euclidean_space" |
|
7651 |
shows "a \<noteq> b \<Longrightarrow> a \<notin> closed_segment (midpoint a b) b" |
|
7652 |
by (auto simp: dist_midpoint dest!: dist_in_closed_segment) |
|
7653 |
||
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7654 |
lemma segment_to_closest_point: |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7655 |
fixes S :: "'a :: euclidean_space set" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7656 |
shows "\<lbrakk>closed S; S \<noteq> {}\<rbrakk> \<Longrightarrow> open_segment a (closest_point S a) \<inter> S = {}"
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7657 |
apply (subst disjoint_iff_not_equal) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7658 |
apply (clarify dest!: dist_in_open_segment) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7659 |
by (metis closest_point_le dist_commute le_less_trans less_irrefl) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7660 |
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7661 |
lemma segment_to_point_exists: |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7662 |
fixes S :: "'a :: euclidean_space set" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7663 |
assumes "closed S" "S \<noteq> {}"
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7664 |
obtains b where "b \<in> S" "open_segment a b \<inter> S = {}"
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7665 |
by (metis assms segment_to_closest_point closest_point_exists that) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
7666 |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7667 |
subsubsection\<open>More lemmas, especially for working with the underlying formula\<close> |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7668 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7669 |
lemma segment_eq_compose: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7670 |
fixes a :: "'a :: real_vector" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7671 |
shows "(\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) = (\<lambda>x. a + x) o (\<lambda>u. u *\<^sub>R (b - a))" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7672 |
by (simp add: o_def algebra_simps) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7673 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7674 |
lemma segment_degen_1: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7675 |
fixes a :: "'a :: real_vector" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7676 |
shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = b \<longleftrightarrow> a=b \<or> u=1" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7677 |
proof - |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7678 |
{ assume "(1 - u) *\<^sub>R a + u *\<^sub>R b = b"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7679 |
then have "(1 - u) *\<^sub>R a = (1 - u) *\<^sub>R b" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7680 |
by (simp add: algebra_simps) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7681 |
then have "a=b \<or> u=1" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7682 |
by simp |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7683 |
} then show ?thesis |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7684 |
by (auto simp: algebra_simps) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7685 |
qed |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7686 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7687 |
lemma segment_degen_0: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7688 |
fixes a :: "'a :: real_vector" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7689 |
shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = a \<longleftrightarrow> a=b \<or> u=0" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7690 |
using segment_degen_1 [of "1-u" b a] |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7691 |
by (auto simp: algebra_simps) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7692 |
|
| 64394 | 7693 |
lemma add_scaleR_degen: |
7694 |
fixes a b ::"'a::real_vector" |
|
7695 |
assumes "(u *\<^sub>R b + v *\<^sub>R a) = (u *\<^sub>R a + v *\<^sub>R b)" "u \<noteq> v" |
|
7696 |
shows "a=b" |
|
7697 |
by (metis (no_types, hide_lams) add.commute add_diff_eq diff_add_cancel real_vector.scale_cancel_left real_vector.scale_left_diff_distrib assms) |
|
7698 |
||
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7699 |
lemma closed_segment_image_interval: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7700 |
"closed_segment a b = (\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0..1}"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7701 |
by (auto simp: set_eq_iff image_iff closed_segment_def) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7702 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7703 |
lemma open_segment_image_interval: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7704 |
"open_segment a b = (if a=b then {} else (\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1})"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7705 |
by (auto simp: open_segment_def closed_segment_def segment_degen_0 segment_degen_1) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7706 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7707 |
lemmas segment_image_interval = closed_segment_image_interval open_segment_image_interval |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7708 |
|
|
61738
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7709 |
lemma open_segment_bound1: |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7710 |
assumes "x \<in> open_segment a b" |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7711 |
shows "norm (x - a) < norm (b - a)" |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7712 |
proof - |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7713 |
obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 < u" "u < 1" "a \<noteq> b" |
|
62618
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7714 |
using assms by (auto simp add: open_segment_image_interval split: if_split_asm) |
|
61738
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7715 |
then show "norm (x - a) < norm (b - a)" |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7716 |
apply clarify |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7717 |
apply (auto simp: algebra_simps) |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7718 |
apply (simp add: scaleR_diff_right [symmetric]) |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7719 |
done |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7720 |
qed |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7721 |
|
| 62950 | 7722 |
lemma compact_segment [simp]: |
7723 |
fixes a :: "'a::real_normed_vector" |
|
7724 |
shows "compact (closed_segment a b)" |
|
7725 |
by (auto simp: segment_image_interval intro!: compact_continuous_image continuous_intros) |
|
7726 |
||
7727 |
lemma closed_segment [simp]: |
|
7728 |
fixes a :: "'a::real_normed_vector" |
|
7729 |
shows "closed (closed_segment a b)" |
|
7730 |
by (simp add: compact_imp_closed) |
|
7731 |
||
7732 |
lemma closure_closed_segment [simp]: |
|
7733 |
fixes a :: "'a::real_normed_vector" |
|
7734 |
shows "closure(closed_segment a b) = closed_segment a b" |
|
7735 |
by simp |
|
7736 |
||
|
61738
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7737 |
lemma open_segment_bound: |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7738 |
assumes "x \<in> open_segment a b" |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7739 |
shows "norm (x - a) < norm (b - a)" "norm (x - b) < norm (b - a)" |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7740 |
apply (simp add: assms open_segment_bound1) |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7741 |
by (metis assms norm_minus_commute open_segment_bound1 open_segment_commute) |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61694
diff
changeset
|
7742 |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7743 |
lemma closure_open_segment [simp]: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7744 |
fixes a :: "'a::euclidean_space" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7745 |
shows "closure(open_segment a b) = (if a = b then {} else closed_segment a b)"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7746 |
proof - |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7747 |
have "closure ((\<lambda>u. u *\<^sub>R (b - a)) ` {0<..<1}) = (\<lambda>u. u *\<^sub>R (b - a)) ` closure {0<..<1}" if "a \<noteq> b"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7748 |
apply (rule closure_injective_linear_image [symmetric]) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7749 |
apply (simp add:) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7750 |
using that by (simp add: inj_on_def) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7751 |
then show ?thesis |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7752 |
by (simp add: segment_image_interval segment_eq_compose closure_greaterThanLessThan [symmetric] |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7753 |
closure_translation image_comp [symmetric] del: closure_greaterThanLessThan) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7754 |
qed |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7755 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7756 |
lemma closed_open_segment_iff [simp]: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7757 |
fixes a :: "'a::euclidean_space" shows "closed(open_segment a b) \<longleftrightarrow> a = b" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7758 |
by (metis open_segment_def DiffE closure_eq closure_open_segment ends_in_segment(1) insert_iff segment_image_interval(2)) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7759 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7760 |
lemma compact_open_segment_iff [simp]: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7761 |
fixes a :: "'a::euclidean_space" shows "compact(open_segment a b) \<longleftrightarrow> a = b" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7762 |
by (simp add: bounded_open_segment compact_eq_bounded_closed) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7763 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7764 |
lemma convex_closed_segment [iff]: "convex (closed_segment a b)" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7765 |
unfolding segment_convex_hull by(rule convex_convex_hull) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7766 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7767 |
lemma convex_open_segment [iff]: "convex(open_segment a b)" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7768 |
proof - |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7769 |
have "convex ((\<lambda>u. u *\<^sub>R (b-a)) ` {0<..<1})"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7770 |
by (rule convex_linear_image) auto |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7771 |
then show ?thesis |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7772 |
apply (simp add: open_segment_image_interval segment_eq_compose) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7773 |
by (metis image_comp convex_translation) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7774 |
qed |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7775 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7776 |
lemmas convex_segment = convex_closed_segment convex_open_segment |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7777 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7778 |
lemma connected_segment [iff]: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7779 |
fixes x :: "'a :: real_normed_vector" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7780 |
shows "connected (closed_segment x y)" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7781 |
by (simp add: convex_connected) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7782 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7783 |
lemma affine_hull_closed_segment [simp]: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7784 |
"affine hull (closed_segment a b) = affine hull {a,b}"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7785 |
by (simp add: segment_convex_hull) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7786 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7787 |
lemma affine_hull_open_segment [simp]: |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7788 |
fixes a :: "'a::euclidean_space" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7789 |
shows "affine hull (open_segment a b) = (if a = b then {} else affine hull {a,b})"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7790 |
by (metis affine_hull_convex_hull affine_hull_empty closure_open_segment closure_same_affine_hull segment_convex_hull) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
7791 |
|
|
62618
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7792 |
lemma rel_interior_closure_convex_segment: |
|
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7793 |
fixes S :: "_::euclidean_space set" |
|
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7794 |
assumes "convex S" "a \<in> rel_interior S" "b \<in> closure S" |
|
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7795 |
shows "open_segment a b \<subseteq> rel_interior S" |
|
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7796 |
proof |
|
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7797 |
fix x |
|
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7798 |
have [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" for u |
|
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7799 |
by (simp add: algebra_simps) |
|
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7800 |
assume "x \<in> open_segment a b" |
|
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7801 |
then show "x \<in> rel_interior S" |
|
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7802 |
unfolding closed_segment_def open_segment_def using assms |
|
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7803 |
by (auto intro: rel_interior_closure_convex_shrink) |
|
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7804 |
qed |
|
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
7805 |
|
|
62620
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7806 |
subsection\<open>More results about segments\<close> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7807 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7808 |
lemma dist_half_times2: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7809 |
fixes a :: "'a :: real_normed_vector" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7810 |
shows "dist ((1 / 2) *\<^sub>R (a + b)) x * 2 = dist (a+b) (2 *\<^sub>R x)" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7811 |
proof - |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7812 |
have "norm ((1 / 2) *\<^sub>R (a + b) - x) * 2 = norm (2 *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x))" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7813 |
by simp |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7814 |
also have "... = norm ((a + b) - 2 *\<^sub>R x)" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7815 |
by (simp add: real_vector.scale_right_diff_distrib) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7816 |
finally show ?thesis |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7817 |
by (simp only: dist_norm) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7818 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7819 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7820 |
lemma closed_segment_as_ball: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7821 |
"closed_segment a b = affine hull {a,b} \<inter> cball(inverse 2 *\<^sub>R (a + b))(norm(b - a) / 2)"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7822 |
proof (cases "b = a") |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7823 |
case True then show ?thesis by (auto simp: hull_inc) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7824 |
next |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7825 |
case False |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7826 |
then have *: "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7827 |
dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \<le> norm (b - a)) = |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7828 |
(\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1)" for x |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7829 |
proof - |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7830 |
have "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7831 |
dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \<le> norm (b - a)) = |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7832 |
((\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b) \<and> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7833 |
dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \<le> norm (b - a))" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7834 |
unfolding eq_diff_eq [symmetric] by simp |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7835 |
also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7836 |
norm ((a+b) - (2 *\<^sub>R x)) \<le> norm (b - a))" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7837 |
by (simp add: dist_half_times2) (simp add: dist_norm) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7838 |
also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7839 |
norm ((a+b) - (2 *\<^sub>R ((1 - u) *\<^sub>R a + u *\<^sub>R b))) \<le> norm (b - a))" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7840 |
by auto |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7841 |
also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7842 |
norm ((1 - u * 2) *\<^sub>R (b - a)) \<le> norm (b - a))" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7843 |
by (simp add: algebra_simps scaleR_2) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7844 |
also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7845 |
\<bar>1 - u * 2\<bar> * norm (b - a) \<le> norm (b - a))" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7846 |
by simp |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7847 |
also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> \<bar>1 - u * 2\<bar> \<le> 1)" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7848 |
by (simp add: mult_le_cancel_right2 False) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7849 |
also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1)" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7850 |
by auto |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7851 |
finally show ?thesis . |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7852 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7853 |
show ?thesis |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7854 |
by (simp add: affine_hull_2 Set.set_eq_iff closed_segment_def *) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7855 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7856 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7857 |
lemma open_segment_as_ball: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7858 |
"open_segment a b = |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7859 |
affine hull {a,b} \<inter> ball(inverse 2 *\<^sub>R (a + b))(norm(b - a) / 2)"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7860 |
proof (cases "b = a") |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7861 |
case True then show ?thesis by (auto simp: hull_inc) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7862 |
next |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7863 |
case False |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7864 |
then have *: "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7865 |
dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a)) = |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7866 |
(\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 < u \<and> u < 1)" for x |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7867 |
proof - |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7868 |
have "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7869 |
dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a)) = |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7870 |
((\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b) \<and> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7871 |
dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a))" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7872 |
unfolding eq_diff_eq [symmetric] by simp |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7873 |
also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7874 |
norm ((a+b) - (2 *\<^sub>R x)) < norm (b - a))" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7875 |
by (simp add: dist_half_times2) (simp add: dist_norm) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7876 |
also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7877 |
norm ((a+b) - (2 *\<^sub>R ((1 - u) *\<^sub>R a + u *\<^sub>R b))) < norm (b - a))" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7878 |
by auto |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7879 |
also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7880 |
norm ((1 - u * 2) *\<^sub>R (b - a)) < norm (b - a))" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7881 |
by (simp add: algebra_simps scaleR_2) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7882 |
also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7883 |
\<bar>1 - u * 2\<bar> * norm (b - a) < norm (b - a))" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7884 |
by simp |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7885 |
also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> \<bar>1 - u * 2\<bar> < 1)" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7886 |
by (simp add: mult_le_cancel_right2 False) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7887 |
also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 < u \<and> u < 1)" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7888 |
by auto |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7889 |
finally show ?thesis . |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7890 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7891 |
show ?thesis |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7892 |
using False by (force simp: affine_hull_2 Set.set_eq_iff open_segment_image_interval *) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7893 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7894 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7895 |
lemmas segment_as_ball = closed_segment_as_ball open_segment_as_ball |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7896 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7897 |
lemma closed_segment_neq_empty [simp]: "closed_segment a b \<noteq> {}"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7898 |
by auto |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7899 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7900 |
lemma open_segment_eq_empty [simp]: "open_segment a b = {} \<longleftrightarrow> a = b"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7901 |
proof - |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7902 |
{ assume a1: "open_segment a b = {}"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7903 |
have "{} \<noteq> {0::real<..<1}"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7904 |
by simp |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7905 |
then have "a = b" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7906 |
using a1 open_segment_image_interval by fastforce |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7907 |
} then show ?thesis by auto |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7908 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7909 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7910 |
lemma open_segment_eq_empty' [simp]: "{} = open_segment a b \<longleftrightarrow> a = b"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7911 |
using open_segment_eq_empty by blast |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7912 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7913 |
lemmas segment_eq_empty = closed_segment_neq_empty open_segment_eq_empty |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7914 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7915 |
lemma inj_segment: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7916 |
fixes a :: "'a :: real_vector" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7917 |
assumes "a \<noteq> b" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7918 |
shows "inj_on (\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) I" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7919 |
proof |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7920 |
fix x y |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7921 |
assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7922 |
then have "x *\<^sub>R (b - a) = y *\<^sub>R (b - a)" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7923 |
by (simp add: algebra_simps) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7924 |
with assms show "x = y" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7925 |
by (simp add: real_vector.scale_right_imp_eq) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7926 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7927 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7928 |
lemma finite_closed_segment [simp]: "finite(closed_segment a b) \<longleftrightarrow> a = b" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7929 |
apply auto |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7930 |
apply (rule ccontr) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7931 |
apply (simp add: segment_image_interval) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7932 |
using infinite_Icc [OF zero_less_one] finite_imageD [OF _ inj_segment] apply blast |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7933 |
done |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7934 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7935 |
lemma finite_open_segment [simp]: "finite(open_segment a b) \<longleftrightarrow> a = b" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7936 |
by (auto simp: open_segment_def) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7937 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7938 |
lemmas finite_segment = finite_closed_segment finite_open_segment |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7939 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7940 |
lemma closed_segment_eq_sing: "closed_segment a b = {c} \<longleftrightarrow> a = c \<and> b = c"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7941 |
by auto |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7942 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7943 |
lemma open_segment_eq_sing: "open_segment a b \<noteq> {c}"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7944 |
by (metis finite_insert finite_open_segment insert_not_empty open_segment_image_interval) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7945 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7946 |
lemmas segment_eq_sing = closed_segment_eq_sing open_segment_eq_sing |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7947 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7948 |
lemma subset_closed_segment: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7949 |
"closed_segment a b \<subseteq> closed_segment c d \<longleftrightarrow> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7950 |
a \<in> closed_segment c d \<and> b \<in> closed_segment c d" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7951 |
by auto (meson contra_subsetD convex_closed_segment convex_contains_segment) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7952 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7953 |
lemma subset_co_segment: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7954 |
"closed_segment a b \<subseteq> open_segment c d \<longleftrightarrow> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7955 |
a \<in> open_segment c d \<and> b \<in> open_segment c d" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7956 |
using closed_segment_subset by blast |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7957 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7958 |
lemma subset_open_segment: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7959 |
fixes a :: "'a::euclidean_space" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7960 |
shows "open_segment a b \<subseteq> open_segment c d \<longleftrightarrow> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7961 |
a = b \<or> a \<in> closed_segment c d \<and> b \<in> closed_segment c d" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7962 |
(is "?lhs = ?rhs") |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7963 |
proof (cases "a = b") |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7964 |
case True then show ?thesis by simp |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7965 |
next |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7966 |
case False show ?thesis |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7967 |
proof |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7968 |
assume rhs: ?rhs |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7969 |
with \<open>a \<noteq> b\<close> have "c \<noteq> d" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7970 |
using closed_segment_idem singleton_iff by auto |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7971 |
have "\<exists>uc. (1 - u) *\<^sub>R ((1 - ua) *\<^sub>R c + ua *\<^sub>R d) + u *\<^sub>R ((1 - ub) *\<^sub>R c + ub *\<^sub>R d) = |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7972 |
(1 - uc) *\<^sub>R c + uc *\<^sub>R d \<and> 0 < uc \<and> uc < 1" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7973 |
if neq: "(1 - ua) *\<^sub>R c + ua *\<^sub>R d \<noteq> (1 - ub) *\<^sub>R c + ub *\<^sub>R d" "c \<noteq> d" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7974 |
and "a = (1 - ua) *\<^sub>R c + ua *\<^sub>R d" "b = (1 - ub) *\<^sub>R c + ub *\<^sub>R d" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7975 |
and u: "0 < u" "u < 1" and uab: "0 \<le> ua" "ua \<le> 1" "0 \<le> ub" "ub \<le> 1" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7976 |
for u ua ub |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7977 |
proof - |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7978 |
have "ua \<noteq> ub" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7979 |
using neq by auto |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7980 |
moreover have "(u - 1) * ua \<le> 0" using u uab |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7981 |
by (simp add: mult_nonpos_nonneg) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7982 |
ultimately have lt: "(u - 1) * ua < u * ub" using u uab |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7983 |
by (metis antisym_conv diff_ge_0_iff_ge le_less_trans mult_eq_0_iff mult_le_0_iff not_less) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7984 |
have "p * ua + q * ub < p+q" if p: "0 < p" and q: "0 < q" for p q |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7985 |
proof - |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7986 |
have "\<not> p \<le> 0" "\<not> q \<le> 0" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7987 |
using p q not_less by blast+ |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7988 |
then show ?thesis |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7989 |
by (metis \<open>ua \<noteq> ub\<close> add_less_cancel_left add_less_cancel_right add_mono_thms_linordered_field(5) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7990 |
less_eq_real_def mult_cancel_left1 mult_less_cancel_left2 uab(2) uab(4)) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7991 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7992 |
then have "(1 - u) * ua + u * ub < 1" using u \<open>ua \<noteq> ub\<close> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7993 |
by (metis diff_add_cancel diff_gt_0_iff_gt) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7994 |
with lt show ?thesis |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7995 |
by (rule_tac x="ua + u*(ub-ua)" in exI) (simp add: algebra_simps) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7996 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7997 |
with rhs \<open>a \<noteq> b\<close> \<open>c \<noteq> d\<close> show ?lhs |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7998 |
unfolding open_segment_image_interval closed_segment_def |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
7999 |
by (fastforce simp add:) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8000 |
next |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8001 |
assume lhs: ?lhs |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8002 |
with \<open>a \<noteq> b\<close> have "c \<noteq> d" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8003 |
by (meson finite_open_segment rev_finite_subset) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8004 |
have "closure (open_segment a b) \<subseteq> closure (open_segment c d)" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8005 |
using lhs closure_mono by blast |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8006 |
then have "closed_segment a b \<subseteq> closed_segment c d" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8007 |
by (simp add: \<open>a \<noteq> b\<close> \<open>c \<noteq> d\<close>) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8008 |
then show ?rhs |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8009 |
by (force simp: \<open>a \<noteq> b\<close>) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8010 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8011 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8012 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8013 |
lemma subset_oc_segment: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8014 |
fixes a :: "'a::euclidean_space" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8015 |
shows "open_segment a b \<subseteq> closed_segment c d \<longleftrightarrow> |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8016 |
a = b \<or> a \<in> closed_segment c d \<and> b \<in> closed_segment c d" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8017 |
apply (simp add: subset_open_segment [symmetric]) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8018 |
apply (rule iffI) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8019 |
apply (metis closure_closed_segment closure_mono closure_open_segment subset_closed_segment subset_open_segment) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8020 |
apply (meson dual_order.trans segment_open_subset_closed) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8021 |
done |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8022 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8023 |
lemmas subset_segment = subset_closed_segment subset_co_segment subset_oc_segment subset_open_segment |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8024 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
8025 |
|
|
62618
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
8026 |
subsection\<open>Betweenness\<close> |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
8027 |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8028 |
definition "between = (\<lambda>(a,b) x. x \<in> closed_segment a b)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8029 |
|
| 33175 | 8030 |
lemma between_mem_segment: "between (a,b) x \<longleftrightarrow> x \<in> closed_segment a b" |
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44142
diff
changeset
|
8031 |
unfolding between_def by auto |
| 33175 | 8032 |
|
| 53348 | 8033 |
lemma between: "between (a, b) (x::'a::euclidean_space) \<longleftrightarrow> dist a b = (dist a x) + (dist x b)" |
8034 |
proof (cases "a = b") |
|
8035 |
case True |
|
8036 |
then show ?thesis |
|
8037 |
unfolding between_def split_conv |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
8038 |
by (auto simp add: dist_commute) |
| 53348 | 8039 |
next |
8040 |
case False |
|
8041 |
then have Fal: "norm (a - b) \<noteq> 0" and Fal2: "norm (a - b) > 0" |
|
8042 |
by auto |
|
8043 |
have *: "\<And>u. a - ((1 - u) *\<^sub>R a + u *\<^sub>R b) = u *\<^sub>R (a - b)" |
|
8044 |
by (auto simp add: algebra_simps) |
|
8045 |
show ?thesis |
|
8046 |
unfolding between_def split_conv closed_segment_def mem_Collect_eq |
|
8047 |
apply rule |
|
8048 |
apply (elim exE conjE) |
|
8049 |
apply (subst dist_triangle_eq) |
|
8050 |
proof - |
|
8051 |
fix u |
|
8052 |
assume as: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1" |
|
8053 |
then have *: "a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)" |
|
8054 |
unfolding as(1) by (auto simp add:algebra_simps) |
|
8055 |
show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)" |
|
8056 |
unfolding norm_minus_commute[of x a] * using as(2,3) |
|
8057 |
by (auto simp add: field_simps) |
|
8058 |
next |
|
8059 |
assume as: "dist a b = dist a x + dist x b" |
|
8060 |
have "norm (a - x) / norm (a - b) \<le> 1" |
|
|
56571
f4635657d66f
added divide_nonneg_nonneg and co; made it a simp rule
hoelzl
parents:
56544
diff
changeset
|
8061 |
using Fal2 unfolding as[unfolded dist_norm] norm_ge_zero by auto |
| 53348 | 8062 |
then show "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1" |
8063 |
apply (rule_tac x="dist a x / dist a b" in exI) |
|
8064 |
unfolding dist_norm |
|
8065 |
apply (subst euclidean_eq_iff) |
|
8066 |
apply rule |
|
8067 |
defer |
|
8068 |
apply rule |
|
|
56571
f4635657d66f
added divide_nonneg_nonneg and co; made it a simp rule
hoelzl
parents:
56544
diff
changeset
|
8069 |
prefer 3 |
| 53348 | 8070 |
apply rule |
8071 |
proof - |
|
8072 |
fix i :: 'a |
|
8073 |
assume i: "i \<in> Basis" |
|
8074 |
have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \<bullet> i = |
|
8075 |
((norm (a - b) - norm (a - x)) * (a \<bullet> i) + norm (a - x) * (b \<bullet> i)) / norm (a - b)" |
|
8076 |
using Fal by (auto simp add: field_simps inner_simps) |
|
8077 |
also have "\<dots> = x\<bullet>i" |
|
8078 |
apply (rule divide_eq_imp[OF Fal]) |
|
8079 |
unfolding as[unfolded dist_norm] |
|
8080 |
using as[unfolded dist_triangle_eq] |
|
8081 |
apply - |
|
8082 |
apply (subst (asm) euclidean_eq_iff) |
|
8083 |
using i |
|
8084 |
apply (erule_tac x=i in ballE) |
|
| 57865 | 8085 |
apply (auto simp add: field_simps inner_simps) |
| 53348 | 8086 |
done |
8087 |
finally show "x \<bullet> i = |
|
8088 |
((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \<bullet> i" |
|
8089 |
by auto |
|
8090 |
qed (insert Fal2, auto) |
|
8091 |
qed |
|
8092 |
qed |
|
8093 |
||
8094 |
lemma between_midpoint: |
|
8095 |
fixes a :: "'a::euclidean_space" |
|
8096 |
shows "between (a,b) (midpoint a b)" (is ?t1) |
|
8097 |
and "between (b,a) (midpoint a b)" (is ?t2) |
|
8098 |
proof - |
|
8099 |
have *: "\<And>x y z. x = (1/2::real) *\<^sub>R z \<Longrightarrow> y = (1/2) *\<^sub>R z \<Longrightarrow> norm z = norm x + norm y" |
|
8100 |
by auto |
|
8101 |
show ?t1 ?t2 |
|
8102 |
unfolding between midpoint_def dist_norm |
|
8103 |
apply(rule_tac[!] *) |
|
|
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
|
8104 |
unfolding euclidean_eq_iff[where 'a='a] |
| 53348 | 8105 |
apply (auto simp add: field_simps inner_simps) |
8106 |
done |
|
8107 |
qed |
|
| 33175 | 8108 |
|
8109 |
lemma between_mem_convex_hull: |
|
8110 |
"between (a,b) x \<longleftrightarrow> x \<in> convex hull {a,b}"
|
|
8111 |
unfolding between_mem_segment segment_convex_hull .. |
|
8112 |
||
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8113 |
lemma between_triv_iff [simp]: "between (a,a) b \<longleftrightarrow> a=b" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8114 |
by (auto simp: between_def) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8115 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8116 |
lemma between_triv1 [simp]: "between (a,b) a" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8117 |
by (auto simp: between_def) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8118 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8119 |
lemma between_triv2 [simp]: "between (a,b) b" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8120 |
by (auto simp: between_def) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8121 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8122 |
lemma between_commute: |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8123 |
"between (a,b) = between (b,a)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8124 |
by (auto simp: between_def closed_segment_commute) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8125 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8126 |
lemma between_antisym: |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8127 |
fixes a :: "'a :: euclidean_space" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8128 |
shows "\<lbrakk>between (b,c) a; between (a,c) b\<rbrakk> \<Longrightarrow> a = b" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8129 |
by (auto simp: between dist_commute) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8130 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8131 |
lemma between_trans: |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8132 |
fixes a :: "'a :: euclidean_space" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8133 |
shows "\<lbrakk>between (b,c) a; between (a,c) d\<rbrakk> \<Longrightarrow> between (b,c) d" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8134 |
using dist_triangle2 [of b c d] dist_triangle3 [of b d a] |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8135 |
by (auto simp: between dist_commute) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8136 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8137 |
lemma between_norm: |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8138 |
fixes a :: "'a :: euclidean_space" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8139 |
shows "between (a,b) x \<longleftrightarrow> norm(x - a) *\<^sub>R (b - x) = norm(b - x) *\<^sub>R (x - a)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8140 |
by (auto simp: between dist_triangle_eq norm_minus_commute algebra_simps) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
8141 |
|
| 53348 | 8142 |
|
| 60420 | 8143 |
subsection \<open>Shrinking towards the interior of a convex set\<close> |
| 33175 | 8144 |
|
8145 |
lemma mem_interior_convex_shrink: |
|
| 53348 | 8146 |
fixes s :: "'a::euclidean_space set" |
8147 |
assumes "convex s" |
|
8148 |
and "c \<in> interior s" |
|
8149 |
and "x \<in> s" |
|
8150 |
and "0 < e" |
|
8151 |
and "e \<le> 1" |
|
| 33175 | 8152 |
shows "x - e *\<^sub>R (x - c) \<in> interior s" |
| 53348 | 8153 |
proof - |
8154 |
obtain d where "d > 0" and d: "ball c d \<subseteq> s" |
|
8155 |
using assms(2) unfolding mem_interior by auto |
|
8156 |
show ?thesis |
|
8157 |
unfolding mem_interior |
|
8158 |
apply (rule_tac x="e*d" in exI) |
|
8159 |
apply rule |
|
8160 |
defer |
|
8161 |
unfolding subset_eq Ball_def mem_ball |
|
8162 |
proof (rule, rule) |
|
8163 |
fix y |
|
8164 |
assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" |
|
8165 |
have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" |
|
| 60420 | 8166 |
using \<open>e > 0\<close> by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) |
| 61945 | 8167 |
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)" |
| 53348 | 8168 |
unfolding dist_norm |
8169 |
unfolding norm_scaleR[symmetric] |
|
8170 |
apply (rule arg_cong[where f=norm]) |
|
| 60420 | 8171 |
using \<open>e > 0\<close> |
| 53348 | 8172 |
by (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) |
| 61945 | 8173 |
also have "\<dots> = \<bar>1/e\<bar> * norm (x - e *\<^sub>R (x - c) - y)" |
| 53348 | 8174 |
by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) |
8175 |
also have "\<dots> < d" |
|
| 60420 | 8176 |
using as[unfolded dist_norm] and \<open>e > 0\<close> |
8177 |
by (auto simp add:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute) |
|
| 53348 | 8178 |
finally show "y \<in> s" |
8179 |
apply (subst *) |
|
8180 |
apply (rule assms(1)[unfolded convex_alt,rule_format]) |
|
8181 |
apply (rule d[unfolded subset_eq,rule_format]) |
|
8182 |
unfolding mem_ball |
|
8183 |
using assms(3-5) |
|
8184 |
apply auto |
|
8185 |
done |
|
| 60420 | 8186 |
qed (insert \<open>e>0\<close> \<open>d>0\<close>, auto) |
| 53348 | 8187 |
qed |
| 33175 | 8188 |
|
8189 |
lemma mem_interior_closure_convex_shrink: |
|
| 53348 | 8190 |
fixes s :: "'a::euclidean_space set" |
8191 |
assumes "convex s" |
|
8192 |
and "c \<in> interior s" |
|
8193 |
and "x \<in> closure s" |
|
8194 |
and "0 < e" |
|
8195 |
and "e \<le> 1" |
|
| 33175 | 8196 |
shows "x - e *\<^sub>R (x - c) \<in> interior s" |
| 53348 | 8197 |
proof - |
8198 |
obtain d where "d > 0" and d: "ball c d \<subseteq> s" |
|
8199 |
using assms(2) unfolding mem_interior by auto |
|
8200 |
have "\<exists>y\<in>s. norm (y - x) * (1 - e) < e * d" |
|
8201 |
proof (cases "x \<in> s") |
|
8202 |
case True |
|
8203 |
then show ?thesis |
|
| 60420 | 8204 |
using \<open>e > 0\<close> \<open>d > 0\<close> |
| 53348 | 8205 |
apply (rule_tac bexI[where x=x]) |
| 56544 | 8206 |
apply (auto) |
| 53348 | 8207 |
done |
8208 |
next |
|
8209 |
case False |
|
8210 |
then have x: "x islimpt s" |
|
8211 |
using assms(3)[unfolded closure_def] by auto |
|
8212 |
show ?thesis |
|
8213 |
proof (cases "e = 1") |
|
8214 |
case True |
|
8215 |
obtain y where "y \<in> s" "y \<noteq> x" "dist y x < 1" |
|
| 33175 | 8216 |
using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto |
| 53348 | 8217 |
then show ?thesis |
8218 |
apply (rule_tac x=y in bexI) |
|
8219 |
unfolding True |
|
| 60420 | 8220 |
using \<open>d > 0\<close> |
| 53348 | 8221 |
apply auto |
8222 |
done |
|
8223 |
next |
|
8224 |
case False |
|
8225 |
then have "0 < e * d / (1 - e)" and *: "1 - e > 0" |
|
| 60420 | 8226 |
using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by auto |
| 53348 | 8227 |
then obtain y where "y \<in> s" "y \<noteq> x" "dist y x < e * d / (1 - e)" |
| 33175 | 8228 |
using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto |
| 53348 | 8229 |
then show ?thesis |
8230 |
apply (rule_tac x=y in bexI) |
|
8231 |
unfolding dist_norm |
|
8232 |
using pos_less_divide_eq[OF *] |
|
8233 |
apply auto |
|
8234 |
done |
|
8235 |
qed |
|
8236 |
qed |
|
8237 |
then obtain y where "y \<in> s" and y: "norm (y - x) * (1 - e) < e * d" |
|
8238 |
by auto |
|
| 63040 | 8239 |
define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" |
| 53348 | 8240 |
have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" |
| 60420 | 8241 |
unfolding z_def using \<open>e > 0\<close> |
| 53348 | 8242 |
by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) |
8243 |
have "z \<in> interior s" |
|
8244 |
apply (rule interior_mono[OF d,unfolded subset_eq,rule_format]) |
|
| 33175 | 8245 |
unfolding interior_open[OF open_ball] mem_ball z_def dist_norm using y and assms(4,5) |
| 53348 | 8246 |
apply (auto simp add:field_simps norm_minus_commute) |
8247 |
done |
|
8248 |
then show ?thesis |
|
8249 |
unfolding * |
|
8250 |
apply - |
|
8251 |
apply (rule mem_interior_convex_shrink) |
|
| 60420 | 8252 |
using assms(1,4-5) \<open>y\<in>s\<close> |
| 53348 | 8253 |
apply auto |
8254 |
done |
|
8255 |
qed |
|
8256 |
||
| 63128 | 8257 |
lemma in_interior_closure_convex_segment: |
8258 |
fixes S :: "'a::euclidean_space set" |
|
8259 |
assumes "convex S" and a: "a \<in> interior S" and b: "b \<in> closure S" |
|
8260 |
shows "open_segment a b \<subseteq> interior S" |
|
8261 |
proof (clarsimp simp: in_segment) |
|
8262 |
fix u::real |
|
8263 |
assume u: "0 < u" "u < 1" |
|
8264 |
have "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" |
|
8265 |
by (simp add: algebra_simps) |
|
8266 |
also have "... \<in> interior S" using mem_interior_closure_convex_shrink [OF assms] u |
|
8267 |
by simp |
|
8268 |
finally show "(1 - u) *\<^sub>R a + u *\<^sub>R b \<in> interior S" . |
|
8269 |
qed |
|
8270 |
||
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8271 |
lemma closure_open_Int_superset: |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8272 |
assumes "open S" "S \<subseteq> closure T" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8273 |
shows "closure(S \<inter> T) = closure S" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8274 |
proof - |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8275 |
have "closure S \<subseteq> closure(S \<inter> T)" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8276 |
by (metis assms closed_closure closure_minimal inf.orderE open_Int_closure_subset) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8277 |
then show ?thesis |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8278 |
by (simp add: closure_mono dual_order.antisym) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8279 |
qed |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8280 |
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8281 |
lemma convex_closure_interior: |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8282 |
fixes S :: "'a::euclidean_space set" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8283 |
assumes "convex S" and int: "interior S \<noteq> {}"
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8284 |
shows "closure(interior S) = closure S" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8285 |
proof - |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8286 |
obtain a where a: "a \<in> interior S" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8287 |
using int by auto |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8288 |
have "closure S \<subseteq> closure(interior S)" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8289 |
proof |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8290 |
fix x |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8291 |
assume x: "x \<in> closure S" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8292 |
show "x \<in> closure (interior S)" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8293 |
proof (cases "x=a") |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8294 |
case True |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8295 |
then show ?thesis |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8296 |
using \<open>a \<in> interior S\<close> closure_subset by blast |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8297 |
next |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8298 |
case False |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8299 |
show ?thesis |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8300 |
proof (clarsimp simp add: closure_def islimpt_approachable) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8301 |
fix e::real |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8302 |
assume xnotS: "x \<notin> interior S" and "0 < e" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8303 |
show "\<exists>x'\<in>interior S. x' \<noteq> x \<and> dist x' x < e" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8304 |
proof (intro bexI conjI) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8305 |
show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \<noteq> x" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8306 |
using False \<open>0 < e\<close> by (auto simp: algebra_simps min_def) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8307 |
show "dist (x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a)) x < e" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8308 |
using \<open>0 < e\<close> by (auto simp: dist_norm min_def) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8309 |
show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \<in> interior S" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8310 |
apply (clarsimp simp add: min_def a) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8311 |
apply (rule mem_interior_closure_convex_shrink [OF \<open>convex S\<close> a x]) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8312 |
using \<open>0 < e\<close> False apply (auto simp: divide_simps) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8313 |
done |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8314 |
qed |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8315 |
qed |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8316 |
qed |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8317 |
qed |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8318 |
then show ?thesis |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8319 |
by (simp add: closure_mono interior_subset subset_antisym) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8320 |
qed |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8321 |
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8322 |
lemma closure_convex_Int_superset: |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8323 |
fixes S :: "'a::euclidean_space set" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8324 |
assumes "convex S" "interior S \<noteq> {}" "interior S \<subseteq> closure T"
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8325 |
shows "closure(S \<inter> T) = closure S" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8326 |
proof - |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8327 |
have "closure S \<subseteq> closure(interior S)" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8328 |
by (simp add: convex_closure_interior assms) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8329 |
also have "... \<subseteq> closure (S \<inter> T)" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8330 |
using interior_subset [of S] assms |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8331 |
by (metis (no_types, lifting) Int_assoc Int_lower2 closure_mono closure_open_Int_superset inf.orderE open_interior) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8332 |
finally show ?thesis |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8333 |
by (simp add: closure_mono dual_order.antisym) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8334 |
qed |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
8335 |
|
| 33175 | 8336 |
|
| 60420 | 8337 |
subsection \<open>Some obvious but surprisingly hard simplex lemmas\<close> |
| 33175 | 8338 |
|
8339 |
lemma simplex: |
|
| 53348 | 8340 |
assumes "finite s" |
8341 |
and "0 \<notin> s" |
|
8342 |
shows "convex hull (insert 0 s) = |
|
| 64267 | 8343 |
{y. (\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s \<le> 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y)}"
|
| 53348 | 8344 |
unfolding convex_hull_finite[OF finite.insertI[OF assms(1)]] |
8345 |
apply (rule set_eqI, rule) |
|
8346 |
unfolding mem_Collect_eq |
|
8347 |
apply (erule_tac[!] exE) |
|
8348 |
apply (erule_tac[!] conjE)+ |
|
| 64267 | 8349 |
unfolding sum_clauses(2)[OF \<open>finite s\<close>] |
| 53348 | 8350 |
apply (rule_tac x=u in exI) |
8351 |
defer |
|
| 64267 | 8352 |
apply (rule_tac x="\<lambda>x. if x = 0 then 1 - sum u s else u x" in exI) |
| 53348 | 8353 |
using assms(2) |
| 64267 | 8354 |
unfolding if_smult and sum_delta_notmem[OF assms(2)] |
| 53348 | 8355 |
apply auto |
8356 |
done |
|
| 33175 | 8357 |
|
|
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
|
8358 |
lemma substd_simplex: |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
8359 |
assumes d: "d \<subseteq> Basis" |
| 53348 | 8360 |
shows "convex hull (insert 0 d) = |
| 54465 | 8361 |
{x. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> (\<Sum>i\<in>d. x\<bullet>i) \<le> 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)}"
|
| 40377 | 8362 |
(is "convex hull (insert 0 ?p) = ?s") |
| 53348 | 8363 |
proof - |
8364 |
let ?D = d |
|
8365 |
have "0 \<notin> ?p" |
|
8366 |
using assms by (auto simp: image_def) |
|
8367 |
from d have "finite d" |
|
8368 |
by (blast intro: finite_subset finite_Basis) |
|
8369 |
show ?thesis |
|
| 60420 | 8370 |
unfolding simplex[OF \<open>finite d\<close> \<open>0 \<notin> ?p\<close>] |
| 53348 | 8371 |
apply (rule set_eqI) |
8372 |
unfolding mem_Collect_eq |
|
8373 |
apply rule |
|
8374 |
apply (elim exE conjE) |
|
8375 |
apply (erule_tac[2] conjE)+ |
|
8376 |
proof - |
|
8377 |
fix x :: "'a::euclidean_space" |
|
8378 |
fix u |
|
| 64267 | 8379 |
assume as: "\<forall>x\<in>?D. 0 \<le> u x" "sum u ?D \<le> 1" "(\<Sum>x\<in>?D. u x *\<^sub>R x) = x" |
| 53348 | 8380 |
have *: "\<forall>i\<in>Basis. i:d \<longrightarrow> u i = x\<bullet>i" |
| 54465 | 8381 |
and "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)" |
| 53348 | 8382 |
using as(3) |
8383 |
unfolding substdbasis_expansion_unique[OF assms] |
|
8384 |
by auto |
|
| 64267 | 8385 |
then have **: "sum u ?D = sum (op \<bullet> x) ?D" |
| 53348 | 8386 |
apply - |
| 64267 | 8387 |
apply (rule sum.cong) |
| 53348 | 8388 |
using assms |
8389 |
apply auto |
|
8390 |
done |
|
| 64267 | 8391 |
have "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> sum (op \<bullet> x) ?D \<le> 1" |
| 53348 | 8392 |
proof (rule,rule) |
8393 |
fix i :: 'a |
|
8394 |
assume i: "i \<in> Basis" |
|
8395 |
have "i \<in> d \<Longrightarrow> 0 \<le> x\<bullet>i" |
|
8396 |
unfolding *[rule_format,OF i,symmetric] |
|
8397 |
apply (rule_tac as(1)[rule_format]) |
|
8398 |
apply auto |
|
8399 |
done |
|
8400 |
moreover have "i \<notin> d \<Longrightarrow> 0 \<le> x\<bullet>i" |
|
| 60420 | 8401 |
using \<open>(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)\<close>[rule_format, OF i] 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
|
8402 |
ultimately show "0 \<le> x\<bullet>i" by auto |
| 53348 | 8403 |
qed (insert as(2)[unfolded **], auto) |
| 64267 | 8404 |
then show "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> sum (op \<bullet> x) ?D \<le> 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)" |
| 60420 | 8405 |
using \<open>(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)\<close> by auto |
| 53348 | 8406 |
next |
8407 |
fix x :: "'a::euclidean_space" |
|
| 64267 | 8408 |
assume as: "\<forall>i\<in>Basis. 0 \<le> x \<bullet> i" "sum (op \<bullet> x) ?D \<le> 1" "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)" |
8409 |
show "\<exists>u. (\<forall>x\<in>?D. 0 \<le> u x) \<and> sum u ?D \<le> 1 \<and> (\<Sum>x\<in>?D. u x *\<^sub>R x) = x" |
|
| 53348 | 8410 |
using as d |
8411 |
unfolding substdbasis_expansion_unique[OF assms] |
|
8412 |
apply (rule_tac x="inner x" in exI) |
|
8413 |
apply auto |
|
8414 |
done |
|
|
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
|
8415 |
qed |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50104
diff
changeset
|
8416 |
qed |
| 40377 | 8417 |
|
| 33175 | 8418 |
lemma std_simplex: |
|
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
|
8419 |
"convex hull (insert 0 Basis) = |
| 64267 | 8420 |
{x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> sum (\<lambda>i. x\<bullet>i) Basis \<le> 1}"
|
|
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
|
8421 |
using substd_simplex[of Basis] by auto |
| 33175 | 8422 |
|
8423 |
lemma interior_std_simplex: |
|
|
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
|
8424 |
"interior (convex hull (insert 0 Basis)) = |
| 64267 | 8425 |
{x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 < x\<bullet>i) \<and> sum (\<lambda>i. x\<bullet>i) Basis < 1}"
|
| 53348 | 8426 |
apply (rule set_eqI) |
8427 |
unfolding mem_interior std_simplex |
|
8428 |
unfolding subset_eq mem_Collect_eq Ball_def mem_ball |
|
8429 |
unfolding Ball_def[symmetric] |
|
8430 |
apply rule |
|
8431 |
apply (elim exE conjE) |
|
8432 |
defer |
|
8433 |
apply (erule conjE) |
|
8434 |
proof - |
|
8435 |
fix x :: 'a |
|
8436 |
fix e |
|
| 64267 | 8437 |
assume "e > 0" and as: "\<forall>xa. dist x xa < e \<longrightarrow> (\<forall>x\<in>Basis. 0 \<le> xa \<bullet> x) \<and> sum (op \<bullet> xa) Basis \<le> 1" |
8438 |
show "(\<forall>xa\<in>Basis. 0 < x \<bullet> xa) \<and> sum (op \<bullet> x) Basis < 1" |
|
| 53348 | 8439 |
apply safe |
8440 |
proof - |
|
8441 |
fix i :: 'a |
|
8442 |
assume i: "i \<in> Basis" |
|
8443 |
then show "0 < x \<bullet> i" |
|
| 60420 | 8444 |
using as[THEN spec[where x="x - (e / 2) *\<^sub>R i"]] and \<open>e > 0\<close> |
| 53348 | 8445 |
unfolding dist_norm |
8446 |
by (auto elim!: ballE[where x=i] simp: inner_simps) |
|
8447 |
next |
|
| 60420 | 8448 |
have **: "dist x (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) < e" using \<open>e > 0\<close> |
|
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
|
8449 |
unfolding dist_norm |
| 53348 | 8450 |
by (auto intro!: mult_strict_left_mono simp: SOME_Basis) |
8451 |
have "\<And>i. i \<in> Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) \<bullet> i = |
|
8452 |
x\<bullet>i + (if i = (SOME i. i\<in>Basis) then e/2 else 0)" |
|
|
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
|
8453 |
by (auto simp: SOME_Basis inner_Basis inner_simps) |
| 64267 | 8454 |
then have *: "sum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis = |
8455 |
sum (\<lambda>i. x\<bullet>i + (if (SOME i. i\<in>Basis) = i then e/2 else 0)) Basis" |
|
8456 |
apply (rule_tac sum.cong) |
|
| 53348 | 8457 |
apply auto |
8458 |
done |
|
| 64267 | 8459 |
have "sum (op \<bullet> x) Basis < sum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis" |
8460 |
unfolding * sum.distrib |
|
| 60420 | 8461 |
using \<open>e > 0\<close> DIM_positive[where 'a='a] |
| 64267 | 8462 |
apply (subst sum.delta') |
| 53348 | 8463 |
apply (auto simp: SOME_Basis) |
8464 |
done |
|
8465 |
also have "\<dots> \<le> 1" |
|
8466 |
using ** |
|
8467 |
apply (drule_tac as[rule_format]) |
|
8468 |
apply auto |
|
8469 |
done |
|
| 64267 | 8470 |
finally show "sum (op \<bullet> x) Basis < 1" by auto |
| 53348 | 8471 |
qed |
8472 |
next |
|
8473 |
fix x :: 'a |
|
| 64267 | 8474 |
assume as: "\<forall>i\<in>Basis. 0 < x \<bullet> i" "sum (op \<bullet> x) Basis < 1" |
| 55697 | 8475 |
obtain a :: 'b where "a \<in> UNIV" using UNIV_witness .. |
| 64267 | 8476 |
let ?d = "(1 - sum (op \<bullet> x) Basis) / real (DIM('a))"
|
| 53348 | 8477 |
have "Min ((op \<bullet> x) ` Basis) > 0" |
8478 |
apply (rule Min_grI) |
|
8479 |
using as(1) |
|
8480 |
apply auto |
|
8481 |
done |
|
8482 |
moreover have "?d > 0" |
|
| 56541 | 8483 |
using as(2) by (auto simp: Suc_le_eq DIM_positive) |
| 64267 | 8484 |
ultimately show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> (\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> sum (op \<bullet> y) Basis \<le> 1" |
| 59807 | 8485 |
apply (rule_tac x="min (Min ((op \<bullet> x) ` Basis)) D" for D in exI) |
| 53348 | 8486 |
apply rule |
8487 |
defer |
|
8488 |
apply (rule, rule) |
|
8489 |
proof - |
|
8490 |
fix y |
|
8491 |
assume y: "dist x y < min (Min (op \<bullet> x ` Basis)) ?d" |
|
| 64267 | 8492 |
have "sum (op \<bullet> y) Basis \<le> sum (\<lambda>i. x\<bullet>i + ?d) Basis" |
8493 |
proof (rule sum_mono) |
|
| 53348 | 8494 |
fix i :: 'a |
8495 |
assume i: "i \<in> Basis" |
|
| 61945 | 8496 |
then have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d" |
| 53348 | 8497 |
apply - |
8498 |
apply (rule le_less_trans) |
|
|
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
|
8499 |
using Basis_le_norm[OF i, of "y - x"] |
| 53348 | 8500 |
using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] |
8501 |
apply (auto simp add: norm_minus_commute inner_diff_left) |
|
8502 |
done |
|
8503 |
then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto |
|
8504 |
qed |
|
8505 |
also have "\<dots> \<le> 1" |
|
| 64267 | 8506 |
unfolding sum.distrib sum_constant |
| 53348 | 8507 |
by (auto simp add: Suc_le_eq) |
| 64267 | 8508 |
finally show "(\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> sum (op \<bullet> y) Basis \<le> 1" |
| 53348 | 8509 |
proof safe |
8510 |
fix i :: 'a |
|
8511 |
assume i: "i \<in> Basis" |
|
8512 |
have "norm (x - y) < x\<bullet>i" |
|
8513 |
apply (rule less_le_trans) |
|
8514 |
apply (rule y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]) |
|
8515 |
using i |
|
8516 |
apply auto |
|
8517 |
done |
|
8518 |
then show "0 \<le> y\<bullet>i" |
|
8519 |
using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format, OF i] |
|
|
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
|
8520 |
by (auto simp: inner_simps) |
| 53348 | 8521 |
qed |
8522 |
qed auto |
|
8523 |
qed |
|
8524 |
||
8525 |
lemma interior_std_simplex_nonempty: |
|
8526 |
obtains a :: "'a::euclidean_space" where |
|
8527 |
"a \<in> interior(convex hull (insert 0 Basis))" |
|
8528 |
proof - |
|
8529 |
let ?D = "Basis :: 'a set" |
|
| 64267 | 8530 |
let ?a = "sum (\<lambda>b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis"
|
| 53348 | 8531 |
{
|
8532 |
fix i :: 'a |
|
8533 |
assume i: "i \<in> Basis" |
|
8534 |
have "?a \<bullet> i = inverse (2 * real DIM('a))"
|
|
| 64267 | 8535 |
by (rule trans[of _ "sum (\<lambda>j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"])
|
8536 |
(simp_all add: sum.If_cases i) } |
|
| 33175 | 8537 |
note ** = this |
| 53348 | 8538 |
show ?thesis |
8539 |
apply (rule that[of ?a]) |
|
8540 |
unfolding interior_std_simplex mem_Collect_eq |
|
8541 |
proof safe |
|
8542 |
fix i :: 'a |
|
8543 |
assume i: "i \<in> Basis" |
|
8544 |
show "0 < ?a \<bullet> i" |
|
8545 |
unfolding **[OF i] by (auto simp add: Suc_le_eq DIM_positive) |
|
8546 |
next |
|
| 64267 | 8547 |
have "sum (op \<bullet> ?a) ?D = sum (\<lambda>i. inverse (2 * real DIM('a))) ?D"
|
8548 |
apply (rule sum.cong) |
|
| 57418 | 8549 |
apply rule |
| 53348 | 8550 |
apply auto |
8551 |
done |
|
8552 |
also have "\<dots> < 1" |
|
| 64267 | 8553 |
unfolding sum_constant divide_inverse[symmetric] |
| 53348 | 8554 |
by (auto simp add: field_simps) |
| 64267 | 8555 |
finally show "sum (op \<bullet> ?a) ?D < 1" by auto |
| 40377 | 8556 |
qed |
| 53348 | 8557 |
qed |
8558 |
||
8559 |
lemma rel_interior_substd_simplex: |
|
8560 |
assumes d: "d \<subseteq> Basis" |
|
8561 |
shows "rel_interior (convex hull (insert 0 d)) = |
|
8562 |
{x::'a::euclidean_space. (\<forall>i\<in>d. 0 < x\<bullet>i) \<and> (\<Sum>i\<in>d. x\<bullet>i) < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)}"
|
|
8563 |
(is "rel_interior (convex hull (insert 0 ?p)) = ?s") |
|
8564 |
proof - |
|
8565 |
have "finite d" |
|
8566 |
apply (rule finite_subset) |
|
8567 |
using assms |
|
8568 |
apply auto |
|
8569 |
done |
|
8570 |
show ?thesis |
|
8571 |
proof (cases "d = {}")
|
|
8572 |
case True |
|
8573 |
then show ?thesis |
|
8574 |
using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto |
|
8575 |
next |
|
8576 |
case False |
|
8577 |
have h0: "affine hull (convex hull (insert 0 ?p)) = |
|
8578 |
{x::'a::euclidean_space. (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)}"
|
|
8579 |
using affine_hull_convex_hull affine_hull_substd_basis assms by auto |
|
8580 |
have aux: "\<And>x::'a. \<forall>i\<in>Basis. (\<forall>i\<in>d. 0 \<le> x\<bullet>i) \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0) \<longrightarrow> 0 \<le> x\<bullet>i" |
|
8581 |
by auto |
|
8582 |
{
|
|
8583 |
fix x :: "'a::euclidean_space" |
|
8584 |
assume x: "x \<in> rel_interior (convex hull (insert 0 ?p))" |
|
8585 |
then obtain e where e0: "e > 0" and |
|
8586 |
"ball x e \<inter> {xa. (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> xa\<bullet>i = 0)} \<subseteq> convex hull (insert 0 ?p)"
|
|
8587 |
using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto |
|
8588 |
then have as: "\<forall>xa. dist x xa < e \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> xa\<bullet>i = 0) \<longrightarrow> |
|
| 64267 | 8589 |
(\<forall>i\<in>d. 0 \<le> xa \<bullet> i) \<and> sum (op \<bullet> xa) d \<le> 1" |
| 53348 | 8590 |
unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto |
8591 |
have x0: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)" |
|
8592 |
using x rel_interior_subset substd_simplex[OF assms] by auto |
|
| 64267 | 8593 |
have "(\<forall>i\<in>d. 0 < x \<bullet> i) \<and> sum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)" |
| 54465 | 8594 |
apply rule |
8595 |
apply rule |
|
| 53348 | 8596 |
proof - |
8597 |
fix i :: 'a |
|
8598 |
assume "i \<in> d" |
|
8599 |
then have "\<forall>ia\<in>d. 0 \<le> (x - (e / 2) *\<^sub>R i) \<bullet> ia" |
|
8600 |
apply - |
|
8601 |
apply (rule as[rule_format,THEN conjunct1]) |
|
8602 |
unfolding dist_norm |
|
| 60420 | 8603 |
using d \<open>e > 0\<close> x0 |
| 53348 | 8604 |
apply (auto simp: inner_simps inner_Basis) |
8605 |
done |
|
8606 |
then show "0 < x \<bullet> i" |
|
8607 |
apply (erule_tac x=i in ballE) |
|
| 60420 | 8608 |
using \<open>e > 0\<close> \<open>i \<in> d\<close> d |
| 53348 | 8609 |
apply (auto simp: inner_simps inner_Basis) |
8610 |
done |
|
8611 |
next |
|
8612 |
obtain a where a: "a \<in> d" |
|
| 60420 | 8613 |
using \<open>d \<noteq> {}\<close> by auto
|
| 53348 | 8614 |
then have **: "dist x (x + (e / 2) *\<^sub>R a) < e" |
| 60420 | 8615 |
using \<open>e > 0\<close> norm_Basis[of a] d |
| 53348 | 8616 |
unfolding dist_norm |
8617 |
by auto |
|
8618 |
have "\<And>i. i \<in> Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = x\<bullet>i + (if i = a then e/2 else 0)" |
|
8619 |
using a d by (auto simp: inner_simps inner_Basis) |
|
| 64267 | 8620 |
then have *: "sum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d = |
8621 |
sum (\<lambda>i. x\<bullet>i + (if a = i then e/2 else 0)) d" |
|
8622 |
using d by (intro sum.cong) auto |
|
| 53348 | 8623 |
have "a \<in> Basis" |
| 60420 | 8624 |
using \<open>a \<in> d\<close> d by auto |
| 53348 | 8625 |
then have h1: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = 0)" |
| 60420 | 8626 |
using x0 d \<open>a\<in>d\<close> by (auto simp add: inner_add_left inner_Basis) |
| 64267 | 8627 |
have "sum (op \<bullet> x) d < sum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d" |
8628 |
unfolding * sum.distrib |
|
| 60420 | 8629 |
using \<open>e > 0\<close> \<open>a \<in> d\<close> |
8630 |
using \<open>finite d\<close> |
|
| 64267 | 8631 |
by (auto simp add: sum.delta') |
| 53348 | 8632 |
also have "\<dots> \<le> 1" |
8633 |
using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"] |
|
8634 |
by auto |
|
| 64267 | 8635 |
finally show "sum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)" |
| 53348 | 8636 |
using x0 by auto |
8637 |
qed |
|
8638 |
} |
|
8639 |
moreover |
|
8640 |
{
|
|
8641 |
fix x :: "'a::euclidean_space" |
|
8642 |
assume as: "x \<in> ?s" |
|
8643 |
have "\<forall>i. 0 < x\<bullet>i \<or> 0 = x\<bullet>i \<longrightarrow> 0 \<le> x\<bullet>i" |
|
8644 |
by auto |
|
8645 |
moreover have "\<forall>i. i \<in> d \<or> i \<notin> d" by auto |
|
8646 |
ultimately |
|
| 54465 | 8647 |
have "\<forall>i. (\<forall>i\<in>d. 0 < x\<bullet>i) \<and> (\<forall>i. i \<notin> d \<longrightarrow> x\<bullet>i = 0) \<longrightarrow> 0 \<le> x\<bullet>i" |
| 53348 | 8648 |
by metis |
8649 |
then have h2: "x \<in> convex hull (insert 0 ?p)" |
|
8650 |
using as assms |
|
8651 |
unfolding substd_simplex[OF assms] by fastforce |
|
8652 |
obtain a where a: "a \<in> d" |
|
| 60420 | 8653 |
using \<open>d \<noteq> {}\<close> by auto
|
| 64267 | 8654 |
let ?d = "(1 - sum (op \<bullet> x) d) / real (card d)" |
| 60420 | 8655 |
have "0 < card d" using \<open>d \<noteq> {}\<close> \<open>finite d\<close>
|
| 44466 | 8656 |
by (simp add: card_gt_0_iff) |
| 53348 | 8657 |
have "Min ((op \<bullet> x) ` d) > 0" |
| 60420 | 8658 |
using as \<open>d \<noteq> {}\<close> \<open>finite d\<close> by (simp add: Min_grI)
|
8659 |
moreover have "?d > 0" using as using \<open>0 < card d\<close> by auto |
|
| 53348 | 8660 |
ultimately have h3: "min (Min ((op \<bullet> x) ` d)) ?d > 0" |
8661 |
by auto |
|
| 54465 | 8662 |
|
| 53348 | 8663 |
have "x \<in> rel_interior (convex hull (insert 0 ?p))" |
8664 |
unfolding rel_interior_ball mem_Collect_eq h0 |
|
8665 |
apply (rule,rule h2) |
|
8666 |
unfolding substd_simplex[OF assms] |
|
8667 |
apply (rule_tac x="min (Min ((op \<bullet> x) ` d)) ?d" in exI) |
|
8668 |
apply (rule, rule h3) |
|
8669 |
apply safe |
|
8670 |
unfolding mem_ball |
|
8671 |
proof - |
|
8672 |
fix y :: 'a |
|
8673 |
assume y: "dist x y < min (Min (op \<bullet> x ` d)) ?d" |
|
8674 |
assume y2: "\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> y\<bullet>i = 0" |
|
| 64267 | 8675 |
have "sum (op \<bullet> y) d \<le> sum (\<lambda>i. x\<bullet>i + ?d) d" |
8676 |
proof (rule sum_mono) |
|
| 53348 | 8677 |
fix i |
8678 |
assume "i \<in> d" |
|
8679 |
with d have i: "i \<in> Basis" |
|
8680 |
by auto |
|
| 61945 | 8681 |
have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d" |
| 53348 | 8682 |
apply (rule le_less_trans) |
8683 |
using Basis_le_norm[OF i, of "y - x"] |
|
8684 |
using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] |
|
8685 |
apply (auto simp add: norm_minus_commute inner_simps) |
|
8686 |
done |
|
8687 |
then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto |
|
8688 |
qed |
|
8689 |
also have "\<dots> \<le> 1" |
|
| 64267 | 8690 |
unfolding sum.distrib sum_constant using \<open>0 < card d\<close> |
| 53348 | 8691 |
by auto |
| 64267 | 8692 |
finally show "sum (op \<bullet> y) d \<le> 1" . |
| 54465 | 8693 |
|
| 53348 | 8694 |
fix i :: 'a |
8695 |
assume i: "i \<in> Basis" |
|
8696 |
then show "0 \<le> y\<bullet>i" |
|
8697 |
proof (cases "i\<in>d") |
|
8698 |
case True |
|
8699 |
have "norm (x - y) < x\<bullet>i" |
|
8700 |
using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1] |
|
| 60420 | 8701 |
using Min_gr_iff[of "op \<bullet> x ` d" "norm (x - y)"] \<open>0 < card d\<close> \<open>i:d\<close> |
| 53348 | 8702 |
by (simp add: card_gt_0_iff) |
8703 |
then show "0 \<le> y\<bullet>i" |
|
8704 |
using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format] |
|
8705 |
by (auto simp: inner_simps) |
|
8706 |
qed (insert y2, auto) |
|
8707 |
qed |
|
8708 |
} |
|
8709 |
ultimately have |
|
8710 |
"\<And>x. x \<in> rel_interior (convex hull insert 0 d) \<longleftrightarrow> |
|
| 64267 | 8711 |
x \<in> {x. (\<forall>i\<in>d. 0 < x \<bullet> i) \<and> sum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)}"
|
| 53348 | 8712 |
by blast |
8713 |
then show ?thesis by (rule set_eqI) |
|
| 40377 | 8714 |
qed |
| 53348 | 8715 |
qed |
8716 |
||
8717 |
lemma rel_interior_substd_simplex_nonempty: |
|
8718 |
assumes "d \<noteq> {}"
|
|
8719 |
and "d \<subseteq> Basis" |
|
8720 |
obtains a :: "'a::euclidean_space" |
|
8721 |
where "a \<in> rel_interior (convex hull (insert 0 d))" |
|
8722 |
proof - |
|
8723 |
let ?D = d |
|
| 64267 | 8724 |
let ?a = "sum (\<lambda>b::'a::euclidean_space. inverse (2 * real (card d)) *\<^sub>R b) ?D" |
| 53348 | 8725 |
have "finite d" |
8726 |
apply (rule finite_subset) |
|
8727 |
using assms(2) |
|
8728 |
apply auto |
|
8729 |
done |
|
8730 |
then have d1: "0 < real (card d)" |
|
| 60420 | 8731 |
using \<open>d \<noteq> {}\<close> by auto
|
| 53348 | 8732 |
{
|
8733 |
fix i |
|
8734 |
assume "i \<in> d" |
|
|
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
|
8735 |
have "?a \<bullet> i = inverse (2 * real (card d))" |
| 64267 | 8736 |
apply (rule trans[of _ "sum (\<lambda>j. if i = j then inverse (2 * real (card d)) else 0) ?D"]) |
8737 |
unfolding inner_sum_left |
|
8738 |
apply (rule sum.cong) |
|
8739 |
using \<open>i \<in> d\<close> \<open>finite d\<close> sum.delta'[of d i "(\<lambda>k. inverse (2 * real (card d)))"] |
|
| 53348 | 8740 |
d1 assms(2) |
| 57418 | 8741 |
by (auto simp: inner_Basis set_rev_mp[OF _ assms(2)]) |
| 53348 | 8742 |
} |
| 40377 | 8743 |
note ** = this |
| 53348 | 8744 |
show ?thesis |
8745 |
apply (rule that[of ?a]) |
|
8746 |
unfolding rel_interior_substd_simplex[OF assms(2)] mem_Collect_eq |
|
8747 |
proof safe |
|
8748 |
fix i |
|
8749 |
assume "i \<in> d" |
|
8750 |
have "0 < inverse (2 * real (card d))" |
|
8751 |
using d1 by auto |
|
| 60420 | 8752 |
also have "\<dots> = ?a \<bullet> i" using **[of i] \<open>i \<in> d\<close> |
| 53348 | 8753 |
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
|
8754 |
finally show "0 < ?a \<bullet> i" by auto |
| 53348 | 8755 |
next |
| 64267 | 8756 |
have "sum (op \<bullet> ?a) ?D = sum (\<lambda>i. inverse (2 * real (card d))) ?D" |
8757 |
by (rule sum.cong) (rule refl, rule **) |
|
| 53348 | 8758 |
also have "\<dots> < 1" |
| 64267 | 8759 |
unfolding sum_constant divide_real_def[symmetric] |
| 53348 | 8760 |
by (auto simp add: field_simps) |
| 64267 | 8761 |
finally show "sum (op \<bullet> ?a) ?D < 1" by auto |
| 53348 | 8762 |
next |
8763 |
fix i |
|
8764 |
assume "i \<in> Basis" and "i \<notin> d" |
|
8765 |
have "?a \<in> span d" |
|
| 64267 | 8766 |
proof (rule span_sum[of d "(\<lambda>b. b /\<^sub>R (2 * real (card d)))" d]) |
| 53348 | 8767 |
{
|
8768 |
fix x :: "'a::euclidean_space" |
|
8769 |
assume "x \<in> d" |
|
8770 |
then have "x \<in> span d" |
|
|
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
|
8771 |
using span_superset[of _ "d"] by auto |
| 53348 | 8772 |
then have "x /\<^sub>R (2 * real (card d)) \<in> span d" |
|
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
|
8773 |
using span_mul[of x "d" "(inverse (real (card d)) / 2)"] by auto |
| 53348 | 8774 |
} |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
8775 |
then show "\<And>x. x\<in>d \<Longrightarrow> x /\<^sub>R (2 * real (card d)) \<in> span d" |
| 53348 | 8776 |
by auto |
| 40377 | 8777 |
qed |
| 53348 | 8778 |
then show "?a \<bullet> i = 0 " |
| 60420 | 8779 |
using \<open>i \<notin> d\<close> unfolding span_substd_basis[OF assms(2)] using \<open>i \<in> Basis\<close> by auto |
| 40377 | 8780 |
qed |
8781 |
qed |
|
8782 |
||
| 53348 | 8783 |
|
| 60420 | 8784 |
subsection \<open>Relative interior of convex set\<close> |
| 40377 | 8785 |
|
| 49531 | 8786 |
lemma rel_interior_convex_nonempty_aux: |
| 53348 | 8787 |
fixes S :: "'n::euclidean_space set" |
8788 |
assumes "convex S" |
|
8789 |
and "0 \<in> S" |
|
8790 |
shows "rel_interior S \<noteq> {}"
|
|
8791 |
proof (cases "S = {0}")
|
|
8792 |
case True |
|
8793 |
then show ?thesis using rel_interior_sing by auto |
|
8794 |
next |
|
8795 |
case False |
|
8796 |
obtain B where B: "independent B \<and> B \<le> S \<and> S \<le> span B \<and> card B = dim S" |
|
8797 |
using basis_exists[of S] by auto |
|
8798 |
then have "B \<noteq> {}"
|
|
| 60420 | 8799 |
using B assms \<open>S \<noteq> {0}\<close> span_empty by auto
|
| 53348 | 8800 |
have "insert 0 B \<le> span B" |
8801 |
using subspace_span[of B] subspace_0[of "span B"] span_inc by auto |
|
8802 |
then have "span (insert 0 B) \<le> span B" |
|
| 40377 | 8803 |
using span_span[of B] span_mono[of "insert 0 B" "span B"] by blast |
| 53348 | 8804 |
then have "convex hull insert 0 B \<le> span B" |
| 40377 | 8805 |
using convex_hull_subset_span[of "insert 0 B"] by auto |
| 53348 | 8806 |
then have "span (convex hull insert 0 B) \<le> span B" |
| 40377 | 8807 |
using span_span[of B] span_mono[of "convex hull insert 0 B" "span B"] by blast |
| 53348 | 8808 |
then have *: "span (convex hull insert 0 B) = span B" |
| 40377 | 8809 |
using span_mono[of B "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto |
| 53348 | 8810 |
then have "span (convex hull insert 0 B) = span S" |
8811 |
using B span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto |
|
8812 |
moreover have "0 \<in> affine hull (convex hull insert 0 B)" |
|
| 40377 | 8813 |
using hull_subset[of "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto |
| 53348 | 8814 |
ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S" |
| 49531 | 8815 |
using affine_hull_span_0[of "convex hull insert 0 B"] affine_hull_span_0[of "S"] |
| 53348 | 8816 |
assms hull_subset[of S] |
8817 |
by auto |
|
8818 |
obtain d and f :: "'n \<Rightarrow> 'n" where |
|
8819 |
fd: "card d = card B" "linear f" "f ` B = d" |
|
8820 |
"f ` span B = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = (0::real)} \<and> inj_on f (span B)"
|
|
8821 |
and d: "d \<subseteq> Basis" |
|
8822 |
using basis_to_substdbasis_subspace_isomorphism[of B,OF _ ] B by auto |
|
8823 |
then have "bounded_linear f" |
|
8824 |
using linear_conv_bounded_linear by auto |
|
8825 |
have "d \<noteq> {}"
|
|
| 60420 | 8826 |
using fd B \<open>B \<noteq> {}\<close> by auto
|
| 53348 | 8827 |
have "insert 0 d = f ` (insert 0 B)" |
8828 |
using fd linear_0 by auto |
|
8829 |
then have "(convex hull (insert 0 d)) = f ` (convex hull (insert 0 B))" |
|
8830 |
using convex_hull_linear_image[of f "(insert 0 d)"] |
|
| 60420 | 8831 |
convex_hull_linear_image[of f "(insert 0 B)"] \<open>linear f\<close> |
| 53348 | 8832 |
by auto |
8833 |
moreover have "rel_interior (f ` (convex hull insert 0 B)) = |
|
8834 |
f ` rel_interior (convex hull insert 0 B)" |
|
8835 |
apply (rule rel_interior_injective_on_span_linear_image[of f "(convex hull insert 0 B)"]) |
|
| 60420 | 8836 |
using \<open>bounded_linear f\<close> fd * |
| 53348 | 8837 |
apply auto |
8838 |
done |
|
8839 |
ultimately have "rel_interior (convex hull insert 0 B) \<noteq> {}"
|
|
| 60420 | 8840 |
using rel_interior_substd_simplex_nonempty[OF \<open>d \<noteq> {}\<close> d]
|
| 53348 | 8841 |
apply auto |
8842 |
apply blast |
|
8843 |
done |
|
8844 |
moreover have "convex hull (insert 0 B) \<subseteq> S" |
|
8845 |
using B assms hull_mono[of "insert 0 B" "S" "convex"] convex_hull_eq |
|
8846 |
by auto |
|
8847 |
ultimately show ?thesis |
|
8848 |
using subset_rel_interior[of "convex hull insert 0 B" S] ** by auto |
|
| 40377 | 8849 |
qed |
8850 |
||
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
8851 |
lemma rel_interior_eq_empty: |
| 53348 | 8852 |
fixes S :: "'n::euclidean_space set" |
8853 |
assumes "convex S" |
|
8854 |
shows "rel_interior S = {} \<longleftrightarrow> S = {}"
|
|
8855 |
proof - |
|
8856 |
{
|
|
8857 |
assume "S \<noteq> {}"
|
|
8858 |
then obtain a where "a \<in> S" by auto |
|
8859 |
then have "0 \<in> op + (-a) ` S" |
|
8860 |
using assms exI[of "(\<lambda>x. x \<in> S \<and> - a + x = 0)" a] by auto |
|
8861 |
then have "rel_interior (op + (-a) ` S) \<noteq> {}"
|
|
8862 |
using rel_interior_convex_nonempty_aux[of "op + (-a) ` S"] |
|
8863 |
convex_translation[of S "-a"] assms |
|
8864 |
by auto |
|
8865 |
then have "rel_interior S \<noteq> {}"
|
|
8866 |
using rel_interior_translation by auto |
|
8867 |
} |
|
8868 |
then show ?thesis |
|
8869 |
using rel_interior_empty by auto |
|
| 40377 | 8870 |
qed |
8871 |
||
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8872 |
lemma interior_simplex_nonempty: |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8873 |
fixes S :: "'N :: euclidean_space set" |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8874 |
assumes "independent S" "finite S" "card S = DIM('N)"
|
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8875 |
obtains a where "a \<in> interior (convex hull (insert 0 S))" |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8876 |
proof - |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8877 |
have "affine hull (insert 0 S) = UNIV" |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8878 |
apply (simp add: hull_inc affine_hull_span_0) |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8879 |
using assms dim_eq_full indep_card_eq_dim_span by fastforce |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8880 |
moreover have "rel_interior (convex hull insert 0 S) \<noteq> {}"
|
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8881 |
using rel_interior_eq_empty [of "convex hull (insert 0 S)"] by auto |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8882 |
ultimately have "interior (convex hull insert 0 S) \<noteq> {}"
|
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8883 |
by (simp add: rel_interior_interior) |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8884 |
with that show ?thesis |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8885 |
by auto |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8886 |
qed |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
8887 |
|
| 40377 | 8888 |
lemma convex_rel_interior: |
| 53348 | 8889 |
fixes S :: "'n::euclidean_space set" |
8890 |
assumes "convex S" |
|
8891 |
shows "convex (rel_interior S)" |
|
8892 |
proof - |
|
8893 |
{
|
|
8894 |
fix x y and u :: real |
|
8895 |
assume assm: "x \<in> rel_interior S" "y \<in> rel_interior S" "0 \<le> u" "u \<le> 1" |
|
8896 |
then have "x \<in> S" |
|
8897 |
using rel_interior_subset by auto |
|
8898 |
have "x - u *\<^sub>R (x-y) \<in> rel_interior S" |
|
8899 |
proof (cases "0 = u") |
|
8900 |
case False |
|
8901 |
then have "0 < u" using assm by auto |
|
8902 |
then show ?thesis |
|
| 60420 | 8903 |
using assm rel_interior_convex_shrink[of S y x u] assms \<open>x \<in> S\<close> by auto |
| 53348 | 8904 |
next |
8905 |
case True |
|
8906 |
then show ?thesis using assm by auto |
|
8907 |
qed |
|
8908 |
then have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> rel_interior S" |
|
8909 |
by (simp add: algebra_simps) |
|
8910 |
} |
|
8911 |
then show ?thesis |
|
8912 |
unfolding convex_alt by auto |
|
| 40377 | 8913 |
qed |
8914 |
||
| 49531 | 8915 |
lemma convex_closure_rel_interior: |
| 53348 | 8916 |
fixes S :: "'n::euclidean_space set" |
8917 |
assumes "convex S" |
|
8918 |
shows "closure (rel_interior S) = closure S" |
|
8919 |
proof - |
|
8920 |
have h1: "closure (rel_interior S) \<le> closure S" |
|
8921 |
using closure_mono[of "rel_interior S" S] rel_interior_subset[of S] by auto |
|
8922 |
show ?thesis |
|
8923 |
proof (cases "S = {}")
|
|
8924 |
case False |
|
8925 |
then obtain a where a: "a \<in> rel_interior S" |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
8926 |
using rel_interior_eq_empty assms by auto |
| 53348 | 8927 |
{ fix x
|
8928 |
assume x: "x \<in> closure S" |
|
8929 |
{
|
|
8930 |
assume "x = a" |
|
8931 |
then have "x \<in> closure (rel_interior S)" |
|
8932 |
using a unfolding closure_def by auto |
|
8933 |
} |
|
8934 |
moreover |
|
8935 |
{
|
|
8936 |
assume "x \<noteq> a" |
|
8937 |
{
|
|
8938 |
fix e :: real |
|
8939 |
assume "e > 0" |
|
| 63040 | 8940 |
define e1 where "e1 = min 1 (e/norm (x - a))" |
| 53348 | 8941 |
then have e1: "e1 > 0" "e1 \<le> 1" "e1 * norm (x - a) \<le> e" |
| 60420 | 8942 |
using \<open>x \<noteq> a\<close> \<open>e > 0\<close> le_divide_eq[of e1 e "norm (x - a)"] |
| 53348 | 8943 |
by simp_all |
8944 |
then have *: "x - e1 *\<^sub>R (x - a) : rel_interior S" |
|
8945 |
using rel_interior_closure_convex_shrink[of S a x e1] assms x a e1_def |
|
8946 |
by auto |
|
8947 |
have "\<exists>y. y \<in> rel_interior S \<and> y \<noteq> x \<and> dist y x \<le> e" |
|
8948 |
apply (rule_tac x="x - e1 *\<^sub>R (x - a)" in exI) |
|
| 60420 | 8949 |
using * e1 dist_norm[of "x - e1 *\<^sub>R (x - a)" x] \<open>x \<noteq> a\<close> |
| 53348 | 8950 |
apply simp |
8951 |
done |
|
8952 |
} |
|
8953 |
then have "x islimpt rel_interior S" |
|
8954 |
unfolding islimpt_approachable_le by auto |
|
8955 |
then have "x \<in> closure(rel_interior S)" |
|
8956 |
unfolding closure_def by auto |
|
8957 |
} |
|
8958 |
ultimately have "x \<in> closure(rel_interior S)" by auto |
|
8959 |
} |
|
8960 |
then show ?thesis using h1 by auto |
|
8961 |
next |
|
8962 |
case True |
|
8963 |
then have "rel_interior S = {}"
|
|
8964 |
using rel_interior_empty by auto |
|
8965 |
then have "closure (rel_interior S) = {}"
|
|
8966 |
using closure_empty by auto |
|
8967 |
with True show ?thesis by auto |
|
8968 |
qed |
|
| 40377 | 8969 |
qed |
8970 |
||
8971 |
lemma rel_interior_same_affine_hull: |
|
| 53348 | 8972 |
fixes S :: "'n::euclidean_space set" |
| 40377 | 8973 |
assumes "convex S" |
8974 |
shows "affine hull (rel_interior S) = affine hull S" |
|
| 53348 | 8975 |
by (metis assms closure_same_affine_hull convex_closure_rel_interior) |
| 40377 | 8976 |
|
| 49531 | 8977 |
lemma rel_interior_aff_dim: |
| 53348 | 8978 |
fixes S :: "'n::euclidean_space set" |
| 40377 | 8979 |
assumes "convex S" |
8980 |
shows "aff_dim (rel_interior S) = aff_dim S" |
|
| 53348 | 8981 |
by (metis aff_dim_affine_hull2 assms rel_interior_same_affine_hull) |
| 40377 | 8982 |
|
8983 |
lemma rel_interior_rel_interior: |
|
| 53348 | 8984 |
fixes S :: "'n::euclidean_space set" |
| 40377 | 8985 |
assumes "convex S" |
8986 |
shows "rel_interior (rel_interior S) = rel_interior S" |
|
| 53348 | 8987 |
proof - |
8988 |
have "openin (subtopology euclidean (affine hull (rel_interior S))) (rel_interior S)" |
|
| 63072 | 8989 |
using openin_rel_interior[of S] rel_interior_same_affine_hull[of S] assms by auto |
| 53348 | 8990 |
then show ?thesis |
8991 |
using rel_interior_def by auto |
|
| 40377 | 8992 |
qed |
8993 |
||
8994 |
lemma rel_interior_rel_open: |
|
| 53348 | 8995 |
fixes S :: "'n::euclidean_space set" |
| 40377 | 8996 |
assumes "convex S" |
8997 |
shows "rel_open (rel_interior S)" |
|
| 53348 | 8998 |
unfolding rel_open_def using rel_interior_rel_interior assms by auto |
| 40377 | 8999 |
|
9000 |
lemma convex_rel_interior_closure_aux: |
|
| 53348 | 9001 |
fixes x y z :: "'n::euclidean_space" |
9002 |
assumes "0 < a" "0 < b" "(a + b) *\<^sub>R z = a *\<^sub>R x + b *\<^sub>R y" |
|
9003 |
obtains e where "0 < e" "e \<le> 1" "z = y - e *\<^sub>R (y - x)" |
|
9004 |
proof - |
|
| 63040 | 9005 |
define e where "e = a / (a + b)" |
| 53348 | 9006 |
have "z = (1 / (a + b)) *\<^sub>R ((a + b) *\<^sub>R z)" |
9007 |
apply auto |
|
9008 |
using assms |
|
9009 |
apply simp |
|
9010 |
done |
|
9011 |
also have "\<dots> = (1 / (a + b)) *\<^sub>R (a *\<^sub>R x + b *\<^sub>R y)" |
|
9012 |
using assms scaleR_cancel_left[of "1/(a+b)" "(a + b) *\<^sub>R z" "a *\<^sub>R x + b *\<^sub>R y"] |
|
9013 |
by auto |
|
9014 |
also have "\<dots> = y - e *\<^sub>R (y-x)" |
|
9015 |
using e_def |
|
9016 |
apply (simp add: algebra_simps) |
|
9017 |
using scaleR_left_distrib[of "a/(a+b)" "b/(a+b)" y] assms add_divide_distrib[of a b "a+b"] |
|
9018 |
apply auto |
|
9019 |
done |
|
9020 |
finally have "z = y - e *\<^sub>R (y-x)" |
|
9021 |
by auto |
|
| 56541 | 9022 |
moreover have "e > 0" using e_def assms by auto |
9023 |
moreover have "e \<le> 1" using e_def assms by auto |
|
9024 |
ultimately show ?thesis using that[of e] by auto |
|
| 40377 | 9025 |
qed |
9026 |
||
| 49531 | 9027 |
lemma convex_rel_interior_closure: |
| 53348 | 9028 |
fixes S :: "'n::euclidean_space set" |
| 40377 | 9029 |
assumes "convex S" |
9030 |
shows "rel_interior (closure S) = rel_interior S" |
|
| 53348 | 9031 |
proof (cases "S = {}")
|
9032 |
case True |
|
9033 |
then show ?thesis |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
9034 |
using assms rel_interior_eq_empty by auto |
| 53348 | 9035 |
next |
9036 |
case False |
|
9037 |
have "rel_interior (closure S) \<supseteq> rel_interior S" |
|
9038 |
using subset_rel_interior[of S "closure S"] closure_same_affine_hull closure_subset |
|
9039 |
by auto |
|
| 40377 | 9040 |
moreover |
| 53348 | 9041 |
{
|
9042 |
fix z |
|
| 54465 | 9043 |
assume z: "z \<in> rel_interior (closure S)" |
| 53348 | 9044 |
obtain x where x: "x \<in> rel_interior S" |
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
9045 |
using \<open>S \<noteq> {}\<close> assms rel_interior_eq_empty by auto
|
| 53348 | 9046 |
have "z \<in> rel_interior S" |
9047 |
proof (cases "x = z") |
|
9048 |
case True |
|
9049 |
then show ?thesis using x by auto |
|
9050 |
next |
|
9051 |
case False |
|
| 54465 | 9052 |
obtain e where e: "e > 0" "cball z e \<inter> affine hull closure S \<le> closure S" |
| 53348 | 9053 |
using z rel_interior_cball[of "closure S"] by auto |
| 56541 | 9054 |
hence *: "0 < e/norm(z-x)" using e False by auto |
| 63040 | 9055 |
define y where "y = z + (e/norm(z-x)) *\<^sub>R (z-x)" |
| 53348 | 9056 |
have yball: "y \<in> cball z e" |
9057 |
using mem_cball y_def dist_norm[of z y] e by auto |
|
9058 |
have "x \<in> affine hull closure S" |
|
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
9059 |
using x rel_interior_subset_closure hull_inc[of x "closure S"] by blast |
| 53348 | 9060 |
moreover have "z \<in> affine hull closure S" |
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
9061 |
using z rel_interior_subset hull_subset[of "closure S"] by blast |
| 53348 | 9062 |
ultimately have "y \<in> affine hull closure S" |
| 49531 | 9063 |
using y_def affine_affine_hull[of "closure S"] |
| 40377 | 9064 |
mem_affine_3_minus [of "affine hull closure S" z z x "e/norm(z-x)"] by auto |
| 53348 | 9065 |
then have "y \<in> closure S" using e yball by auto |
9066 |
have "(1 + (e/norm(z-x))) *\<^sub>R z = (e/norm(z-x)) *\<^sub>R x + y" |
|
| 49531 | 9067 |
using y_def by (simp add: algebra_simps) |
| 53348 | 9068 |
then obtain e1 where "0 < e1" "e1 \<le> 1" "z = y - e1 *\<^sub>R (y - x)" |
| 49531 | 9069 |
using * convex_rel_interior_closure_aux[of "e / norm (z - x)" 1 z x y] |
| 53348 | 9070 |
by (auto simp add: algebra_simps) |
9071 |
then show ?thesis |
|
| 60420 | 9072 |
using rel_interior_closure_convex_shrink assms x \<open>y \<in> closure S\<close> |
| 53348 | 9073 |
by auto |
9074 |
qed |
|
9075 |
} |
|
9076 |
ultimately show ?thesis by auto |
|
| 40377 | 9077 |
qed |
9078 |
||
| 49531 | 9079 |
lemma convex_interior_closure: |
| 53348 | 9080 |
fixes S :: "'n::euclidean_space set" |
9081 |
assumes "convex S" |
|
9082 |
shows "interior (closure S) = interior S" |
|
9083 |
using closure_aff_dim[of S] interior_rel_interior_gen[of S] |
|
9084 |
interior_rel_interior_gen[of "closure S"] |
|
9085 |
convex_rel_interior_closure[of S] assms |
|
9086 |
by auto |
|
| 40377 | 9087 |
|
9088 |
lemma closure_eq_rel_interior_eq: |
|
| 53348 | 9089 |
fixes S1 S2 :: "'n::euclidean_space set" |
9090 |
assumes "convex S1" |
|
9091 |
and "convex S2" |
|
9092 |
shows "closure S1 = closure S2 \<longleftrightarrow> rel_interior S1 = rel_interior S2" |
|
9093 |
by (metis convex_rel_interior_closure convex_closure_rel_interior assms) |
|
| 40377 | 9094 |
|
9095 |
lemma closure_eq_between: |
|
| 53348 | 9096 |
fixes S1 S2 :: "'n::euclidean_space set" |
9097 |
assumes "convex S1" |
|
9098 |
and "convex S2" |
|
9099 |
shows "closure S1 = closure S2 \<longleftrightarrow> rel_interior S1 \<le> S2 \<and> S2 \<subseteq> closure S1" |
|
| 54465 | 9100 |
(is "?A \<longleftrightarrow> ?B") |
| 53348 | 9101 |
proof |
9102 |
assume ?A |
|
9103 |
then show ?B |
|
9104 |
by (metis assms closure_subset convex_rel_interior_closure rel_interior_subset) |
|
9105 |
next |
|
9106 |
assume ?B |
|
9107 |
then have "closure S1 \<subseteq> closure S2" |
|
9108 |
by (metis assms(1) convex_closure_rel_interior closure_mono) |
|
| 60420 | 9109 |
moreover from \<open>?B\<close> have "closure S1 \<supseteq> closure S2" |
| 53348 | 9110 |
by (metis closed_closure closure_minimal) |
9111 |
ultimately show ?A .. |
|
| 40377 | 9112 |
qed |
9113 |
||
9114 |
lemma open_inter_closure_rel_interior: |
|
| 53348 | 9115 |
fixes S A :: "'n::euclidean_space set" |
9116 |
assumes "convex S" |
|
9117 |
and "open A" |
|
9118 |
shows "A \<inter> closure S = {} \<longleftrightarrow> A \<inter> rel_interior S = {}"
|
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62626
diff
changeset
|
9119 |
by (metis assms convex_closure_rel_interior open_Int_closure_eq_empty) |
| 40377 | 9120 |
|
|
62620
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9121 |
lemma rel_interior_open_segment: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9122 |
fixes a :: "'a :: euclidean_space" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9123 |
shows "rel_interior(open_segment a b) = open_segment a b" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9124 |
proof (cases "a = b") |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9125 |
case True then show ?thesis by auto |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9126 |
next |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9127 |
case False then show ?thesis |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9128 |
apply (simp add: rel_interior_eq openin_open) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9129 |
apply (rule_tac x="ball (inverse 2 *\<^sub>R (a + b)) (norm(b - a) / 2)" in exI) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9130 |
apply (simp add: open_segment_as_ball) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9131 |
done |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9132 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9133 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9134 |
lemma rel_interior_closed_segment: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9135 |
fixes a :: "'a :: euclidean_space" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9136 |
shows "rel_interior(closed_segment a b) = |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9137 |
(if a = b then {a} else open_segment a b)"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9138 |
proof (cases "a = b") |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9139 |
case True then show ?thesis by auto |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9140 |
next |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9141 |
case False then show ?thesis |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9142 |
by simp |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9143 |
(metis closure_open_segment convex_open_segment convex_rel_interior_closure |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9144 |
rel_interior_open_segment) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9145 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9146 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9147 |
lemmas rel_interior_segment = rel_interior_closed_segment rel_interior_open_segment |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9148 |
|
|
62626
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9149 |
lemma starlike_convex_tweak_boundary_points: |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9150 |
fixes S :: "'a::euclidean_space set" |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9151 |
assumes "convex S" "S \<noteq> {}" and ST: "rel_interior S \<subseteq> T" and TS: "T \<subseteq> closure S"
|
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9152 |
shows "starlike T" |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9153 |
proof - |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9154 |
have "rel_interior S \<noteq> {}"
|
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9155 |
by (simp add: assms rel_interior_eq_empty) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
9156 |
then obtain a where a: "a \<in> rel_interior S" by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
9157 |
with ST have "a \<in> T" by blast |
|
62626
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9158 |
have *: "\<And>x. x \<in> T \<Longrightarrow> open_segment a x \<subseteq> rel_interior S" |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9159 |
apply (rule rel_interior_closure_convex_segment [OF \<open>convex S\<close> a]) |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9160 |
using assms by blast |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9161 |
show ?thesis |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9162 |
unfolding starlike_def |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9163 |
apply (rule bexI [OF _ \<open>a \<in> T\<close>]) |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9164 |
apply (simp add: closed_segment_eq_open) |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9165 |
apply (intro conjI ballI a \<open>a \<in> T\<close> rel_interior_closure_convex_segment [OF \<open>convex S\<close> a]) |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9166 |
apply (simp add: order_trans [OF * ST]) |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9167 |
done |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9168 |
qed |
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9169 |
|
|
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
9170 |
subsection\<open>The relative frontier of a set\<close> |
|
62620
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9171 |
|
| 40377 | 9172 |
definition "rel_frontier S = closure S - rel_interior S" |
9173 |
||
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9174 |
lemma rel_frontier_empty [simp]: "rel_frontier {} = {}"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9175 |
by (simp add: rel_frontier_def) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9176 |
|
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
9177 |
lemma rel_frontier_eq_empty: |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
9178 |
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
|
9179 |
shows "rel_frontier S = {} \<longleftrightarrow> affine S"
|
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
9180 |
apply (simp add: rel_frontier_def) |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
9181 |
apply (simp add: rel_interior_eq_closure [symmetric]) |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
9182 |
using rel_interior_subset_closure by blast |
|
63593
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63566
diff
changeset
|
9183 |
|
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9184 |
lemma rel_frontier_sing [simp]: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9185 |
fixes a :: "'n::euclidean_space" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9186 |
shows "rel_frontier {a} = {}"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9187 |
by (simp add: rel_frontier_def) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9188 |
|
|
63492
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
9189 |
lemma rel_frontier_affine_hull: |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
9190 |
fixes S :: "'a::euclidean_space set" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
9191 |
shows "rel_frontier S \<subseteq> affine hull S" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
9192 |
using closure_affine_hull rel_frontier_def by fastforce |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
9193 |
|
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9194 |
lemma rel_frontier_cball [simp]: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9195 |
fixes a :: "'n::euclidean_space" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9196 |
shows "rel_frontier(cball a r) = (if r = 0 then {} else sphere a r)"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9197 |
proof (cases rule: linorder_cases [of r 0]) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9198 |
case less then show ?thesis |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9199 |
by (force simp: sphere_def) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9200 |
next |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9201 |
case equal then show ?thesis by simp |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9202 |
next |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9203 |
case greater then show ?thesis |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9204 |
apply simp |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9205 |
by (metis centre_in_ball empty_iff frontier_cball frontier_def interior_cball interior_rel_interior_gen rel_frontier_def) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9206 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9207 |
|
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
9208 |
lemma rel_frontier_translation: |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
9209 |
fixes a :: "'a::euclidean_space" |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
9210 |
shows "rel_frontier((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` (rel_frontier S)" |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
9211 |
by (simp add: rel_frontier_def translation_diff rel_interior_translation closure_translation) |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
9212 |
|
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
9213 |
lemma closed_affine_hull [iff]: |
| 53348 | 9214 |
fixes S :: "'n::euclidean_space set" |
9215 |
shows "closed (affine hull S)" |
|
9216 |
by (metis affine_affine_hull affine_closed) |
|
9217 |
||
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
9218 |
lemma rel_frontier_nonempty_interior: |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
9219 |
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
|
9220 |
shows "interior S \<noteq> {} \<Longrightarrow> rel_frontier S = frontier S"
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
9221 |
by (metis frontier_def interior_rel_interior_gen rel_frontier_def) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
9222 |
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
9223 |
lemma rel_frontier_frontier: |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
9224 |
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
|
9225 |
shows "affine hull S = UNIV \<Longrightarrow> rel_frontier S = frontier S" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
9226 |
by (simp add: frontier_def rel_frontier_def rel_interior_interior) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
9227 |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
9228 |
lemma closest_point_in_rel_frontier: |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
9229 |
"\<lbrakk>closed S; S \<noteq> {}; x \<in> affine hull S - rel_interior S\<rbrakk>
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
9230 |
\<Longrightarrow> closest_point S x \<in> rel_frontier S" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
9231 |
by (simp add: closest_point_in_rel_interior closest_point_in_set rel_frontier_def) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
9232 |
|
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
9233 |
lemma closed_rel_frontier [iff]: |
| 53348 | 9234 |
fixes S :: "'n::euclidean_space set" |
9235 |
shows "closed (rel_frontier S)" |
|
9236 |
proof - |
|
9237 |
have *: "closedin (subtopology euclidean (affine hull S)) (closure S - rel_interior S)" |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
9238 |
by (simp add: closed_subset closedin_diff closure_affine_hull openin_rel_interior) |
| 53348 | 9239 |
show ?thesis |
9240 |
apply (rule closedin_closed_trans[of "affine hull S" "rel_frontier S"]) |
|
9241 |
unfolding rel_frontier_def |
|
9242 |
using * closed_affine_hull |
|
9243 |
apply auto |
|
9244 |
done |
|
| 49531 | 9245 |
qed |
9246 |
||
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9247 |
lemma closed_rel_boundary: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9248 |
fixes S :: "'n::euclidean_space set" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9249 |
shows "closed S \<Longrightarrow> closed(S - rel_interior S)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9250 |
by (metis closed_rel_frontier closure_closed rel_frontier_def) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9251 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9252 |
lemma compact_rel_boundary: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9253 |
fixes S :: "'n::euclidean_space set" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9254 |
shows "compact S \<Longrightarrow> compact(S - rel_interior S)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9255 |
by (metis bounded_diff closed_rel_boundary closure_eq compact_closure compact_imp_closed) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9256 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9257 |
lemma bounded_rel_frontier: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9258 |
fixes S :: "'n::euclidean_space set" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9259 |
shows "bounded S \<Longrightarrow> bounded(rel_frontier S)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9260 |
by (simp add: bounded_closure bounded_diff rel_frontier_def) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9261 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9262 |
lemma compact_rel_frontier_bounded: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9263 |
fixes S :: "'n::euclidean_space set" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9264 |
shows "bounded S \<Longrightarrow> compact(rel_frontier S)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9265 |
using bounded_rel_frontier closed_rel_frontier compact_eq_bounded_closed by blast |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9266 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9267 |
lemma compact_rel_frontier: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9268 |
fixes S :: "'n::euclidean_space set" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9269 |
shows "compact S \<Longrightarrow> compact(rel_frontier S)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9270 |
by (meson compact_eq_bounded_closed compact_rel_frontier_bounded) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9271 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9272 |
lemma convex_same_rel_interior_closure: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9273 |
fixes S :: "'n::euclidean_space set" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9274 |
shows "\<lbrakk>convex S; convex T\<rbrakk> |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9275 |
\<Longrightarrow> rel_interior S = rel_interior T \<longleftrightarrow> closure S = closure T" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9276 |
by (simp add: closure_eq_rel_interior_eq) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9277 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9278 |
lemma convex_same_rel_interior_closure_straddle: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9279 |
fixes S :: "'n::euclidean_space set" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9280 |
shows "\<lbrakk>convex S; convex T\<rbrakk> |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9281 |
\<Longrightarrow> rel_interior S = rel_interior T \<longleftrightarrow> |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9282 |
rel_interior S \<subseteq> T \<and> T \<subseteq> closure S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
9283 |
by (simp add: closure_eq_between convex_same_rel_interior_closure) |
| 40377 | 9284 |
|
9285 |
lemma convex_rel_frontier_aff_dim: |
|
| 53348 | 9286 |
fixes S1 S2 :: "'n::euclidean_space set" |
9287 |
assumes "convex S1" |
|
9288 |
and "convex S2" |
|
9289 |
and "S2 \<noteq> {}"
|
|
9290 |
and "S1 \<le> rel_frontier S2" |
|
9291 |
shows "aff_dim S1 < aff_dim S2" |
|
9292 |
proof - |
|
9293 |
have "S1 \<subseteq> closure S2" |
|
9294 |
using assms unfolding rel_frontier_def by auto |
|
9295 |
then have *: "affine hull S1 \<subseteq> affine hull S2" |
|
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
9296 |
using hull_mono[of "S1" "closure S2"] closure_same_affine_hull[of S2] by blast |
| 53348 | 9297 |
then have "aff_dim S1 \<le> aff_dim S2" |
9298 |
using * aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] |
|
9299 |
aff_dim_subset[of "affine hull S1" "affine hull S2"] |
|
9300 |
by auto |
|
9301 |
moreover |
|
9302 |
{
|
|
9303 |
assume eq: "aff_dim S1 = aff_dim S2" |
|
9304 |
then have "S1 \<noteq> {}"
|
|
| 60420 | 9305 |
using aff_dim_empty[of S1] aff_dim_empty[of S2] \<open>S2 \<noteq> {}\<close> by auto
|
| 53348 | 9306 |
have **: "affine hull S1 = affine hull S2" |
9307 |
apply (rule affine_dim_equal) |
|
9308 |
using * affine_affine_hull |
|
9309 |
apply auto |
|
| 60420 | 9310 |
using \<open>S1 \<noteq> {}\<close> hull_subset[of S1]
|
| 53348 | 9311 |
apply auto |
9312 |
using eq aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] |
|
9313 |
apply auto |
|
9314 |
done |
|
9315 |
obtain a where a: "a \<in> rel_interior S1" |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
9316 |
using \<open>S1 \<noteq> {}\<close> rel_interior_eq_empty assms by auto
|
| 53348 | 9317 |
obtain T where T: "open T" "a \<in> T \<inter> S1" "T \<inter> affine hull S1 \<subseteq> S1" |
9318 |
using mem_rel_interior[of a S1] a by auto |
|
9319 |
then have "a \<in> T \<inter> closure S2" |
|
9320 |
using a assms unfolding rel_frontier_def by auto |
|
| 54465 | 9321 |
then obtain b where b: "b \<in> T \<inter> rel_interior S2" |
| 53348 | 9322 |
using open_inter_closure_rel_interior[of S2 T] assms T by auto |
9323 |
then have "b \<in> affine hull S1" |
|
9324 |
using rel_interior_subset hull_subset[of S2] ** by auto |
|
9325 |
then have "b \<in> S1" |
|
9326 |
using T b by auto |
|
9327 |
then have False |
|
9328 |
using b assms unfolding rel_frontier_def by auto |
|
9329 |
} |
|
9330 |
ultimately show ?thesis |
|
9331 |
using less_le by auto |
|
| 40377 | 9332 |
qed |
9333 |
||
9334 |
lemma convex_rel_interior_if: |
|
| 53348 | 9335 |
fixes S :: "'n::euclidean_space set" |
9336 |
assumes "convex S" |
|
9337 |
and "z \<in> rel_interior S" |
|
9338 |
shows "\<forall>x\<in>affine hull S. \<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)" |
|
9339 |
proof - |
|
| 54465 | 9340 |
obtain e1 where e1: "e1 > 0 \<and> cball z e1 \<inter> affine hull S \<subseteq> S" |
9341 |
using mem_rel_interior_cball[of z S] assms by auto |
|
| 53348 | 9342 |
{
|
9343 |
fix x |
|
9344 |
assume x: "x \<in> affine hull S" |
|
| 54465 | 9345 |
{
|
9346 |
assume "x \<noteq> z" |
|
| 63040 | 9347 |
define m where "m = 1 + e1/norm(x-z)" |
| 60420 | 9348 |
hence "m > 1" using e1 \<open>x \<noteq> z\<close> by auto |
| 53348 | 9349 |
{
|
9350 |
fix e |
|
9351 |
assume e: "e > 1 \<and> e \<le> m" |
|
9352 |
have "z \<in> affine hull S" |
|
9353 |
using assms rel_interior_subset hull_subset[of S] by auto |
|
9354 |
then have *: "(1 - e)*\<^sub>R x + e *\<^sub>R z \<in> affine hull S" |
|
9355 |
using mem_affine[of "affine hull S" x z "(1-e)" e] affine_affine_hull[of S] x |
|
9356 |
by auto |
|
9357 |
have "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) = norm ((e - 1) *\<^sub>R (x - z))" |
|
9358 |
by (simp add: algebra_simps) |
|
9359 |
also have "\<dots> = (e - 1) * norm (x-z)" |
|
9360 |
using norm_scaleR e by auto |
|
9361 |
also have "\<dots> \<le> (m - 1) * norm (x - z)" |
|
9362 |
using e mult_right_mono[of _ _ "norm(x-z)"] by auto |
|
9363 |
also have "\<dots> = (e1 / norm (x - z)) * norm (x - z)" |
|
9364 |
using m_def by auto |
|
9365 |
also have "\<dots> = e1" |
|
| 60420 | 9366 |
using \<open>x \<noteq> z\<close> e1 by simp |
| 53348 | 9367 |
finally have **: "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) \<le> e1" |
9368 |
by auto |
|
9369 |
have "(1 - e)*\<^sub>R x+ e *\<^sub>R z \<in> cball z e1" |
|
9370 |
using m_def ** |
|
9371 |
unfolding cball_def dist_norm |
|
9372 |
by (auto simp add: algebra_simps) |
|
9373 |
then have "(1 - e) *\<^sub>R x+ e *\<^sub>R z \<in> S" |
|
9374 |
using e * e1 by auto |
|
9375 |
} |
|
9376 |
then have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S )" |
|
| 60420 | 9377 |
using \<open>m> 1 \<close> by auto |
| 53348 | 9378 |
} |
9379 |
moreover |
|
9380 |
{
|
|
9381 |
assume "x = z" |
|
| 63040 | 9382 |
define m where "m = 1 + e1" |
| 53348 | 9383 |
then have "m > 1" |
9384 |
using e1 by auto |
|
9385 |
{
|
|
9386 |
fix e |
|
9387 |
assume e: "e > 1 \<and> e \<le> m" |
|
9388 |
then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" |
|
| 60420 | 9389 |
using e1 x \<open>x = z\<close> by (auto simp add: algebra_simps) |
| 53348 | 9390 |
then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" |
9391 |
using e by auto |
|
9392 |
} |
|
9393 |
then have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)" |
|
| 60420 | 9394 |
using \<open>m > 1\<close> by auto |
| 53348 | 9395 |
} |
9396 |
ultimately have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S )" |
|
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
9397 |
by blast |
| 40377 | 9398 |
} |
| 53348 | 9399 |
then show ?thesis by auto |
| 40377 | 9400 |
qed |
9401 |
||
9402 |
lemma convex_rel_interior_if2: |
|
| 53348 | 9403 |
fixes S :: "'n::euclidean_space set" |
9404 |
assumes "convex S" |
|
9405 |
assumes "z \<in> rel_interior S" |
|
9406 |
shows "\<forall>x\<in>affine hull S. \<exists>e. e > 1 \<and> (1 - e)*\<^sub>R x + e *\<^sub>R z \<in> S" |
|
9407 |
using convex_rel_interior_if[of S z] assms by auto |
|
| 40377 | 9408 |
|
9409 |
lemma convex_rel_interior_only_if: |
|
| 53348 | 9410 |
fixes S :: "'n::euclidean_space set" |
9411 |
assumes "convex S" |
|
9412 |
and "S \<noteq> {}"
|
|
9413 |
assumes "\<forall>x\<in>S. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" |
|
9414 |
shows "z \<in> rel_interior S" |
|
9415 |
proof - |
|
9416 |
obtain x where x: "x \<in> rel_interior S" |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
9417 |
using rel_interior_eq_empty assms by auto |
| 53348 | 9418 |
then have "x \<in> S" |
9419 |
using rel_interior_subset by auto |
|
9420 |
then obtain e where e: "e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" |
|
9421 |
using assms by auto |
|
| 63040 | 9422 |
define y where [abs_def]: "y = (1 - e) *\<^sub>R x + e *\<^sub>R z" |
| 53348 | 9423 |
then have "y \<in> S" using e by auto |
| 63040 | 9424 |
define e1 where "e1 = 1/e" |
| 53348 | 9425 |
then have "0 < e1 \<and> e1 < 1" using e by auto |
9426 |
then have "z =y - (1 - e1) *\<^sub>R (y - x)" |
|
9427 |
using e1_def y_def by (auto simp add: algebra_simps) |
|
9428 |
then show ?thesis |
|
| 60420 | 9429 |
using rel_interior_convex_shrink[of S x y "1-e1"] \<open>0 < e1 \<and> e1 < 1\<close> \<open>y \<in> S\<close> x assms |
| 53348 | 9430 |
by auto |
| 40377 | 9431 |
qed |
9432 |
||
9433 |
lemma convex_rel_interior_iff: |
|
| 53348 | 9434 |
fixes S :: "'n::euclidean_space set" |
9435 |
assumes "convex S" |
|
9436 |
and "S \<noteq> {}"
|
|
9437 |
shows "z \<in> rel_interior S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)" |
|
9438 |
using assms hull_subset[of S "affine"] |
|
9439 |
convex_rel_interior_if[of S z] convex_rel_interior_only_if[of S z] |
|
9440 |
by auto |
|
| 40377 | 9441 |
|
9442 |
lemma convex_rel_interior_iff2: |
|
| 53348 | 9443 |
fixes S :: "'n::euclidean_space set" |
9444 |
assumes "convex S" |
|
9445 |
and "S \<noteq> {}"
|
|
9446 |
shows "z \<in> rel_interior S \<longleftrightarrow> (\<forall>x\<in>affine hull S. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)" |
|
9447 |
using assms hull_subset[of S] convex_rel_interior_if2[of S z] convex_rel_interior_only_if[of S z] |
|
9448 |
by auto |
|
| 40377 | 9449 |
|
9450 |
lemma convex_interior_iff: |
|
| 53348 | 9451 |
fixes S :: "'n::euclidean_space set" |
9452 |
assumes "convex S" |
|
9453 |
shows "z \<in> interior S \<longleftrightarrow> (\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S)" |
|
9454 |
proof (cases "aff_dim S = int DIM('n)")
|
|
9455 |
case False |
|
9456 |
{
|
|
9457 |
assume "z \<in> interior S" |
|
9458 |
then have False |
|
9459 |
using False interior_rel_interior_gen[of S] by auto |
|
| 40377 | 9460 |
} |
9461 |
moreover |
|
| 53348 | 9462 |
{
|
9463 |
assume r: "\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S" |
|
9464 |
{
|
|
9465 |
fix x |
|
| 54465 | 9466 |
obtain e1 where e1: "e1 > 0 \<and> z + e1 *\<^sub>R (x - z) \<in> S" |
| 53348 | 9467 |
using r by auto |
| 54465 | 9468 |
obtain e2 where e2: "e2 > 0 \<and> z + e2 *\<^sub>R (z - x) \<in> S" |
| 53348 | 9469 |
using r by auto |
| 63040 | 9470 |
define x1 where [abs_def]: "x1 = z + e1 *\<^sub>R (x - z)" |
| 53348 | 9471 |
then have x1: "x1 \<in> affine hull S" |
| 54465 | 9472 |
using e1 hull_subset[of S] by auto |
| 63040 | 9473 |
define x2 where [abs_def]: "x2 = z + e2 *\<^sub>R (z - x)" |
| 53348 | 9474 |
then have x2: "x2 \<in> affine hull S" |
| 54465 | 9475 |
using e2 hull_subset[of S] by auto |
| 53348 | 9476 |
have *: "e1/(e1+e2) + e2/(e1+e2) = 1" |
| 54465 | 9477 |
using add_divide_distrib[of e1 e2 "e1+e2"] e1 e2 by simp |
| 53348 | 9478 |
then have "z = (e2/(e1+e2)) *\<^sub>R x1 + (e1/(e1+e2)) *\<^sub>R x2" |
9479 |
using x1_def x2_def |
|
9480 |
apply (auto simp add: algebra_simps) |
|
9481 |
using scaleR_left_distrib[of "e1/(e1+e2)" "e2/(e1+e2)" z] |
|
9482 |
apply auto |
|
9483 |
done |
|
9484 |
then have z: "z \<in> affine hull S" |
|
9485 |
using mem_affine[of "affine hull S" x1 x2 "e2/(e1+e2)" "e1/(e1+e2)"] |
|
9486 |
x1 x2 affine_affine_hull[of S] * |
|
9487 |
by auto |
|
9488 |
have "x1 - x2 = (e1 + e2) *\<^sub>R (x - z)" |
|
9489 |
using x1_def x2_def by (auto simp add: algebra_simps) |
|
9490 |
then have "x = z+(1/(e1+e2)) *\<^sub>R (x1-x2)" |
|
| 54465 | 9491 |
using e1 e2 by simp |
| 53348 | 9492 |
then have "x \<in> affine hull S" |
9493 |
using mem_affine_3_minus[of "affine hull S" z x1 x2 "1/(e1+e2)"] |
|
9494 |
x1 x2 z affine_affine_hull[of S] |
|
9495 |
by auto |
|
9496 |
} |
|
9497 |
then have "affine hull S = UNIV" |
|
9498 |
by auto |
|
9499 |
then have "aff_dim S = int DIM('n)"
|
|
| 63072 | 9500 |
using aff_dim_affine_hull[of S] by (simp add: aff_dim_UNIV) |
| 53348 | 9501 |
then have False |
9502 |
using False by auto |
|
9503 |
} |
|
9504 |
ultimately show ?thesis by auto |
|
9505 |
next |
|
9506 |
case True |
|
9507 |
then have "S \<noteq> {}"
|
|
9508 |
using aff_dim_empty[of S] by auto |
|
9509 |
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
|
9510 |
using True affine_hull_UNIV by auto |
| 53348 | 9511 |
{
|
9512 |
assume "z \<in> interior S" |
|
9513 |
then have "z \<in> rel_interior S" |
|
9514 |
using True interior_rel_interior_gen[of S] by auto |
|
9515 |
then have **: "\<forall>x. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" |
|
| 60420 | 9516 |
using convex_rel_interior_iff2[of S z] assms \<open>S \<noteq> {}\<close> * by auto
|
| 53348 | 9517 |
fix x |
9518 |
obtain e1 where e1: "e1 > 1" "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z \<in> S" |
|
| 40377 | 9519 |
using **[rule_format, of "z-x"] by auto |
| 63040 | 9520 |
define e where [abs_def]: "e = e1 - 1" |
| 53348 | 9521 |
then have "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z = z + e *\<^sub>R x" |
9522 |
by (simp add: algebra_simps) |
|
9523 |
then have "e > 0" "z + e *\<^sub>R x \<in> S" |
|
9524 |
using e1 e_def by auto |
|
9525 |
then have "\<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S" |
|
9526 |
by auto |
|
| 40377 | 9527 |
} |
9528 |
moreover |
|
| 53348 | 9529 |
{
|
9530 |
assume r: "\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S" |
|
9531 |
{
|
|
9532 |
fix x |
|
9533 |
obtain e1 where e1: "e1 > 0" "z + e1 *\<^sub>R (z - x) \<in> S" |
|
9534 |
using r[rule_format, of "z-x"] by auto |
|
| 63040 | 9535 |
define e where "e = e1 + 1" |
| 53348 | 9536 |
then have "z + e1 *\<^sub>R (z - x) = (1 - e) *\<^sub>R x + e *\<^sub>R z" |
9537 |
by (simp add: algebra_simps) |
|
9538 |
then have "e > 1" "(1 - e)*\<^sub>R x + e *\<^sub>R z \<in> S" |
|
9539 |
using e1 e_def by auto |
|
9540 |
then have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" by auto |
|
| 40377 | 9541 |
} |
| 53348 | 9542 |
then have "z \<in> rel_interior S" |
| 60420 | 9543 |
using convex_rel_interior_iff2[of S z] assms \<open>S \<noteq> {}\<close> by auto
|
| 53348 | 9544 |
then have "z \<in> interior S" |
9545 |
using True interior_rel_interior_gen[of S] by auto |
|
9546 |
} |
|
9547 |
ultimately show ?thesis by auto |
|
9548 |
qed |
|
9549 |
||
| 40377 | 9550 |
|
| 60420 | 9551 |
subsubsection \<open>Relative interior and closure under common operations\<close> |
| 40377 | 9552 |
|
| 53348 | 9553 |
lemma rel_interior_inter_aux: "\<Inter>{rel_interior S |S. S : I} \<subseteq> \<Inter>I"
|
9554 |
proof - |
|
9555 |
{
|
|
9556 |
fix y |
|
9557 |
assume "y \<in> \<Inter>{rel_interior S |S. S : I}"
|
|
9558 |
then have y: "\<forall>S \<in> I. y \<in> rel_interior S" |
|
9559 |
by auto |
|
9560 |
{
|
|
9561 |
fix S |
|
9562 |
assume "S \<in> I" |
|
9563 |
then have "y \<in> S" |
|
9564 |
using rel_interior_subset y by auto |
|
9565 |
} |
|
9566 |
then have "y \<in> \<Inter>I" by auto |
|
9567 |
} |
|
9568 |
then show ?thesis by auto |
|
9569 |
qed |
|
9570 |
||
| 63128 | 9571 |
lemma closure_Int: "closure (\<Inter>I) \<le> \<Inter>{closure S |S. S \<in> I}"
|
| 53348 | 9572 |
proof - |
9573 |
{
|
|
9574 |
fix y |
|
9575 |
assume "y \<in> \<Inter>I" |
|
9576 |
then have y: "\<forall>S \<in> I. y \<in> S" by auto |
|
9577 |
{
|
|
9578 |
fix S |
|
9579 |
assume "S \<in> I" |
|
9580 |
then have "y \<in> closure S" |
|
9581 |
using closure_subset y by auto |
|
9582 |
} |
|
9583 |
then have "y \<in> \<Inter>{closure S |S. S \<in> I}"
|
|
9584 |
by auto |
|
9585 |
} |
|
9586 |
then have "\<Inter>I \<subseteq> \<Inter>{closure S |S. S \<in> I}"
|
|
9587 |
by auto |
|
| 54465 | 9588 |
moreover have "closed (\<Inter>{closure S |S. S \<in> I})"
|
| 53348 | 9589 |
unfolding closed_Inter closed_closure by auto |
| 54465 | 9590 |
ultimately show ?thesis using closure_hull[of "\<Inter>I"] |
| 53348 | 9591 |
hull_minimal[of "\<Inter>I" "\<Inter>{closure S |S. S \<in> I}" "closed"] by auto
|
| 40377 | 9592 |
qed |
9593 |
||
| 49531 | 9594 |
lemma convex_closure_rel_interior_inter: |
| 53348 | 9595 |
assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)" |
9596 |
and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
|
|
9597 |
shows "\<Inter>{closure S |S. S \<in> I} \<le> closure (\<Inter>{rel_interior S |S. S \<in> I})"
|
|
9598 |
proof - |
|
9599 |
obtain x where x: "\<forall>S\<in>I. x \<in> rel_interior S" |
|
9600 |
using assms by auto |
|
9601 |
{
|
|
9602 |
fix y |
|
9603 |
assume "y \<in> \<Inter>{closure S |S. S \<in> I}"
|
|
9604 |
then have y: "\<forall>S \<in> I. y \<in> closure S" |
|
9605 |
by auto |
|
9606 |
{
|
|
9607 |
assume "y = x" |
|
9608 |
then have "y \<in> closure (\<Inter>{rel_interior S |S. S \<in> I})"
|
|
| 54465 | 9609 |
using x closure_subset[of "\<Inter>{rel_interior S |S. S \<in> I}"] by auto
|
| 53348 | 9610 |
} |
9611 |
moreover |
|
9612 |
{
|
|
9613 |
assume "y \<noteq> x" |
|
9614 |
{ fix e :: real
|
|
9615 |
assume e: "e > 0" |
|
| 63040 | 9616 |
define e1 where "e1 = min 1 (e/norm (y - x))" |
| 53348 | 9617 |
then have e1: "e1 > 0" "e1 \<le> 1" "e1 * norm (y - x) \<le> e" |
| 60420 | 9618 |
using \<open>y \<noteq> x\<close> \<open>e > 0\<close> le_divide_eq[of e1 e "norm (y - x)"] |
| 53348 | 9619 |
by simp_all |
| 63040 | 9620 |
define z where "z = y - e1 *\<^sub>R (y - x)" |
| 53348 | 9621 |
{
|
9622 |
fix S |
|
9623 |
assume "S \<in> I" |
|
9624 |
then have "z \<in> rel_interior S" |
|
9625 |
using rel_interior_closure_convex_shrink[of S x y e1] assms x y e1 z_def |
|
9626 |
by auto |
|
9627 |
} |
|
9628 |
then have *: "z \<in> \<Inter>{rel_interior S |S. S \<in> I}"
|
|
9629 |
by auto |
|
9630 |
have "\<exists>z. z \<in> \<Inter>{rel_interior S |S. S \<in> I} \<and> z \<noteq> y \<and> dist z y \<le> e"
|
|
9631 |
apply (rule_tac x="z" in exI) |
|
| 60420 | 9632 |
using \<open>y \<noteq> x\<close> z_def * e1 e dist_norm[of z y] |
| 53348 | 9633 |
apply simp |
9634 |
done |
|
9635 |
} |
|
9636 |
then have "y islimpt \<Inter>{rel_interior S |S. S \<in> I}"
|
|
9637 |
unfolding islimpt_approachable_le by blast |
|
9638 |
then have "y \<in> closure (\<Inter>{rel_interior S |S. S \<in> I})"
|
|
9639 |
unfolding closure_def by auto |
|
9640 |
} |
|
9641 |
ultimately have "y \<in> closure (\<Inter>{rel_interior S |S. S \<in> I})"
|
|
9642 |
by auto |
|
| 40377 | 9643 |
} |
| 53348 | 9644 |
then show ?thesis by auto |
| 40377 | 9645 |
qed |
9646 |
||
| 49531 | 9647 |
lemma convex_closure_inter: |
| 53348 | 9648 |
assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)" |
9649 |
and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
|
|
| 54465 | 9650 |
shows "closure (\<Inter>I) = \<Inter>{closure S |S. S \<in> I}"
|
| 53348 | 9651 |
proof - |
9652 |
have "\<Inter>{closure S |S. S \<in> I} \<le> closure (\<Inter>{rel_interior S |S. S \<in> I})"
|
|
9653 |
using convex_closure_rel_interior_inter assms by auto |
|
9654 |
moreover |
|
| 60585 | 9655 |
have "closure (\<Inter>{rel_interior S |S. S \<in> I}) \<le> closure (\<Inter>I)"
|
| 54465 | 9656 |
using rel_interior_inter_aux closure_mono[of "\<Inter>{rel_interior S |S. S \<in> I}" "\<Inter>I"]
|
| 53348 | 9657 |
by auto |
9658 |
ultimately show ?thesis |
|
| 63128 | 9659 |
using closure_Int[of I] by auto |
| 40377 | 9660 |
qed |
9661 |
||
| 49531 | 9662 |
lemma convex_inter_rel_interior_same_closure: |
| 53348 | 9663 |
assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)" |
9664 |
and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
|
|
| 54465 | 9665 |
shows "closure (\<Inter>{rel_interior S |S. S \<in> I}) = closure (\<Inter>I)"
|
| 53348 | 9666 |
proof - |
9667 |
have "\<Inter>{closure S |S. S \<in> I} \<le> closure (\<Inter>{rel_interior S |S. S \<in> I})"
|
|
9668 |
using convex_closure_rel_interior_inter assms by auto |
|
9669 |
moreover |
|
9670 |
have "closure (\<Inter>{rel_interior S |S. S \<in> I}) \<le> closure (\<Inter>I)"
|
|
| 54465 | 9671 |
using rel_interior_inter_aux closure_mono[of "\<Inter>{rel_interior S |S. S \<in> I}" "\<Inter>I"]
|
| 53348 | 9672 |
by auto |
9673 |
ultimately show ?thesis |
|
| 63128 | 9674 |
using closure_Int[of I] by auto |
| 40377 | 9675 |
qed |
9676 |
||
| 49531 | 9677 |
lemma convex_rel_interior_inter: |
| 53348 | 9678 |
assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)" |
9679 |
and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
|
|
| 54465 | 9680 |
shows "rel_interior (\<Inter>I) \<subseteq> \<Inter>{rel_interior S |S. S \<in> I}"
|
| 53348 | 9681 |
proof - |
9682 |
have "convex (\<Inter>I)" |
|
9683 |
using assms convex_Inter by auto |
|
9684 |
moreover |
|
| 54465 | 9685 |
have "convex (\<Inter>{rel_interior S |S. S \<in> I})"
|
| 53348 | 9686 |
apply (rule convex_Inter) |
9687 |
using assms convex_rel_interior |
|
9688 |
apply auto |
|
9689 |
done |
|
9690 |
ultimately |
|
9691 |
have "rel_interior (\<Inter>{rel_interior S |S. S \<in> I}) = rel_interior (\<Inter>I)"
|
|
9692 |
using convex_inter_rel_interior_same_closure assms |
|
9693 |
closure_eq_rel_interior_eq[of "\<Inter>{rel_interior S |S. S \<in> I}" "\<Inter>I"]
|
|
9694 |
by blast |
|
9695 |
then show ?thesis |
|
9696 |
using rel_interior_subset[of "\<Inter>{rel_interior S |S. S \<in> I}"] by auto
|
|
| 40377 | 9697 |
qed |
9698 |
||
| 49531 | 9699 |
lemma convex_rel_interior_finite_inter: |
| 53348 | 9700 |
assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)" |
9701 |
and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
|
|
9702 |
and "finite I" |
|
9703 |
shows "rel_interior (\<Inter>I) = \<Inter>{rel_interior S |S. S \<in> I}"
|
|
9704 |
proof - |
|
9705 |
have "\<Inter>I \<noteq> {}"
|
|
9706 |
using assms rel_interior_inter_aux[of I] by auto |
|
9707 |
have "convex (\<Inter>I)" |
|
9708 |
using convex_Inter assms by auto |
|
9709 |
show ?thesis |
|
9710 |
proof (cases "I = {}")
|
|
9711 |
case True |
|
9712 |
then show ?thesis |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
9713 |
using Inter_empty rel_interior_UNIV by auto |
| 53348 | 9714 |
next |
9715 |
case False |
|
9716 |
{
|
|
9717 |
fix z |
|
9718 |
assume z: "z \<in> \<Inter>{rel_interior S |S. S \<in> I}"
|
|
9719 |
{
|
|
9720 |
fix x |
|
| 61952 | 9721 |
assume x: "x \<in> \<Inter>I" |
| 53348 | 9722 |
{
|
9723 |
fix S |
|
9724 |
assume S: "S \<in> I" |
|
9725 |
then have "z \<in> rel_interior S" "x \<in> S" |
|
9726 |
using z x by auto |
|
9727 |
then have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e)*\<^sub>R x + e *\<^sub>R z \<in> S)" |
|
9728 |
using convex_rel_interior_if[of S z] S assms hull_subset[of S] by auto |
|
9729 |
} |
|
9730 |
then obtain mS where |
|
9731 |
mS: "\<forall>S\<in>I. mS S > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> mS S \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)" by metis |
|
| 63040 | 9732 |
define e where "e = Min (mS ` I)" |
| 60420 | 9733 |
then have "e \<in> mS ` I" using assms \<open>I \<noteq> {}\<close> by simp
|
| 53348 | 9734 |
then have "e > 1" using mS by auto |
9735 |
moreover have "\<forall>S\<in>I. e \<le> mS S" |
|
9736 |
using e_def assms by auto |
|
9737 |
ultimately have "\<exists>e > 1. (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> \<Inter>I" |
|
9738 |
using mS by auto |
|
9739 |
} |
|
9740 |
then have "z \<in> rel_interior (\<Inter>I)" |
|
| 60420 | 9741 |
using convex_rel_interior_iff[of "\<Inter>I" z] \<open>\<Inter>I \<noteq> {}\<close> \<open>convex (\<Inter>I)\<close> by auto
|
| 53348 | 9742 |
} |
9743 |
then show ?thesis |
|
9744 |
using convex_rel_interior_inter[of I] assms by auto |
|
9745 |
qed |
|
| 40377 | 9746 |
qed |
9747 |
||
| 49531 | 9748 |
lemma convex_closure_inter_two: |
| 53348 | 9749 |
fixes S T :: "'n::euclidean_space set" |
9750 |
assumes "convex S" |
|
9751 |
and "convex T" |
|
9752 |
assumes "rel_interior S \<inter> rel_interior T \<noteq> {}"
|
|
9753 |
shows "closure (S \<inter> T) = closure S \<inter> closure T" |
|
9754 |
using convex_closure_inter[of "{S,T}"] assms by auto
|
|
| 40377 | 9755 |
|
| 49531 | 9756 |
lemma convex_rel_interior_inter_two: |
| 53348 | 9757 |
fixes S T :: "'n::euclidean_space set" |
9758 |
assumes "convex S" |
|
9759 |
and "convex T" |
|
9760 |
and "rel_interior S \<inter> rel_interior T \<noteq> {}"
|
|
9761 |
shows "rel_interior (S \<inter> T) = rel_interior S \<inter> rel_interior T" |
|
9762 |
using convex_rel_interior_finite_inter[of "{S,T}"] assms by auto
|
|
| 40377 | 9763 |
|
| 63128 | 9764 |
lemma convex_affine_closure_Int: |
| 53348 | 9765 |
fixes S T :: "'n::euclidean_space set" |
9766 |
assumes "convex S" |
|
9767 |
and "affine T" |
|
9768 |
and "rel_interior S \<inter> T \<noteq> {}"
|
|
9769 |
shows "closure (S \<inter> T) = closure S \<inter> T" |
|
9770 |
proof - |
|
9771 |
have "affine hull T = T" |
|
9772 |
using assms by auto |
|
9773 |
then have "rel_interior T = T" |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
9774 |
using rel_interior_affine_hull[of T] by metis |
| 53348 | 9775 |
moreover have "closure T = T" |
9776 |
using assms affine_closed[of T] by auto |
|
9777 |
ultimately show ?thesis |
|
9778 |
using convex_closure_inter_two[of S T] assms affine_imp_convex by auto |
|
| 49531 | 9779 |
qed |
9780 |
||
|
62620
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9781 |
lemma connected_component_1_gen: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9782 |
fixes S :: "'a :: euclidean_space set" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9783 |
assumes "DIM('a) = 1"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9784 |
shows "connected_component S a b \<longleftrightarrow> closed_segment a b \<subseteq> S" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9785 |
unfolding connected_component_def |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9786 |
by (metis (no_types, lifting) assms subsetD subsetI convex_contains_segment convex_segment(1) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9787 |
ends_in_segment connected_convex_1_gen) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9788 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9789 |
lemma connected_component_1: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9790 |
fixes S :: "real set" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9791 |
shows "connected_component S a b \<longleftrightarrow> closed_segment a b \<subseteq> S" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9792 |
by (simp add: connected_component_1_gen) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
9793 |
|
| 63128 | 9794 |
lemma convex_affine_rel_interior_Int: |
| 53348 | 9795 |
fixes S T :: "'n::euclidean_space set" |
9796 |
assumes "convex S" |
|
9797 |
and "affine T" |
|
9798 |
and "rel_interior S \<inter> T \<noteq> {}"
|
|
9799 |
shows "rel_interior (S \<inter> T) = rel_interior S \<inter> T" |
|
9800 |
proof - |
|
9801 |
have "affine hull T = T" |
|
9802 |
using assms by auto |
|
9803 |
then have "rel_interior T = T" |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
9804 |
using rel_interior_affine_hull[of T] by metis |
| 53348 | 9805 |
moreover have "closure T = T" |
9806 |
using assms affine_closed[of T] by auto |
|
9807 |
ultimately show ?thesis |
|
9808 |
using convex_rel_interior_inter_two[of S T] assms affine_imp_convex by auto |
|
| 40377 | 9809 |
qed |
9810 |
||
| 63128 | 9811 |
lemma convex_affine_rel_frontier_Int: |
9812 |
fixes S T :: "'n::euclidean_space set" |
|
9813 |
assumes "convex S" |
|
9814 |
and "affine T" |
|
9815 |
and "interior S \<inter> T \<noteq> {}"
|
|
9816 |
shows "rel_frontier(S \<inter> T) = frontier S \<inter> T" |
|
9817 |
using assms |
|
9818 |
apply (simp add: rel_frontier_def convex_affine_closure_Int frontier_def) |
|
9819 |
by (metis Diff_Int_distrib2 Int_emptyI convex_affine_closure_Int convex_affine_rel_interior_Int empty_iff interior_rel_interior_gen) |
|
9820 |
||
| 63955 | 9821 |
lemma rel_interior_convex_Int_affine: |
9822 |
fixes S :: "'a::euclidean_space set" |
|
9823 |
assumes "convex S" "affine T" "interior S \<inter> T \<noteq> {}"
|
|
9824 |
shows "rel_interior(S \<inter> T) = interior S \<inter> T" |
|
9825 |
proof - |
|
9826 |
obtain a where aS: "a \<in> interior S" and aT:"a \<in> T" |
|
9827 |
using assms by force |
|
9828 |
have "rel_interior S = interior S" |
|
9829 |
by (metis (no_types) aS affine_hull_nonempty_interior equals0D rel_interior_interior) |
|
9830 |
then show ?thesis |
|
9831 |
by (metis (no_types) affine_imp_convex assms convex_rel_interior_inter_two hull_same rel_interior_affine_hull) |
|
9832 |
qed |
|
9833 |
||
9834 |
lemma closure_convex_Int_affine: |
|
9835 |
fixes S :: "'a::euclidean_space set" |
|
9836 |
assumes "convex S" "affine T" "rel_interior S \<inter> T \<noteq> {}"
|
|
9837 |
shows "closure(S \<inter> T) = closure S \<inter> T" |
|
9838 |
proof |
|
9839 |
have "closure (S \<inter> T) \<subseteq> closure T" |
|
9840 |
by (simp add: closure_mono) |
|
9841 |
also have "... \<subseteq> T" |
|
9842 |
by (simp add: affine_closed assms) |
|
9843 |
finally show "closure(S \<inter> T) \<subseteq> closure S \<inter> T" |
|
9844 |
by (simp add: closure_mono) |
|
9845 |
next |
|
9846 |
obtain a where "a \<in> rel_interior S" "a \<in> T" |
|
9847 |
using assms by auto |
|
9848 |
then have ssT: "subspace ((\<lambda>x. (-a)+x) ` T)" and "a \<in> S" |
|
9849 |
using affine_diffs_subspace rel_interior_subset assms by blast+ |
|
9850 |
show "closure S \<inter> T \<subseteq> closure (S \<inter> T)" |
|
9851 |
proof |
|
9852 |
fix x assume "x \<in> closure S \<inter> T" |
|
9853 |
show "x \<in> closure (S \<inter> T)" |
|
9854 |
proof (cases "x = a") |
|
9855 |
case True |
|
9856 |
then show ?thesis |
|
9857 |
using \<open>a \<in> S\<close> \<open>a \<in> T\<close> closure_subset by fastforce |
|
9858 |
next |
|
9859 |
case False |
|
9860 |
then have "x \<in> closure(open_segment a x)" |
|
9861 |
by auto |
|
9862 |
then show ?thesis |
|
9863 |
using \<open>x \<in> closure S \<inter> T\<close> assms convex_affine_closure_Int by blast |
|
9864 |
qed |
|
9865 |
qed |
|
9866 |
qed |
|
9867 |
||
| 40377 | 9868 |
lemma subset_rel_interior_convex: |
| 53348 | 9869 |
fixes S T :: "'n::euclidean_space set" |
9870 |
assumes "convex S" |
|
9871 |
and "convex T" |
|
9872 |
and "S \<le> closure T" |
|
9873 |
and "\<not> S \<subseteq> rel_frontier T" |
|
9874 |
shows "rel_interior S \<subseteq> rel_interior T" |
|
9875 |
proof - |
|
9876 |
have *: "S \<inter> closure T = S" |
|
9877 |
using assms by auto |
|
9878 |
have "\<not> rel_interior S \<subseteq> rel_frontier T" |
|
9879 |
using closure_mono[of "rel_interior S" "rel_frontier T"] closed_rel_frontier[of T] |
|
9880 |
closure_closed[of S] convex_closure_rel_interior[of S] closure_subset[of S] assms |
|
9881 |
by auto |
|
9882 |
then have "rel_interior S \<inter> rel_interior (closure T) \<noteq> {}"
|
|
9883 |
using assms rel_frontier_def[of T] rel_interior_subset convex_rel_interior_closure[of T] |
|
9884 |
by auto |
|
9885 |
then have "rel_interior S \<inter> rel_interior T = rel_interior (S \<inter> closure T)" |
|
9886 |
using assms convex_closure convex_rel_interior_inter_two[of S "closure T"] |
|
9887 |
convex_rel_interior_closure[of T] |
|
9888 |
by auto |
|
9889 |
also have "\<dots> = rel_interior S" |
|
9890 |
using * by auto |
|
9891 |
finally show ?thesis |
|
9892 |
by auto |
|
9893 |
qed |
|
| 40377 | 9894 |
|
9895 |
lemma rel_interior_convex_linear_image: |
|
| 53348 | 9896 |
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" |
9897 |
assumes "linear f" |
|
9898 |
and "convex S" |
|
9899 |
shows "f ` (rel_interior S) = rel_interior (f ` S)" |
|
9900 |
proof (cases "S = {}")
|
|
9901 |
case True |
|
9902 |
then show ?thesis |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
9903 |
using assms rel_interior_empty rel_interior_eq_empty by auto |
| 53348 | 9904 |
next |
9905 |
case False |
|
9906 |
have *: "f ` (rel_interior S) \<subseteq> f ` S" |
|
9907 |
unfolding image_mono using rel_interior_subset by auto |
|
9908 |
have "f ` S \<subseteq> f ` (closure S)" |
|
9909 |
unfolding image_mono using closure_subset by auto |
|
9910 |
also have "\<dots> = f ` (closure (rel_interior S))" |
|
9911 |
using convex_closure_rel_interior assms by auto |
|
9912 |
also have "\<dots> \<subseteq> closure (f ` (rel_interior S))" |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
9913 |
using closure_linear_image_subset assms by auto |
| 53348 | 9914 |
finally have "closure (f ` S) = closure (f ` rel_interior S)" |
9915 |
using closure_mono[of "f ` S" "closure (f ` rel_interior S)"] closure_closure |
|
9916 |
closure_mono[of "f ` rel_interior S" "f ` S"] * |
|
9917 |
by auto |
|
9918 |
then have "rel_interior (f ` S) = rel_interior (f ` rel_interior S)" |
|
9919 |
using assms convex_rel_interior |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
9920 |
linear_conv_bounded_linear[of f] convex_linear_image[of _ S] |
|
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
9921 |
convex_linear_image[of _ "rel_interior S"] |
| 53348 | 9922 |
closure_eq_rel_interior_eq[of "f ` S" "f ` rel_interior S"] |
9923 |
by auto |
|
9924 |
then have "rel_interior (f ` S) \<subseteq> f ` rel_interior S" |
|
9925 |
using rel_interior_subset by auto |
|
9926 |
moreover |
|
9927 |
{
|
|
9928 |
fix z |
|
9929 |
assume "z \<in> f ` rel_interior S" |
|
9930 |
then obtain z1 where z1: "z1 \<in> rel_interior S" "f z1 = z" by auto |
|
9931 |
{
|
|
9932 |
fix x |
|
9933 |
assume "x \<in> f ` S" |
|
9934 |
then obtain x1 where x1: "x1 \<in> S" "f x1 = x" by auto |
|
| 54465 | 9935 |
then obtain e where e: "e > 1" "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1 : S" |
| 60420 | 9936 |
using convex_rel_interior_iff[of S z1] \<open>convex S\<close> x1 z1 by auto |
| 53348 | 9937 |
moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z" |
| 60420 | 9938 |
using x1 z1 \<open>linear f\<close> by (simp add: linear_add_cmul) |
| 53348 | 9939 |
ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S" |
| 40377 | 9940 |
using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto |
| 53348 | 9941 |
then have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S" |
| 54465 | 9942 |
using e by auto |
| 53348 | 9943 |
} |
9944 |
then have "z \<in> rel_interior (f ` S)" |
|
| 60420 | 9945 |
using convex_rel_interior_iff[of "f ` S" z] \<open>convex S\<close> |
9946 |
\<open>linear f\<close> \<open>S \<noteq> {}\<close> convex_linear_image[of f S] linear_conv_bounded_linear[of f]
|
|
| 53348 | 9947 |
by auto |
9948 |
} |
|
9949 |
ultimately show ?thesis by auto |
|
| 40377 | 9950 |
qed |
9951 |
||
9952 |
lemma rel_interior_convex_linear_preimage: |
|
| 53348 | 9953 |
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" |
9954 |
assumes "linear f" |
|
9955 |
and "convex S" |
|
9956 |
and "f -` (rel_interior S) \<noteq> {}"
|
|
9957 |
shows "rel_interior (f -` S) = f -` (rel_interior S)" |
|
9958 |
proof - |
|
9959 |
have "S \<noteq> {}"
|
|
9960 |
using assms rel_interior_empty by auto |
|
9961 |
have nonemp: "f -` S \<noteq> {}"
|
|
9962 |
by (metis assms(3) rel_interior_subset subset_empty vimage_mono) |
|
9963 |
then have "S \<inter> (range f) \<noteq> {}"
|
|
9964 |
by auto |
|
9965 |
have conv: "convex (f -` S)" |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
9966 |
using convex_linear_vimage assms by auto |
| 53348 | 9967 |
then have "convex (S \<inter> range f)" |
9968 |
by (metis assms(1) assms(2) convex_Int subspace_UNIV subspace_imp_convex subspace_linear_image) |
|
9969 |
{
|
|
9970 |
fix z |
|
9971 |
assume "z \<in> f -` (rel_interior S)" |
|
9972 |
then have z: "f z : rel_interior S" |
|
9973 |
by auto |
|
9974 |
{
|
|
9975 |
fix x |
|
9976 |
assume "x \<in> f -` S" |
|
9977 |
then have "f x \<in> S" by auto |
|
9978 |
then obtain e where e: "e > 1" "(1 - e) *\<^sub>R f x + e *\<^sub>R f z \<in> S" |
|
| 60420 | 9979 |
using convex_rel_interior_iff[of S "f z"] z assms \<open>S \<noteq> {}\<close> by auto
|
| 53348 | 9980 |
moreover have "(1 - e) *\<^sub>R f x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R x + e *\<^sub>R z)" |
| 60420 | 9981 |
using \<open>linear f\<close> by (simp add: linear_iff) |
| 53348 | 9982 |
ultimately have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> f -` S" |
9983 |
using e by auto |
|
9984 |
} |
|
9985 |
then have "z \<in> rel_interior (f -` S)" |
|
9986 |
using convex_rel_interior_iff[of "f -` S" z] conv nonemp by auto |
|
9987 |
} |
|
9988 |
moreover |
|
| 54465 | 9989 |
{
|
| 53348 | 9990 |
fix z |
9991 |
assume z: "z \<in> rel_interior (f -` S)" |
|
9992 |
{
|
|
9993 |
fix x |
|
9994 |
assume "x \<in> S \<inter> range f" |
|
9995 |
then obtain y where y: "f y = x" "y \<in> f -` S" by auto |
|
9996 |
then obtain e where e: "e > 1" "(1 - e) *\<^sub>R y + e *\<^sub>R z \<in> f -` S" |
|
9997 |
using convex_rel_interior_iff[of "f -` S" z] z conv by auto |
|
9998 |
moreover have "(1 - e) *\<^sub>R x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R y + e *\<^sub>R z)" |
|
| 60420 | 9999 |
using \<open>linear f\<close> y by (simp add: linear_iff) |
| 53348 | 10000 |
ultimately have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R f z \<in> S \<inter> range f" |
10001 |
using e by auto |
|
10002 |
} |
|
10003 |
then have "f z \<in> rel_interior (S \<inter> range f)" |
|
| 60420 | 10004 |
using \<open>convex (S \<inter> (range f))\<close> \<open>S \<inter> range f \<noteq> {}\<close>
|
| 53348 | 10005 |
convex_rel_interior_iff[of "S \<inter> (range f)" "f z"] |
10006 |
by auto |
|
10007 |
moreover have "affine (range f)" |
|
10008 |
by (metis assms(1) subspace_UNIV subspace_imp_affine subspace_linear_image) |
|
10009 |
ultimately have "f z \<in> rel_interior S" |
|
| 63128 | 10010 |
using convex_affine_rel_interior_Int[of S "range f"] assms by auto |
| 53348 | 10011 |
then have "z \<in> f -` (rel_interior S)" |
10012 |
by auto |
|
10013 |
} |
|
10014 |
ultimately show ?thesis by auto |
|
| 40377 | 10015 |
qed |
| 49531 | 10016 |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
10017 |
lemma rel_interior_Times: |
| 53348 | 10018 |
fixes S :: "'n::euclidean_space set" |
10019 |
and T :: "'m::euclidean_space set" |
|
10020 |
assumes "convex S" |
|
10021 |
and "convex T" |
|
10022 |
shows "rel_interior (S \<times> T) = rel_interior S \<times> rel_interior T" |
|
10023 |
proof - |
|
| 60303 | 10024 |
{ assume "S = {}"
|
| 53348 | 10025 |
then have ?thesis |
| 60303 | 10026 |
by auto |
| 53348 | 10027 |
} |
10028 |
moreover |
|
| 60303 | 10029 |
{ assume "T = {}"
|
| 53348 | 10030 |
then have ?thesis |
| 60303 | 10031 |
by auto |
| 53348 | 10032 |
} |
10033 |
moreover |
|
10034 |
{
|
|
10035 |
assume "S \<noteq> {}" "T \<noteq> {}"
|
|
10036 |
then have ri: "rel_interior S \<noteq> {}" "rel_interior T \<noteq> {}"
|
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
10037 |
using rel_interior_eq_empty assms by auto |
| 53348 | 10038 |
then have "fst -` rel_interior S \<noteq> {}"
|
10039 |
using fst_vimage_eq_Times[of "rel_interior S"] by auto |
|
10040 |
then have "rel_interior ((fst :: 'n * 'm \<Rightarrow> 'n) -` S) = fst -` rel_interior S" |
|
| 60420 | 10041 |
using fst_linear \<open>convex S\<close> rel_interior_convex_linear_preimage[of fst S] by auto |
| 53348 | 10042 |
then have s: "rel_interior (S \<times> (UNIV :: 'm set)) = rel_interior S \<times> UNIV" |
10043 |
by (simp add: fst_vimage_eq_Times) |
|
10044 |
from ri have "snd -` rel_interior T \<noteq> {}"
|
|
10045 |
using snd_vimage_eq_Times[of "rel_interior T"] by auto |
|
10046 |
then have "rel_interior ((snd :: 'n * 'm \<Rightarrow> 'm) -` T) = snd -` rel_interior T" |
|
| 60420 | 10047 |
using snd_linear \<open>convex T\<close> rel_interior_convex_linear_preimage[of snd T] by auto |
| 53348 | 10048 |
then have t: "rel_interior ((UNIV :: 'n set) \<times> T) = UNIV \<times> rel_interior T" |
10049 |
by (simp add: snd_vimage_eq_Times) |
|
10050 |
from s t have *: "rel_interior (S \<times> (UNIV :: 'm set)) \<inter> rel_interior ((UNIV :: 'n set) \<times> T) = |
|
10051 |
rel_interior S \<times> rel_interior T" by auto |
|
10052 |
have "S \<times> T = S \<times> (UNIV :: 'm set) \<inter> (UNIV :: 'n set) \<times> T" |
|
10053 |
by auto |
|
10054 |
then have "rel_interior (S \<times> T) = rel_interior ((S \<times> (UNIV :: 'm set)) \<inter> ((UNIV :: 'n set) \<times> T))" |
|
10055 |
by auto |
|
10056 |
also have "\<dots> = rel_interior (S \<times> (UNIV :: 'm set)) \<inter> rel_interior ((UNIV :: 'n set) \<times> T)" |
|
| 55787 | 10057 |
apply (subst convex_rel_interior_inter_two[of "S \<times> (UNIV :: 'm set)" "(UNIV :: 'n set) \<times> T"]) |
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
10058 |
using * ri assms convex_Times |
| 53348 | 10059 |
apply auto |
10060 |
done |
|
10061 |
finally have ?thesis using * by auto |
|
10062 |
} |
|
10063 |
ultimately show ?thesis by blast |
|
| 40377 | 10064 |
qed |
10065 |
||
| 49531 | 10066 |
lemma rel_interior_scaleR: |
| 53348 | 10067 |
fixes S :: "'n::euclidean_space set" |
10068 |
assumes "c \<noteq> 0" |
|
10069 |
shows "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^sub>R c) ` S)" |
|
10070 |
using rel_interior_injective_linear_image[of "(op *\<^sub>R c)" S] |
|
10071 |
linear_conv_bounded_linear[of "op *\<^sub>R c"] linear_scaleR injective_scaleR[of c] assms |
|
10072 |
by auto |
|
| 40377 | 10073 |
|
| 49531 | 10074 |
lemma rel_interior_convex_scaleR: |
| 53348 | 10075 |
fixes S :: "'n::euclidean_space set" |
10076 |
assumes "convex S" |
|
10077 |
shows "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^sub>R c) ` S)" |
|
10078 |
by (metis assms linear_scaleR rel_interior_convex_linear_image) |
|
| 40377 | 10079 |
|
| 49531 | 10080 |
lemma convex_rel_open_scaleR: |
| 53348 | 10081 |
fixes S :: "'n::euclidean_space set" |
10082 |
assumes "convex S" |
|
10083 |
and "rel_open S" |
|
10084 |
shows "convex ((op *\<^sub>R c) ` S) \<and> rel_open ((op *\<^sub>R c) ` S)" |
|
10085 |
by (metis assms convex_scaling rel_interior_convex_scaleR rel_open_def) |
|
| 40377 | 10086 |
|
| 49531 | 10087 |
lemma convex_rel_open_finite_inter: |
| 53348 | 10088 |
assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set) \<and> rel_open S" |
10089 |
and "finite I" |
|
10090 |
shows "convex (\<Inter>I) \<and> rel_open (\<Inter>I)" |
|
| 54465 | 10091 |
proof (cases "\<Inter>{rel_interior S |S. S \<in> I} = {}")
|
| 53348 | 10092 |
case True |
10093 |
then have "\<Inter>I = {}"
|
|
10094 |
using assms unfolding rel_open_def by auto |
|
10095 |
then show ?thesis |
|
10096 |
unfolding rel_open_def using rel_interior_empty by auto |
|
10097 |
next |
|
10098 |
case False |
|
| 54465 | 10099 |
then have "rel_open (\<Inter>I)" |
| 53348 | 10100 |
using assms unfolding rel_open_def |
10101 |
using convex_rel_interior_finite_inter[of I] |
|
10102 |
by auto |
|
10103 |
then show ?thesis |
|
10104 |
using convex_Inter assms by auto |
|
| 40377 | 10105 |
qed |
10106 |
||
10107 |
lemma convex_rel_open_linear_image: |
|
| 53348 | 10108 |
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" |
10109 |
assumes "linear f" |
|
10110 |
and "convex S" |
|
10111 |
and "rel_open S" |
|
10112 |
shows "convex (f ` S) \<and> rel_open (f ` S)" |
|
| 57865 | 10113 |
by (metis assms convex_linear_image rel_interior_convex_linear_image rel_open_def) |
| 40377 | 10114 |
|
10115 |
lemma convex_rel_open_linear_preimage: |
|
| 53348 | 10116 |
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" |
10117 |
assumes "linear f" |
|
10118 |
and "convex S" |
|
10119 |
and "rel_open S" |
|
10120 |
shows "convex (f -` S) \<and> rel_open (f -` S)" |
|
10121 |
proof (cases "f -` (rel_interior S) = {}")
|
|
10122 |
case True |
|
10123 |
then have "f -` S = {}"
|
|
10124 |
using assms unfolding rel_open_def by auto |
|
10125 |
then show ?thesis |
|
10126 |
unfolding rel_open_def using rel_interior_empty by auto |
|
10127 |
next |
|
10128 |
case False |
|
10129 |
then have "rel_open (f -` S)" |
|
10130 |
using assms unfolding rel_open_def |
|
10131 |
using rel_interior_convex_linear_preimage[of f S] |
|
10132 |
by auto |
|
10133 |
then show ?thesis |
|
|
53620
3c7f5e7926dc
generalized and simplified proofs of several theorems about convex sets
huffman
parents:
53600
diff
changeset
|
10134 |
using convex_linear_vimage assms |
| 53348 | 10135 |
by auto |
| 40377 | 10136 |
qed |
10137 |
||
10138 |
lemma rel_interior_projection: |
|
| 53348 | 10139 |
fixes S :: "('m::euclidean_space \<times> 'n::euclidean_space) set"
|
10140 |
and f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space set" |
|
10141 |
assumes "convex S" |
|
10142 |
and "f = (\<lambda>y. {z. (y, z) \<in> S})"
|
|
10143 |
shows "(y, z) \<in> rel_interior S \<longleftrightarrow> (y \<in> rel_interior {y. (f y \<noteq> {})} \<and> z \<in> rel_interior (f y))"
|
|
10144 |
proof - |
|
10145 |
{
|
|
10146 |
fix y |
|
10147 |
assume "y \<in> {y. f y \<noteq> {}}"
|
|
10148 |
then obtain z where "(y, z) \<in> S" |
|
10149 |
using assms by auto |
|
10150 |
then have "\<exists>x. x \<in> S \<and> y = fst x" |
|
10151 |
apply (rule_tac x="(y, z)" in exI) |
|
10152 |
apply auto |
|
10153 |
done |
|
10154 |
then obtain x where "x \<in> S" "y = fst x" |
|
10155 |
by blast |
|
10156 |
then have "y \<in> fst ` S" |
|
10157 |
unfolding image_def by auto |
|
| 40377 | 10158 |
} |
| 53348 | 10159 |
then have "fst ` S = {y. f y \<noteq> {}}"
|
10160 |
unfolding fst_def using assms by auto |
|
10161 |
then have h1: "fst ` rel_interior S = rel_interior {y. f y \<noteq> {}}"
|
|
10162 |
using rel_interior_convex_linear_image[of fst S] assms fst_linear by auto |
|
10163 |
{
|
|
10164 |
fix y |
|
10165 |
assume "y \<in> rel_interior {y. f y \<noteq> {}}"
|
|
10166 |
then have "y \<in> fst ` rel_interior S" |
|
10167 |
using h1 by auto |
|
10168 |
then have *: "rel_interior S \<inter> fst -` {y} \<noteq> {}"
|
|
10169 |
by auto |
|
10170 |
moreover have aff: "affine (fst -` {y})"
|
|
10171 |
unfolding affine_alt by (simp add: algebra_simps) |
|
10172 |
ultimately have **: "rel_interior (S \<inter> fst -` {y}) = rel_interior S \<inter> fst -` {y}"
|
|
| 63128 | 10173 |
using convex_affine_rel_interior_Int[of S "fst -` {y}"] assms by auto
|
| 53348 | 10174 |
have conv: "convex (S \<inter> fst -` {y})"
|
10175 |
using convex_Int assms aff affine_imp_convex by auto |
|
10176 |
{
|
|
10177 |
fix x |
|
10178 |
assume "x \<in> f y" |
|
10179 |
then have "(y, x) \<in> S \<inter> (fst -` {y})"
|
|
10180 |
using assms by auto |
|
10181 |
moreover have "x = snd (y, x)" by auto |
|
10182 |
ultimately have "x \<in> snd ` (S \<inter> fst -` {y})"
|
|
10183 |
by blast |
|
10184 |
} |
|
10185 |
then have "snd ` (S \<inter> fst -` {y}) = f y"
|
|
10186 |
using assms by auto |
|
10187 |
then have ***: "rel_interior (f y) = snd ` rel_interior (S \<inter> fst -` {y})"
|
|
10188 |
using rel_interior_convex_linear_image[of snd "S \<inter> fst -` {y}"] snd_linear conv
|
|
10189 |
by auto |
|
10190 |
{
|
|
10191 |
fix z |
|
10192 |
assume "z \<in> rel_interior (f y)" |
|
10193 |
then have "z \<in> snd ` rel_interior (S \<inter> fst -` {y})"
|
|
10194 |
using *** by auto |
|
10195 |
moreover have "{y} = fst ` rel_interior (S \<inter> fst -` {y})"
|
|
10196 |
using * ** rel_interior_subset by auto |
|
10197 |
ultimately have "(y, z) \<in> rel_interior (S \<inter> fst -` {y})"
|
|
10198 |
by force |
|
10199 |
then have "(y,z) \<in> rel_interior S" |
|
10200 |
using ** by auto |
|
10201 |
} |
|
10202 |
moreover |
|
10203 |
{
|
|
10204 |
fix z |
|
10205 |
assume "(y, z) \<in> rel_interior S" |
|
10206 |
then have "(y, z) \<in> rel_interior (S \<inter> fst -` {y})"
|
|
10207 |
using ** by auto |
|
10208 |
then have "z \<in> snd ` rel_interior (S \<inter> fst -` {y})"
|
|
10209 |
by (metis Range_iff snd_eq_Range) |
|
10210 |
then have "z \<in> rel_interior (f y)" |
|
10211 |
using *** by auto |
|
10212 |
} |
|
10213 |
ultimately have "\<And>z. (y, z) \<in> rel_interior S \<longleftrightarrow> z \<in> rel_interior (f y)" |
|
10214 |
by auto |
|
| 40377 | 10215 |
} |
| 53348 | 10216 |
then have h2: "\<And>y z. y \<in> rel_interior {t. f t \<noteq> {}} \<Longrightarrow>
|
10217 |
(y, z) \<in> rel_interior S \<longleftrightarrow> z \<in> rel_interior (f y)" |
|
10218 |
by auto |
|
10219 |
{
|
|
10220 |
fix y z |
|
10221 |
assume asm: "(y, z) \<in> rel_interior S" |
|
10222 |
then have "y \<in> fst ` rel_interior S" |
|
10223 |
by (metis Domain_iff fst_eq_Domain) |
|
10224 |
then have "y \<in> rel_interior {t. f t \<noteq> {}}"
|
|
10225 |
using h1 by auto |
|
10226 |
then have "y \<in> rel_interior {t. f t \<noteq> {}}" and "(z : rel_interior (f y))"
|
|
10227 |
using h2 asm by auto |
|
| 40377 | 10228 |
} |
| 53348 | 10229 |
then show ?thesis using h2 by blast |
10230 |
qed |
|
10231 |
||
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
10232 |
lemma rel_frontier_Times: |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
10233 |
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
|
10234 |
and T :: "'m::euclidean_space set" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
10235 |
assumes "convex S" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
10236 |
and "convex T" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
10237 |
shows "rel_frontier S \<times> rel_frontier T \<subseteq> rel_frontier (S \<times> T)" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
10238 |
by (force simp: rel_frontier_def rel_interior_Times assms closure_Times) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
10239 |
|
| 40377 | 10240 |
|
| 60420 | 10241 |
subsubsection \<open>Relative interior of convex cone\<close> |
| 40377 | 10242 |
|
10243 |
lemma cone_rel_interior: |
|
| 53348 | 10244 |
fixes S :: "'m::euclidean_space set" |
10245 |
assumes "cone S" |
|
10246 |
shows "cone ({0} \<union> rel_interior S)"
|
|
10247 |
proof (cases "S = {}")
|
|
10248 |
case True |
|
10249 |
then show ?thesis |
|
10250 |
by (simp add: rel_interior_empty cone_0) |
|
10251 |
next |
|
10252 |
case False |
|
10253 |
then have *: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)" |
|
10254 |
using cone_iff[of S] assms by auto |
|
10255 |
then have *: "0 \<in> ({0} \<union> rel_interior S)"
|
|
| 54465 | 10256 |
and "\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` ({0} \<union> rel_interior S) = ({0} \<union> rel_interior S)"
|
| 53348 | 10257 |
by (auto simp add: rel_interior_scaleR) |
10258 |
then show ?thesis |
|
| 54465 | 10259 |
using cone_iff[of "{0} \<union> rel_interior S"] by auto
|
| 40377 | 10260 |
qed |
10261 |
||
10262 |
lemma rel_interior_convex_cone_aux: |
|
| 54465 | 10263 |
fixes S :: "'m::euclidean_space set" |
10264 |
assumes "convex S" |
|
| 55787 | 10265 |
shows "(c, x) \<in> rel_interior (cone hull ({(1 :: real)} \<times> S)) \<longleftrightarrow>
|
| 54465 | 10266 |
c > 0 \<and> x \<in> ((op *\<^sub>R c) ` (rel_interior S))" |
10267 |
proof (cases "S = {}")
|
|
10268 |
case True |
|
10269 |
then show ?thesis |
|
10270 |
by (simp add: rel_interior_empty cone_hull_empty) |
|
10271 |
next |
|
10272 |
case False |
|
10273 |
then obtain s where "s \<in> S" by auto |
|
| 55787 | 10274 |
have conv: "convex ({(1 :: real)} \<times> S)"
|
| 54465 | 10275 |
using convex_Times[of "{(1 :: real)}" S] assms convex_singleton[of "1 :: real"]
|
10276 |
by auto |
|
| 63040 | 10277 |
define f where "f y = {z. (y, z) \<in> cone hull ({1 :: real} \<times> S)}" for y
|
| 55787 | 10278 |
then have *: "(c, x) \<in> rel_interior (cone hull ({(1 :: real)} \<times> S)) =
|
| 54465 | 10279 |
(c \<in> rel_interior {y. f y \<noteq> {}} \<and> x \<in> rel_interior (f c))"
|
| 55787 | 10280 |
apply (subst rel_interior_projection[of "cone hull ({(1 :: real)} \<times> S)" f c x])
|
10281 |
using convex_cone_hull[of "{(1 :: real)} \<times> S"] conv
|
|
| 54465 | 10282 |
apply auto |
10283 |
done |
|
10284 |
{
|
|
10285 |
fix y :: real |
|
10286 |
assume "y \<ge> 0" |
|
| 55787 | 10287 |
then have "y *\<^sub>R (1,s) \<in> cone hull ({1 :: real} \<times> S)"
|
| 60420 | 10288 |
using cone_hull_expl[of "{(1 :: real)} \<times> S"] \<open>s \<in> S\<close> by auto
|
| 54465 | 10289 |
then have "f y \<noteq> {}"
|
10290 |
using f_def by auto |
|
10291 |
} |
|
10292 |
then have "{y. f y \<noteq> {}} = {0..}"
|
|
| 55787 | 10293 |
using f_def cone_hull_expl[of "{1 :: real} \<times> S"] by auto
|
| 54465 | 10294 |
then have **: "rel_interior {y. f y \<noteq> {}} = {0<..}"
|
10295 |
using rel_interior_real_semiline by auto |
|
10296 |
{
|
|
10297 |
fix c :: real |
|
10298 |
assume "c > 0" |
|
10299 |
then have "f c = (op *\<^sub>R c ` S)" |
|
| 55787 | 10300 |
using f_def cone_hull_expl[of "{1 :: real} \<times> S"] by auto
|
| 54465 | 10301 |
then have "rel_interior (f c) = op *\<^sub>R c ` rel_interior S" |
10302 |
using rel_interior_convex_scaleR[of S c] assms by auto |
|
10303 |
} |
|
10304 |
then show ?thesis using * ** by auto |
|
10305 |
qed |
|
| 40377 | 10306 |
|
10307 |
lemma rel_interior_convex_cone: |
|
| 54465 | 10308 |
fixes S :: "'m::euclidean_space set" |
10309 |
assumes "convex S" |
|
| 55787 | 10310 |
shows "rel_interior (cone hull ({1 :: real} \<times> S)) =
|
| 54465 | 10311 |
{(c, c *\<^sub>R x) | c x. c > 0 \<and> x \<in> rel_interior S}"
|
10312 |
(is "?lhs = ?rhs") |
|
10313 |
proof - |
|
10314 |
{
|
|
10315 |
fix z |
|
10316 |
assume "z \<in> ?lhs" |
|
10317 |
have *: "z = (fst z, snd z)" |
|
10318 |
by auto |
|
10319 |
have "z \<in> ?rhs" |
|
| 60420 | 10320 |
using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \<open>z \<in> ?lhs\<close> |
| 54465 | 10321 |
apply auto |
10322 |
apply (rule_tac x = "fst z" in exI) |
|
10323 |
apply (rule_tac x = x in exI) |
|
10324 |
using * |
|
10325 |
apply auto |
|
10326 |
done |
|
10327 |
} |
|
10328 |
moreover |
|
10329 |
{
|
|
10330 |
fix z |
|
10331 |
assume "z \<in> ?rhs" |
|
10332 |
then have "z \<in> ?lhs" |
|
10333 |
using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms |
|
10334 |
by auto |
|
10335 |
} |
|
10336 |
ultimately show ?thesis by blast |
|
| 40377 | 10337 |
qed |
10338 |
||
10339 |
lemma convex_hull_finite_union: |
|
| 54465 | 10340 |
assumes "finite I" |
10341 |
assumes "\<forall>i\<in>I. convex (S i) \<and> (S i) \<noteq> {}"
|
|
10342 |
shows "convex hull (\<Union>(S ` I)) = |
|
| 64267 | 10343 |
{sum (\<lambda>i. c i *\<^sub>R s i) I | c s. (\<forall>i\<in>I. c i \<ge> 0) \<and> sum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> S i)}"
|
| 40377 | 10344 |
(is "?lhs = ?rhs") |
| 54465 | 10345 |
proof - |
10346 |
have "?lhs \<supseteq> ?rhs" |
|
10347 |
proof |
|
10348 |
fix x |
|
10349 |
assume "x : ?rhs" |
|
| 64267 | 10350 |
then obtain c s where *: "sum (\<lambda>i. c i *\<^sub>R s i) I = x" "sum c I = 1" |
| 54465 | 10351 |
"(\<forall>i\<in>I. c i \<ge> 0) \<and> (\<forall>i\<in>I. s i \<in> S i)" by auto |
10352 |
then have "\<forall>i\<in>I. s i \<in> convex hull (\<Union>(S ` I))" |
|
10353 |
using hull_subset[of "\<Union>(S ` I)" convex] by auto |
|
10354 |
then show "x \<in> ?lhs" |
|
10355 |
unfolding *(1)[symmetric] |
|
| 64267 | 10356 |
apply (subst convex_sum[of I "convex hull \<Union>(S ` I)" c s]) |
| 54465 | 10357 |
using * assms convex_convex_hull |
10358 |
apply auto |
|
10359 |
done |
|
10360 |
qed |
|
10361 |
||
10362 |
{
|
|
10363 |
fix i |
|
10364 |
assume "i \<in> I" |
|
10365 |
with assms have "\<exists>p. p \<in> S i" by auto |
|
10366 |
} |
|
10367 |
then obtain p where p: "\<forall>i\<in>I. p i \<in> S i" by metis |
|
10368 |
||
10369 |
{
|
|
10370 |
fix i |
|
10371 |
assume "i \<in> I" |
|
10372 |
{
|
|
10373 |
fix x |
|
10374 |
assume "x \<in> S i" |
|
| 63040 | 10375 |
define c where "c j = (if j = i then 1::real else 0)" for j |
| 64267 | 10376 |
then have *: "sum c I = 1" |
10377 |
using \<open>finite I\<close> \<open>i \<in> I\<close> sum.delta[of I i "\<lambda>j::'a. 1::real"] |
|
| 54465 | 10378 |
by auto |
| 63040 | 10379 |
define s where "s j = (if j = i then x else p j)" for j |
| 54465 | 10380 |
then have "\<forall>j. c j *\<^sub>R s j = (if j = i then x else 0)" |
10381 |
using c_def by (auto simp add: algebra_simps) |
|
| 64267 | 10382 |
then have "x = sum (\<lambda>i. c i *\<^sub>R s i) I" |
10383 |
using s_def c_def \<open>finite I\<close> \<open>i \<in> I\<close> sum.delta[of I i "\<lambda>j::'a. x"] |
|
| 54465 | 10384 |
by auto |
10385 |
then have "x \<in> ?rhs" |
|
10386 |
apply auto |
|
10387 |
apply (rule_tac x = c in exI) |
|
10388 |
apply (rule_tac x = s in exI) |
|
| 60420 | 10389 |
using * c_def s_def p \<open>x \<in> S i\<close> |
| 54465 | 10390 |
apply auto |
10391 |
done |
|
| 40377 | 10392 |
} |
| 54465 | 10393 |
then have "?rhs \<supseteq> S i" by auto |
10394 |
} |
|
10395 |
then have *: "?rhs \<supseteq> \<Union>(S ` I)" by auto |
|
10396 |
||
10397 |
{
|
|
10398 |
fix u v :: real |
|
10399 |
assume uv: "u \<ge> 0 \<and> v \<ge> 0 \<and> u + v = 1" |
|
10400 |
fix x y |
|
10401 |
assume xy: "x \<in> ?rhs \<and> y \<in> ?rhs" |
|
10402 |
from xy obtain c s where |
|
| 64267 | 10403 |
xc: "x = sum (\<lambda>i. c i *\<^sub>R s i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> sum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> S i)" |
| 54465 | 10404 |
by auto |
10405 |
from xy obtain d t where |
|
| 64267 | 10406 |
yc: "y = sum (\<lambda>i. d i *\<^sub>R t i) I \<and> (\<forall>i\<in>I. d i \<ge> 0) \<and> sum d I = 1 \<and> (\<forall>i\<in>I. t i \<in> S i)" |
| 54465 | 10407 |
by auto |
| 63040 | 10408 |
define e where "e i = u * c i + v * d i" for i |
| 54465 | 10409 |
have ge0: "\<forall>i\<in>I. e i \<ge> 0" |
| 56536 | 10410 |
using e_def xc yc uv by simp |
| 64267 | 10411 |
have "sum (\<lambda>i. u * c i) I = u * sum c I" |
10412 |
by (simp add: sum_distrib_left) |
|
10413 |
moreover have "sum (\<lambda>i. v * d i) I = v * sum d I" |
|
10414 |
by (simp add: sum_distrib_left) |
|
10415 |
ultimately have sum1: "sum e I = 1" |
|
10416 |
using e_def xc yc uv by (simp add: sum.distrib) |
|
| 63040 | 10417 |
define q where "q i = (if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i)" |
10418 |
for i |
|
| 54465 | 10419 |
{
|
10420 |
fix i |
|
10421 |
assume i: "i \<in> I" |
|
10422 |
have "q i \<in> S i" |
|
10423 |
proof (cases "e i = 0") |
|
10424 |
case True |
|
10425 |
then show ?thesis using i p q_def by auto |
|
10426 |
next |
|
10427 |
case False |
|
10428 |
then show ?thesis |
|
10429 |
using mem_convex_alt[of "S i" "s i" "t i" "u * (c i)" "v * (d i)"] |
|
10430 |
mult_nonneg_nonneg[of u "c i"] mult_nonneg_nonneg[of v "d i"] |
|
10431 |
assms q_def e_def i False xc yc uv |
|
| 56536 | 10432 |
by (auto simp del: mult_nonneg_nonneg) |
| 54465 | 10433 |
qed |
10434 |
} |
|
10435 |
then have qs: "\<forall>i\<in>I. q i \<in> S i" by auto |
|
10436 |
{
|
|
10437 |
fix i |
|
10438 |
assume i: "i \<in> I" |
|
10439 |
have "(u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" |
|
10440 |
proof (cases "e i = 0") |
|
10441 |
case True |
|
10442 |
have ge: "u * (c i) \<ge> 0 \<and> v * d i \<ge> 0" |
|
| 56536 | 10443 |
using xc yc uv i by simp |
| 54465 | 10444 |
moreover from ge have "u * c i \<le> 0 \<and> v * d i \<le> 0" |
10445 |
using True e_def i by simp |
|
10446 |
ultimately have "u * c i = 0 \<and> v * d i = 0" by auto |
|
10447 |
with True show ?thesis by auto |
|
10448 |
next |
|
10449 |
case False |
|
10450 |
then have "(u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i) = q i" |
|
10451 |
using q_def by auto |
|
10452 |
then have "e i *\<^sub>R ((u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i)) |
|
10453 |
= (e i) *\<^sub>R (q i)" by auto |
|
10454 |
with False show ?thesis by (simp add: algebra_simps) |
|
10455 |
qed |
|
10456 |
} |
|
10457 |
then have *: "\<forall>i\<in>I. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" |
|
10458 |
by auto |
|
| 64267 | 10459 |
have "u *\<^sub>R x + v *\<^sub>R y = sum (\<lambda>i. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i) I" |
10460 |
using xc yc by (simp add: algebra_simps scaleR_right.sum sum.distrib) |
|
10461 |
also have "\<dots> = sum (\<lambda>i. e i *\<^sub>R q i) I" |
|
| 54465 | 10462 |
using * by auto |
| 64267 | 10463 |
finally have "u *\<^sub>R x + v *\<^sub>R y = sum (\<lambda>i. (e i) *\<^sub>R (q i)) I" |
| 54465 | 10464 |
by auto |
10465 |
then have "u *\<^sub>R x + v *\<^sub>R y \<in> ?rhs" |
|
10466 |
using ge0 sum1 qs by auto |
|
10467 |
} |
|
10468 |
then have "convex ?rhs" unfolding convex_def by auto |
|
10469 |
then show ?thesis |
|
| 60420 | 10470 |
using \<open>?lhs \<supseteq> ?rhs\<close> * hull_minimal[of "\<Union>(S ` I)" ?rhs convex] |
| 54465 | 10471 |
by blast |
| 40377 | 10472 |
qed |
10473 |
||
10474 |
lemma convex_hull_union_two: |
|
| 54465 | 10475 |
fixes S T :: "'m::euclidean_space set" |
10476 |
assumes "convex S" |
|
10477 |
and "S \<noteq> {}"
|
|
10478 |
and "convex T" |
|
10479 |
and "T \<noteq> {}"
|
|
10480 |
shows "convex hull (S \<union> T) = |
|
10481 |
{u *\<^sub>R s + v *\<^sub>R t | u v s t. u \<ge> 0 \<and> v \<ge> 0 \<and> u + v = 1 \<and> s \<in> S \<and> t \<in> T}"
|
|
| 40377 | 10482 |
(is "?lhs = ?rhs") |
| 54465 | 10483 |
proof |
| 63040 | 10484 |
define I :: "nat set" where "I = {1, 2}"
|
10485 |
define s where "s i = (if i = (1::nat) then S else T)" for i |
|
| 54465 | 10486 |
have "\<Union>(s ` I) = S \<union> T" |
10487 |
using s_def I_def by auto |
|
10488 |
then have "convex hull (\<Union>(s ` I)) = convex hull (S \<union> T)" |
|
10489 |
by auto |
|
10490 |
moreover have "convex hull \<Union>(s ` I) = |
|
| 64267 | 10491 |
{\<Sum> i\<in>I. c i *\<^sub>R sa i | c sa. (\<forall>i\<in>I. 0 \<le> c i) \<and> sum c I = 1 \<and> (\<forall>i\<in>I. sa i \<in> s i)}"
|
| 54465 | 10492 |
apply (subst convex_hull_finite_union[of I s]) |
10493 |
using assms s_def I_def |
|
10494 |
apply auto |
|
10495 |
done |
|
10496 |
moreover have |
|
| 64267 | 10497 |
"{\<Sum>i\<in>I. c i *\<^sub>R sa i | c sa. (\<forall>i\<in>I. 0 \<le> c i) \<and> sum c I = 1 \<and> (\<forall>i\<in>I. sa i \<in> s i)} \<le> ?rhs"
|
| 54465 | 10498 |
using s_def I_def by auto |
10499 |
ultimately show "?lhs \<subseteq> ?rhs" by auto |
|
10500 |
{
|
|
10501 |
fix x |
|
10502 |
assume "x \<in> ?rhs" |
|
10503 |
then obtain u v s t where *: "x = u *\<^sub>R s + v *\<^sub>R t \<and> u \<ge> 0 \<and> v \<ge> 0 \<and> u + v = 1 \<and> s \<in> S \<and> t \<in> T" |
|
10504 |
by auto |
|
10505 |
then have "x \<in> convex hull {s, t}"
|
|
10506 |
using convex_hull_2[of s t] by auto |
|
10507 |
then have "x \<in> convex hull (S \<union> T)" |
|
10508 |
using * hull_mono[of "{s, t}" "S \<union> T"] by auto
|
|
10509 |
} |
|
10510 |
then show "?lhs \<supseteq> ?rhs" by blast |
|
10511 |
qed |
|
10512 |
||
| 40377 | 10513 |
|
| 60420 | 10514 |
subsection \<open>Convexity on direct sums\<close> |
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10515 |
|
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10516 |
lemma closure_sum: |
| 55928 | 10517 |
fixes S T :: "'a::real_normed_vector set" |
|
47445
69e96e5500df
Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents:
47444
diff
changeset
|
10518 |
shows "closure S + closure T \<subseteq> closure (S + T)" |
| 55928 | 10519 |
unfolding set_plus_image closure_Times [symmetric] split_def |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
10520 |
by (intro closure_bounded_linear_image_subset bounded_linear_add |
| 55928 | 10521 |
bounded_linear_fst bounded_linear_snd) |
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10522 |
|
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10523 |
lemma rel_interior_sum: |
| 54465 | 10524 |
fixes S T :: "'n::euclidean_space set" |
10525 |
assumes "convex S" |
|
10526 |
and "convex T" |
|
10527 |
shows "rel_interior (S + T) = rel_interior S + rel_interior T" |
|
10528 |
proof - |
|
| 55787 | 10529 |
have "rel_interior S + rel_interior T = (\<lambda>(x,y). x + y) ` (rel_interior S \<times> rel_interior T)" |
| 54465 | 10530 |
by (simp add: set_plus_image) |
| 55787 | 10531 |
also have "\<dots> = (\<lambda>(x,y). x + y) ` rel_interior (S \<times> T)" |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
10532 |
using rel_interior_Times assms by auto |
| 54465 | 10533 |
also have "\<dots> = rel_interior (S + T)" |
10534 |
using fst_snd_linear convex_Times assms |
|
| 55787 | 10535 |
rel_interior_convex_linear_image[of "(\<lambda>(x,y). x + y)" "S \<times> T"] |
| 54465 | 10536 |
by (auto simp add: set_plus_image) |
10537 |
finally show ?thesis .. |
|
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10538 |
qed |
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10539 |
|
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10540 |
lemma rel_interior_sum_gen: |
| 54465 | 10541 |
fixes S :: "'a \<Rightarrow> 'n::euclidean_space set" |
10542 |
assumes "\<forall>i\<in>I. convex (S i)" |
|
| 64267 | 10543 |
shows "rel_interior (sum S I) = sum (\<lambda>i. rel_interior (S i)) I" |
10544 |
apply (subst sum_set_cond_linear[of convex]) |
|
| 54465 | 10545 |
using rel_interior_sum rel_interior_sing[of "0"] assms |
|
55929
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents:
55928
diff
changeset
|
10546 |
apply (auto simp add: convex_set_plus) |
| 54465 | 10547 |
done |
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10548 |
|
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10549 |
lemma convex_rel_open_direct_sum: |
| 54465 | 10550 |
fixes S T :: "'n::euclidean_space set" |
10551 |
assumes "convex S" |
|
10552 |
and "rel_open S" |
|
10553 |
and "convex T" |
|
10554 |
and "rel_open T" |
|
| 55787 | 10555 |
shows "convex (S \<times> T) \<and> rel_open (S \<times> T)" |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62950
diff
changeset
|
10556 |
by (metis assms convex_Times rel_interior_Times rel_open_def) |
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10557 |
|
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10558 |
lemma convex_rel_open_sum: |
| 54465 | 10559 |
fixes S T :: "'n::euclidean_space set" |
10560 |
assumes "convex S" |
|
10561 |
and "rel_open S" |
|
10562 |
and "convex T" |
|
10563 |
and "rel_open T" |
|
10564 |
shows "convex (S + T) \<and> rel_open (S + T)" |
|
|
55929
91f245c23bc5
remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents:
55928
diff
changeset
|
10565 |
by (metis assms convex_set_plus rel_interior_sum rel_open_def) |
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10566 |
|
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10567 |
lemma convex_hull_finite_union_cones: |
| 54465 | 10568 |
assumes "finite I" |
10569 |
and "I \<noteq> {}"
|
|
10570 |
assumes "\<forall>i\<in>I. convex (S i) \<and> cone (S i) \<and> S i \<noteq> {}"
|
|
| 64267 | 10571 |
shows "convex hull (\<Union>(S ` I)) = sum S I" |
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10572 |
(is "?lhs = ?rhs") |
| 54465 | 10573 |
proof - |
10574 |
{
|
|
10575 |
fix x |
|
10576 |
assume "x \<in> ?lhs" |
|
10577 |
then obtain c xs where |
|
| 64267 | 10578 |
x: "x = sum (\<lambda>i. c i *\<^sub>R xs i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> sum c I = 1 \<and> (\<forall>i\<in>I. xs i \<in> S i)" |
| 54465 | 10579 |
using convex_hull_finite_union[of I S] assms by auto |
| 63040 | 10580 |
define s where "s i = c i *\<^sub>R xs i" for i |
| 54465 | 10581 |
{
|
10582 |
fix i |
|
10583 |
assume "i \<in> I" |
|
10584 |
then have "s i \<in> S i" |
|
10585 |
using s_def x assms mem_cone[of "S i" "xs i" "c i"] by auto |
|
10586 |
} |
|
10587 |
then have "\<forall>i\<in>I. s i \<in> S i" by auto |
|
| 64267 | 10588 |
moreover have "x = sum s I" using x s_def by auto |
| 54465 | 10589 |
ultimately have "x \<in> ?rhs" |
| 64267 | 10590 |
using set_sum_alt[of I S] assms by auto |
| 54465 | 10591 |
} |
10592 |
moreover |
|
10593 |
{
|
|
10594 |
fix x |
|
10595 |
assume "x \<in> ?rhs" |
|
| 64267 | 10596 |
then obtain s where x: "x = sum s I \<and> (\<forall>i\<in>I. s i \<in> S i)" |
10597 |
using set_sum_alt[of I S] assms by auto |
|
| 63040 | 10598 |
define xs where "xs i = of_nat(card I) *\<^sub>R s i" for i |
| 64267 | 10599 |
then have "x = sum (\<lambda>i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I" |
| 54465 | 10600 |
using x assms by auto |
10601 |
moreover have "\<forall>i\<in>I. xs i \<in> S i" |
|
10602 |
using x xs_def assms by (simp add: cone_def) |
|
10603 |
moreover have "\<forall>i\<in>I. (1 :: real) / of_nat (card I) \<ge> 0" |
|
10604 |
by auto |
|
| 64267 | 10605 |
moreover have "sum (\<lambda>i. (1 :: real) / of_nat (card I)) I = 1" |
| 54465 | 10606 |
using assms by auto |
10607 |
ultimately have "x \<in> ?lhs" |
|
10608 |
apply (subst convex_hull_finite_union[of I S]) |
|
10609 |
using assms |
|
10610 |
apply blast |
|
10611 |
using assms |
|
10612 |
apply blast |
|
10613 |
apply rule |
|
10614 |
apply (rule_tac x = "(\<lambda>i. (1 :: real) / of_nat (card I))" in exI) |
|
10615 |
apply auto |
|
10616 |
done |
|
10617 |
} |
|
10618 |
ultimately show ?thesis by auto |
|
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10619 |
qed |
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10620 |
|
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10621 |
lemma convex_hull_union_cones_two: |
| 54465 | 10622 |
fixes S T :: "'m::euclidean_space set" |
10623 |
assumes "convex S" |
|
10624 |
and "cone S" |
|
10625 |
and "S \<noteq> {}"
|
|
10626 |
assumes "convex T" |
|
10627 |
and "cone T" |
|
10628 |
and "T \<noteq> {}"
|
|
10629 |
shows "convex hull (S \<union> T) = S + T" |
|
10630 |
proof - |
|
| 63040 | 10631 |
define I :: "nat set" where "I = {1, 2}"
|
10632 |
define A where "A i = (if i = (1::nat) then S else T)" for i |
|
| 54465 | 10633 |
have "\<Union>(A ` I) = S \<union> T" |
10634 |
using A_def I_def by auto |
|
10635 |
then have "convex hull (\<Union>(A ` I)) = convex hull (S \<union> T)" |
|
10636 |
by auto |
|
| 64267 | 10637 |
moreover have "convex hull \<Union>(A ` I) = sum A I" |
| 54465 | 10638 |
apply (subst convex_hull_finite_union_cones[of I A]) |
10639 |
using assms A_def I_def |
|
10640 |
apply auto |
|
10641 |
done |
|
| 64267 | 10642 |
moreover have "sum A I = S + T" |
| 54465 | 10643 |
using A_def I_def |
10644 |
unfolding set_plus_def |
|
10645 |
apply auto |
|
10646 |
unfolding set_plus_def |
|
10647 |
apply auto |
|
10648 |
done |
|
10649 |
ultimately show ?thesis by auto |
|
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10650 |
qed |
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10651 |
|
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10652 |
lemma rel_interior_convex_hull_union: |
| 54465 | 10653 |
fixes S :: "'a \<Rightarrow> 'n::euclidean_space set" |
10654 |
assumes "finite I" |
|
10655 |
and "\<forall>i\<in>I. convex (S i) \<and> S i \<noteq> {}"
|
|
10656 |
shows "rel_interior (convex hull (\<Union>(S ` I))) = |
|
| 64267 | 10657 |
{sum (\<lambda>i. c i *\<^sub>R s i) I | c s. (\<forall>i\<in>I. c i > 0) \<and> sum c I = 1 \<and>
|
| 54465 | 10658 |
(\<forall>i\<in>I. s i \<in> rel_interior(S i))}" |
10659 |
(is "?lhs = ?rhs") |
|
10660 |
proof (cases "I = {}")
|
|
10661 |
case True |
|
10662 |
then show ?thesis |
|
10663 |
using convex_hull_empty rel_interior_empty by auto |
|
10664 |
next |
|
10665 |
case False |
|
| 63040 | 10666 |
define C0 where "C0 = convex hull (\<Union>(S ` I))" |
| 54465 | 10667 |
have "\<forall>i\<in>I. C0 \<ge> S i" |
10668 |
unfolding C0_def using hull_subset[of "\<Union>(S ` I)"] by auto |
|
| 63040 | 10669 |
define K0 where "K0 = cone hull ({1 :: real} \<times> C0)"
|
10670 |
define K where "K i = cone hull ({1 :: real} \<times> S i)" for i
|
|
| 54465 | 10671 |
have "\<forall>i\<in>I. K i \<noteq> {}"
|
10672 |
unfolding K_def using assms |
|
10673 |
by (simp add: cone_hull_empty_iff[symmetric]) |
|
10674 |
{
|
|
10675 |
fix i |
|
10676 |
assume "i \<in> I" |
|
10677 |
then have "convex (K i)" |
|
10678 |
unfolding K_def |
|
10679 |
apply (subst convex_cone_hull) |
|
10680 |
apply (subst convex_Times) |
|
10681 |
using assms |
|
10682 |
apply auto |
|
10683 |
done |
|
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10684 |
} |
| 54465 | 10685 |
then have convK: "\<forall>i\<in>I. convex (K i)" |
10686 |
by auto |
|
10687 |
{
|
|
10688 |
fix i |
|
10689 |
assume "i \<in> I" |
|
10690 |
then have "K0 \<supseteq> K i" |
|
10691 |
unfolding K0_def K_def |
|
10692 |
apply (subst hull_mono) |
|
| 60420 | 10693 |
using \<open>\<forall>i\<in>I. C0 \<ge> S i\<close> |
| 54465 | 10694 |
apply auto |
10695 |
done |
|
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10696 |
} |
| 54465 | 10697 |
then have "K0 \<supseteq> \<Union>(K ` I)" by auto |
10698 |
moreover have "convex K0" |
|
10699 |
unfolding K0_def |
|
10700 |
apply (subst convex_cone_hull) |
|
10701 |
apply (subst convex_Times) |
|
10702 |
unfolding C0_def |
|
10703 |
using convex_convex_hull |
|
10704 |
apply auto |
|
10705 |
done |
|
10706 |
ultimately have geq: "K0 \<supseteq> convex hull (\<Union>(K ` I))" |
|
10707 |
using hull_minimal[of _ "K0" "convex"] by blast |
|
| 55787 | 10708 |
have "\<forall>i\<in>I. K i \<supseteq> {1 :: real} \<times> S i"
|
| 54465 | 10709 |
using K_def by (simp add: hull_subset) |
| 55787 | 10710 |
then have "\<Union>(K ` I) \<supseteq> {1 :: real} \<times> \<Union>(S ` I)"
|
| 54465 | 10711 |
by auto |
| 55787 | 10712 |
then have "convex hull \<Union>(K ` I) \<supseteq> convex hull ({1 :: real} \<times> \<Union>(S ` I))"
|
| 54465 | 10713 |
by (simp add: hull_mono) |
| 55787 | 10714 |
then have "convex hull \<Union>(K ` I) \<supseteq> {1 :: real} \<times> C0"
|
| 54465 | 10715 |
unfolding C0_def |
10716 |
using convex_hull_Times[of "{(1 :: real)}" "\<Union>(S ` I)"] convex_hull_singleton
|
|
10717 |
by auto |
|
10718 |
moreover have "cone (convex hull (\<Union>(K ` I)))" |
|
10719 |
apply (subst cone_convex_hull) |
|
10720 |
using cone_Union[of "K ` I"] |
|
10721 |
apply auto |
|
10722 |
unfolding K_def |
|
10723 |
using cone_cone_hull |
|
10724 |
apply auto |
|
10725 |
done |
|
10726 |
ultimately have "convex hull (\<Union>(K ` I)) \<supseteq> K0" |
|
10727 |
unfolding K0_def |
|
| 60585 | 10728 |
using hull_minimal[of _ "convex hull (\<Union>(K ` I))" "cone"] |
| 54465 | 10729 |
by blast |
10730 |
then have "K0 = convex hull (\<Union>(K ` I))" |
|
10731 |
using geq by auto |
|
| 64267 | 10732 |
also have "\<dots> = sum K I" |
| 54465 | 10733 |
apply (subst convex_hull_finite_union_cones[of I K]) |
10734 |
using assms |
|
10735 |
apply blast |
|
10736 |
using False |
|
10737 |
apply blast |
|
10738 |
unfolding K_def |
|
10739 |
apply rule |
|
10740 |
apply (subst convex_cone_hull) |
|
10741 |
apply (subst convex_Times) |
|
| 60420 | 10742 |
using assms cone_cone_hull \<open>\<forall>i\<in>I. K i \<noteq> {}\<close> K_def
|
| 54465 | 10743 |
apply auto |
10744 |
done |
|
| 64267 | 10745 |
finally have "K0 = sum K I" by auto |
10746 |
then have *: "rel_interior K0 = sum (\<lambda>i. (rel_interior (K i))) I" |
|
| 54465 | 10747 |
using rel_interior_sum_gen[of I K] convK by auto |
10748 |
{
|
|
10749 |
fix x |
|
10750 |
assume "x \<in> ?lhs" |
|
10751 |
then have "(1::real, x) \<in> rel_interior K0" |
|
10752 |
using K0_def C0_def rel_interior_convex_cone_aux[of C0 "1::real" x] convex_convex_hull |
|
10753 |
by auto |
|
| 64267 | 10754 |
then obtain k where k: "(1::real, x) = sum k I \<and> (\<forall>i\<in>I. k i \<in> rel_interior (K i))" |
10755 |
using \<open>finite I\<close> * set_sum_alt[of I "\<lambda>i. rel_interior (K i)"] by auto |
|
| 54465 | 10756 |
{
|
10757 |
fix i |
|
10758 |
assume "i \<in> I" |
|
| 55787 | 10759 |
then have "convex (S i) \<and> k i \<in> rel_interior (cone hull {1} \<times> S i)"
|
| 54465 | 10760 |
using k K_def assms by auto |
10761 |
then have "\<exists>ci si. k i = (ci, ci *\<^sub>R si) \<and> 0 < ci \<and> si \<in> rel_interior (S i)" |
|
10762 |
using rel_interior_convex_cone[of "S i"] by auto |
|
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10763 |
} |
| 54465 | 10764 |
then obtain c s where |
10765 |
cs: "\<forall>i\<in>I. k i = (c i, c i *\<^sub>R s i) \<and> 0 < c i \<and> s i \<in> rel_interior (S i)" |
|
10766 |
by metis |
|
| 64267 | 10767 |
then have "x = (\<Sum>i\<in>I. c i *\<^sub>R s i) \<and> sum c I = 1" |
10768 |
using k by (simp add: sum_prod) |
|
| 54465 | 10769 |
then have "x \<in> ?rhs" |
10770 |
using k |
|
10771 |
apply auto |
|
10772 |
apply (rule_tac x = c in exI) |
|
10773 |
apply (rule_tac x = s in exI) |
|
10774 |
using cs |
|
10775 |
apply auto |
|
10776 |
done |
|
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10777 |
} |
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10778 |
moreover |
| 54465 | 10779 |
{
|
10780 |
fix x |
|
10781 |
assume "x \<in> ?rhs" |
|
| 64267 | 10782 |
then obtain c s where cs: "x = sum (\<lambda>i. c i *\<^sub>R s i) I \<and> |
10783 |
(\<forall>i\<in>I. c i > 0) \<and> sum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> rel_interior (S i))" |
|
| 54465 | 10784 |
by auto |
| 63040 | 10785 |
define k where "k i = (c i, c i *\<^sub>R s i)" for i |
| 54465 | 10786 |
{
|
10787 |
fix i assume "i:I" |
|
10788 |
then have "k i \<in> rel_interior (K i)" |
|
10789 |
using k_def K_def assms cs rel_interior_convex_cone[of "S i"] |
|
10790 |
by auto |
|
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10791 |
} |
| 54465 | 10792 |
then have "(1::real, x) \<in> rel_interior K0" |
| 64267 | 10793 |
using K0_def * set_sum_alt[of I "(\<lambda>i. rel_interior (K i))"] assms k_def cs |
| 54465 | 10794 |
apply auto |
10795 |
apply (rule_tac x = k in exI) |
|
| 64267 | 10796 |
apply (simp add: sum_prod) |
| 54465 | 10797 |
done |
10798 |
then have "x \<in> ?lhs" |
|
10799 |
using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x] |
|
10800 |
by (auto simp add: convex_convex_hull) |
|
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10801 |
} |
| 54465 | 10802 |
ultimately show ?thesis by blast |
|
40887
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10803 |
qed |
|
ee8d0548c148
Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents:
40719
diff
changeset
|
10804 |
|
| 50104 | 10805 |
|
10806 |
lemma convex_le_Inf_differential: |
|
10807 |
fixes f :: "real \<Rightarrow> real" |
|
10808 |
assumes "convex_on I f" |
|
| 54465 | 10809 |
and "x \<in> interior I" |
10810 |
and "y \<in> I" |
|
| 50104 | 10811 |
shows "f y \<ge> f x + Inf ((\<lambda>t. (f x - f t) / (x - t)) ` ({x<..} \<inter> I)) * (y - x)"
|
| 54465 | 10812 |
(is "_ \<ge> _ + Inf (?F x) * (y - x)") |
| 50104 | 10813 |
proof (cases rule: linorder_cases) |
10814 |
assume "x < y" |
|
10815 |
moreover |
|
10816 |
have "open (interior I)" by auto |
|
| 60420 | 10817 |
from openE[OF this \<open>x \<in> interior I\<close>] |
| 55697 | 10818 |
obtain e where e: "0 < e" "ball x e \<subseteq> interior I" . |
| 63040 | 10819 |
moreover define t where "t = min (x + e / 2) ((x + y) / 2)" |
| 50104 | 10820 |
ultimately have "x < t" "t < y" "t \<in> ball x e" |
10821 |
by (auto simp: dist_real_def field_simps split: split_min) |
|
| 60420 | 10822 |
with \<open>x \<in> interior I\<close> e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto |
| 50104 | 10823 |
|
10824 |
have "open (interior I)" by auto |
|
| 60420 | 10825 |
from openE[OF this \<open>x \<in> interior I\<close>] |
| 55697 | 10826 |
obtain e where "0 < e" "ball x e \<subseteq> interior I" . |
| 63040 | 10827 |
moreover define K where "K = x - e / 2" |
| 60420 | 10828 |
with \<open>0 < e\<close> have "K \<in> ball x e" "K < x" |
| 54465 | 10829 |
by (auto simp: dist_real_def) |
| 50104 | 10830 |
ultimately have "K \<in> I" "K < x" "x \<in> I" |
| 60420 | 10831 |
using interior_subset[of I] \<open>x \<in> interior I\<close> by auto |
| 50104 | 10832 |
|
10833 |
have "Inf (?F x) \<le> (f x - f y) / (x - y)" |
|
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
10834 |
proof (intro bdd_belowI cInf_lower2) |
| 50104 | 10835 |
show "(f x - f t) / (x - t) \<in> ?F x" |
| 60420 | 10836 |
using \<open>t \<in> I\<close> \<open>x < t\<close> by auto |
| 50104 | 10837 |
show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)" |
| 60420 | 10838 |
using \<open>convex_on I f\<close> \<open>x \<in> I\<close> \<open>y \<in> I\<close> \<open>x < t\<close> \<open>t < y\<close> |
| 54465 | 10839 |
by (rule convex_on_diff) |
| 50104 | 10840 |
next |
| 54465 | 10841 |
fix y |
10842 |
assume "y \<in> ?F x" |
|
| 60420 | 10843 |
with order_trans[OF convex_on_diff[OF \<open>convex_on I f\<close> \<open>K \<in> I\<close> _ \<open>K < x\<close> _]] |
| 50104 | 10844 |
show "(f K - f x) / (K - x) \<le> y" by auto |
10845 |
qed |
|
10846 |
then show ?thesis |
|
| 60420 | 10847 |
using \<open>x < y\<close> by (simp add: field_simps) |
| 50104 | 10848 |
next |
10849 |
assume "y < x" |
|
10850 |
moreover |
|
10851 |
have "open (interior I)" by auto |
|
| 60420 | 10852 |
from openE[OF this \<open>x \<in> interior I\<close>] |
| 55697 | 10853 |
obtain e where e: "0 < e" "ball x e \<subseteq> interior I" . |
| 63040 | 10854 |
moreover define t where "t = x + e / 2" |
| 50104 | 10855 |
ultimately have "x < t" "t \<in> ball x e" |
10856 |
by (auto simp: dist_real_def field_simps) |
|
| 60420 | 10857 |
with \<open>x \<in> interior I\<close> e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto |
| 50104 | 10858 |
|
10859 |
have "(f x - f y) / (x - y) \<le> Inf (?F x)" |
|
|
51475
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
hoelzl
parents:
50979
diff
changeset
|
10860 |
proof (rule cInf_greatest) |
| 50104 | 10861 |
have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" |
| 60420 | 10862 |
using \<open>y < x\<close> by (auto simp: field_simps) |
| 50104 | 10863 |
also |
| 54465 | 10864 |
fix z |
10865 |
assume "z \<in> ?F x" |
|
| 60420 | 10866 |
with order_trans[OF convex_on_diff[OF \<open>convex_on I f\<close> \<open>y \<in> I\<close> _ \<open>y < x\<close>]] |
| 54465 | 10867 |
have "(f y - f x) / (y - x) \<le> z" |
10868 |
by auto |
|
| 50104 | 10869 |
finally show "(f x - f y) / (x - y) \<le> z" . |
10870 |
next |
|
10871 |
have "open (interior I)" by auto |
|
| 60420 | 10872 |
from openE[OF this \<open>x \<in> interior I\<close>] |
| 55697 | 10873 |
obtain e where e: "0 < e" "ball x e \<subseteq> interior I" . |
| 54465 | 10874 |
then have "x + e / 2 \<in> ball x e" |
10875 |
by (auto simp: dist_real_def) |
|
10876 |
with e interior_subset[of I] have "x + e / 2 \<in> {x<..} \<inter> I"
|
|
10877 |
by auto |
|
10878 |
then show "?F x \<noteq> {}"
|
|
10879 |
by blast |
|
| 50104 | 10880 |
qed |
10881 |
then show ?thesis |
|
| 60420 | 10882 |
using \<open>y < x\<close> by (simp add: field_simps) |
| 50104 | 10883 |
qed simp |
| 60762 | 10884 |
|
| 60420 | 10885 |
subsection\<open>Explicit formulas for interior and relative interior of convex hull\<close> |
| 60762 | 10886 |
|
10887 |
lemma interior_atLeastAtMost [simp]: |
|
10888 |
fixes a::real shows "interior {a..b} = {a<..<b}"
|
|
10889 |
using interior_cbox [of a b] by auto |
|
10890 |
||
10891 |
lemma interior_atLeastLessThan [simp]: |
|
10892 |
fixes a::real shows "interior {a..<b} = {a<..<b}"
|
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
10893 |
by (metis atLeastLessThan_def greaterThanLessThan_def interior_atLeastAtMost interior_Int interior_interior interior_real_semiline) |
| 60762 | 10894 |
|
10895 |
lemma interior_lessThanAtMost [simp]: |
|
10896 |
fixes a::real shows "interior {a<..b} = {a<..<b}"
|
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
10897 |
by (metis atLeastAtMost_def greaterThanAtMost_def interior_atLeastAtMost interior_Int |
| 60762 | 10898 |
interior_interior interior_real_semiline) |
10899 |
||
10900 |
lemma at_within_closed_interval: |
|
10901 |
fixes x::real |
|
10902 |
shows "a < x \<Longrightarrow> x < b \<Longrightarrow> (at x within {a..b}) = at x"
|
|
10903 |
by (metis at_within_interior greaterThanLessThan_iff interior_atLeastAtMost) |
|
10904 |
||
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10905 |
lemma affine_independent_convex_affine_hull: |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10906 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10907 |
assumes "~affine_dependent s" "t \<subseteq> s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10908 |
shows "convex hull t = affine hull t \<inter> convex hull s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10909 |
proof - |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10910 |
have fin: "finite s" "finite t" using assms aff_independent_finite finite_subset by auto |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10911 |
{ fix u v x
|
| 64267 | 10912 |
assume uv: "sum u t = 1" "\<forall>x\<in>s. 0 \<le> v x" "sum v s = 1" |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10913 |
"(\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>v\<in>t. u v *\<^sub>R v)" "x \<in> t" |
| 61808 | 10914 |
then have s: "s = (s - t) \<union> t" \<comment>\<open>split into separate cases\<close> |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10915 |
using assms by auto |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10916 |
have [simp]: "(\<Sum>x\<in>t. v x *\<^sub>R x) + (\<Sum>x\<in>s - t. v x *\<^sub>R x) = (\<Sum>x\<in>t. u x *\<^sub>R x)" |
| 64267 | 10917 |
"sum v t + sum v (s - t) = 1" |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10918 |
using uv fin s |
| 64267 | 10919 |
by (auto simp: sum.union_disjoint [symmetric] Un_commute) |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
10920 |
have "(\<Sum>x\<in>s. if x \<in> t then v x - u x else v x) = 0" |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10921 |
"(\<Sum>x\<in>s. (if x \<in> t then v x - u x else v x) *\<^sub>R x) = 0" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10922 |
using uv fin |
| 64267 | 10923 |
by (subst s, subst sum.union_disjoint, auto simp: algebra_simps sum_subtractf)+ |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10924 |
} note [simp] = this |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
10925 |
have "convex hull t \<subseteq> affine hull t" |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10926 |
using convex_hull_subset_affine_hull by blast |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10927 |
moreover have "convex hull t \<subseteq> convex hull s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10928 |
using assms hull_mono by blast |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10929 |
moreover have "affine hull t \<inter> convex hull s \<subseteq> convex hull t" |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
10930 |
using assms |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10931 |
apply (simp add: convex_hull_finite affine_hull_finite fin affine_dependent_explicit) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10932 |
apply (drule_tac x=s in spec) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10933 |
apply (auto simp: fin) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10934 |
apply (rule_tac x=u in exI) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10935 |
apply (rename_tac v) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10936 |
apply (drule_tac x="\<lambda>x. if x \<in> t then v x - u x else v x" in spec) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10937 |
apply (force)+ |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10938 |
done |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10939 |
ultimately show ?thesis |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10940 |
by blast |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10941 |
qed |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10942 |
|
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
10943 |
lemma affine_independent_span_eq: |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10944 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10945 |
assumes "~affine_dependent s" "card s = Suc (DIM ('a))"
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10946 |
shows "affine hull s = UNIV" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10947 |
proof (cases "s = {}")
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10948 |
case True then show ?thesis |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10949 |
using assms by simp |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10950 |
next |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10951 |
case False |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10952 |
then obtain a t where t: "a \<notin> t" "s = insert a t" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10953 |
by blast |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
10954 |
then have fin: "finite t" using assms |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10955 |
by (metis finite_insert aff_independent_finite) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10956 |
show ?thesis |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10957 |
using assms t fin |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10958 |
apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10959 |
apply (rule subset_antisym) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10960 |
apply force |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10961 |
apply (rule Fun.vimage_subsetD) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10962 |
apply (metis add.commute diff_add_cancel surj_def) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10963 |
apply (rule card_ge_dim_independent) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10964 |
apply (auto simp: card_image inj_on_def dim_subset_UNIV) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10965 |
done |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10966 |
qed |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10967 |
|
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
10968 |
lemma affine_independent_span_gt: |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10969 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10970 |
assumes ind: "~ affine_dependent s" and dim: "DIM ('a) < card s"
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10971 |
shows "affine hull s = UNIV" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10972 |
apply (rule affine_independent_span_eq [OF ind]) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10973 |
apply (rule antisym) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10974 |
using assms |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10975 |
apply auto |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10976 |
apply (metis add_2_eq_Suc' not_less_eq_eq affine_dependent_biggerset aff_independent_finite) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10977 |
done |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10978 |
|
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
10979 |
lemma empty_interior_affine_hull: |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10980 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10981 |
assumes "finite s" and dim: "card s \<le> DIM ('a)"
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10982 |
shows "interior(affine hull s) = {}"
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10983 |
using assms |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10984 |
apply (induct s rule: finite_induct) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10985 |
apply (simp_all add: affine_dependent_iff_dependent affine_hull_insert_span_gen interior_translation) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10986 |
apply (rule empty_interior_lowdim) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10987 |
apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10988 |
apply (metis Suc_le_lessD not_less order_trans card_image_le finite_imageI dim_le_card) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10989 |
done |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10990 |
|
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
10991 |
lemma empty_interior_convex_hull: |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10992 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10993 |
assumes "finite s" and dim: "card s \<le> DIM ('a)"
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10994 |
shows "interior(convex hull s) = {}"
|
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
10995 |
by (metis Diff_empty Diff_eq_empty_iff convex_hull_subset_affine_hull |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10996 |
interior_mono empty_interior_affine_hull [OF assms]) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10997 |
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10998 |
lemma explicit_subset_rel_interior_convex_hull: |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
10999 |
fixes s :: "'a::euclidean_space set" |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11000 |
shows "finite s |
| 64267 | 11001 |
\<Longrightarrow> {y. \<exists>u. (\<forall>x \<in> s. 0 < u x \<and> u x < 1) \<and> sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11002 |
\<subseteq> rel_interior (convex hull s)" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11003 |
by (force simp add: rel_interior_convex_hull_union [where S="\<lambda>x. {x}" and I=s, simplified])
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11004 |
|
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11005 |
lemma explicit_subset_rel_interior_convex_hull_minimal: |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11006 |
fixes s :: "'a::euclidean_space set" |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11007 |
shows "finite s |
| 64267 | 11008 |
\<Longrightarrow> {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11009 |
\<subseteq> rel_interior (convex hull s)" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11010 |
by (force simp add: rel_interior_convex_hull_union [where S="\<lambda>x. {x}" and I=s, simplified])
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11011 |
|
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11012 |
lemma rel_interior_convex_hull_explicit: |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11013 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11014 |
assumes "~ affine_dependent s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11015 |
shows "rel_interior(convex hull s) = |
| 64267 | 11016 |
{y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}"
|
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11017 |
(is "?lhs = ?rhs") |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11018 |
proof |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11019 |
show "?rhs \<le> ?lhs" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11020 |
by (simp add: aff_independent_finite explicit_subset_rel_interior_convex_hull_minimal assms) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11021 |
next |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11022 |
show "?lhs \<le> ?rhs" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11023 |
proof (cases "\<exists>a. s = {a}")
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11024 |
case True then show "?lhs \<le> ?rhs" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11025 |
by force |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11026 |
next |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11027 |
case False |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11028 |
have fs: "finite s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11029 |
using assms by (simp add: aff_independent_finite) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11030 |
{ fix a b and d::real
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11031 |
assume ab: "a \<in> s" "b \<in> s" "a \<noteq> b" |
| 61808 | 11032 |
then have s: "s = (s - {a,b}) \<union> {a,b}" \<comment>\<open>split into separate cases\<close>
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11033 |
by auto |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11034 |
have "(\<Sum>x\<in>s. if x = a then - d else if x = b then d else 0) = 0" |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11035 |
"(\<Sum>x\<in>s. (if x = a then - d else if x = b then d else 0) *\<^sub>R x) = d *\<^sub>R b - d *\<^sub>R a" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11036 |
using ab fs |
| 64267 | 11037 |
by (subst s, subst sum.union_disjoint, auto)+ |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11038 |
} note [simp] = this |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11039 |
{ fix y
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11040 |
assume y: "y \<in> convex hull s" "y \<notin> ?rhs" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11041 |
{ fix u T a
|
| 64267 | 11042 |
assume ua: "\<forall>x\<in>s. 0 \<le> u x" "sum u s = 1" "\<not> 0 < u a" "a \<in> s" |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11043 |
and yT: "y = (\<Sum>x\<in>s. u x *\<^sub>R x)" "y \<in> T" "open T" |
| 64267 | 11044 |
and sb: "T \<inter> affine hull s \<subseteq> {w. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = w}"
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11045 |
have ua0: "u a = 0" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11046 |
using ua by auto |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11047 |
obtain b where b: "b\<in>s" "a \<noteq> b" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11048 |
using ua False by auto |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11049 |
obtain e where e: "0 < e" "ball (\<Sum>x\<in>s. u x *\<^sub>R x) e \<subseteq> T" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11050 |
using yT by (auto elim: openE) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11051 |
with b obtain d where d: "0 < d" "norm(d *\<^sub>R (a-b)) < e" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11052 |
by (auto intro: that [of "e / 2 / norm(a-b)"]) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11053 |
have "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> affine hull s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11054 |
using yT y by (metis affine_hull_convex_hull hull_redundant_eq) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11055 |
then have "(\<Sum>x\<in>s. u x *\<^sub>R x) - d *\<^sub>R (a - b) \<in> affine hull s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11056 |
using ua b by (auto simp: hull_inc intro: mem_affine_3_minus2) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11057 |
then have "y - d *\<^sub>R (a - b) \<in> T \<inter> affine hull s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11058 |
using d e yT by auto |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11059 |
then obtain v where "\<forall>x\<in>s. 0 \<le> v x" |
| 64267 | 11060 |
"sum v s = 1" |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11061 |
"(\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x) - d *\<^sub>R (a - b)" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11062 |
using subsetD [OF sb] yT |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11063 |
by auto |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11064 |
then have False |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11065 |
using assms |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11066 |
apply (simp add: affine_dependent_explicit_finite fs) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11067 |
apply (drule_tac x="\<lambda>x. (v x - u x) - (if x = a then -d else if x = b then d else 0)" in spec) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11068 |
using ua b d |
| 64267 | 11069 |
apply (auto simp: algebra_simps sum_subtractf sum.distrib) |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11070 |
done |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11071 |
} note * = this |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11072 |
have "y \<notin> rel_interior (convex hull s)" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11073 |
using y |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11074 |
apply (simp add: mem_rel_interior affine_hull_convex_hull) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11075 |
apply (auto simp: convex_hull_finite [OF fs]) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11076 |
apply (drule_tac x=u in spec) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11077 |
apply (auto intro: *) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11078 |
done |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11079 |
} with rel_interior_subset show "?lhs \<le> ?rhs" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11080 |
by blast |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11081 |
qed |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11082 |
qed |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11083 |
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11084 |
lemma interior_convex_hull_explicit_minimal: |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11085 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11086 |
shows |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11087 |
"~ affine_dependent s |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11088 |
==> interior(convex hull s) = |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11089 |
(if card(s) \<le> DIM('a) then {}
|
| 64267 | 11090 |
else {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> sum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = y})"
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11091 |
apply (simp add: aff_independent_finite empty_interior_convex_hull, clarify) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11092 |
apply (rule trans [of _ "rel_interior(convex hull s)"]) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11093 |
apply (simp add: affine_hull_convex_hull affine_independent_span_gt rel_interior_interior) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11094 |
by (simp add: rel_interior_convex_hull_explicit) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11095 |
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11096 |
lemma interior_convex_hull_explicit: |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11097 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11098 |
assumes "~ affine_dependent s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11099 |
shows |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11100 |
"interior(convex hull s) = |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11101 |
(if card(s) \<le> DIM('a) then {}
|
| 64267 | 11102 |
else {y. \<exists>u. (\<forall>x \<in> s. 0 < u x \<and> u x < 1) \<and> sum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = y})"
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11103 |
proof - |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11104 |
{ fix u :: "'a \<Rightarrow> real" and a
|
| 64267 | 11105 |
assume "card Basis < card s" and u: "\<And>x. x\<in>s \<Longrightarrow> 0 < u x" "sum u s = 1" and a: "a \<in> s" |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11106 |
then have cs: "Suc 0 < card s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11107 |
by (metis DIM_positive less_trans_Suc) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11108 |
obtain b where b: "b \<in> s" "a \<noteq> b" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11109 |
proof (cases "s \<le> {a}")
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11110 |
case True |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11111 |
then show thesis |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11112 |
using cs subset_singletonD by fastforce |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11113 |
next |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11114 |
case False |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11115 |
then show thesis |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11116 |
by (blast intro: that) |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11117 |
qed |
| 64267 | 11118 |
have "u a + u b \<le> sum u {a,b}"
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11119 |
using a b by simp |
| 64267 | 11120 |
also have "... \<le> sum u s" |
11121 |
apply (rule Groups_Big.sum_mono2) |
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11122 |
using a b u |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11123 |
apply (auto simp: less_imp_le aff_independent_finite assms) |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11124 |
done |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11125 |
finally have "u a < 1" |
| 60420 | 11126 |
using \<open>b \<in> s\<close> u by fastforce |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11127 |
} note [simp] = this |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11128 |
show ?thesis |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11129 |
using assms |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11130 |
apply (auto simp: interior_convex_hull_explicit_minimal) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11131 |
apply (rule_tac x=u in exI) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11132 |
apply (auto simp: not_le) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11133 |
done |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11134 |
qed |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11135 |
|
|
62620
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11136 |
lemma interior_closed_segment_ge2: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11137 |
fixes a :: "'a::euclidean_space" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11138 |
assumes "2 \<le> DIM('a)"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11139 |
shows "interior(closed_segment a b) = {}"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11140 |
using assms unfolding segment_convex_hull |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11141 |
proof - |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11142 |
have "card {a, b} \<le> DIM('a)"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11143 |
using assms |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11144 |
by (simp add: card_insert_if linear not_less_eq_eq numeral_2_eq_2) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11145 |
then show "interior (convex hull {a, b}) = {}"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11146 |
by (metis empty_interior_convex_hull finite.insertI finite.emptyI) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11147 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11148 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11149 |
lemma interior_open_segment: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11150 |
fixes a :: "'a::euclidean_space" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11151 |
shows "interior(open_segment a b) = |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11152 |
(if 2 \<le> DIM('a) then {} else open_segment a b)"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11153 |
proof (simp add: not_le, intro conjI impI) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11154 |
assume "2 \<le> DIM('a)"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11155 |
then show "interior (open_segment a b) = {}"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11156 |
apply (simp add: segment_convex_hull open_segment_def) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11157 |
apply (metis Diff_subset interior_mono segment_convex_hull subset_empty interior_closed_segment_ge2) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11158 |
done |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11159 |
next |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11160 |
assume le2: "DIM('a) < 2"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11161 |
show "interior (open_segment a b) = open_segment a b" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11162 |
proof (cases "a = b") |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11163 |
case True then show ?thesis by auto |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11164 |
next |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11165 |
case False |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11166 |
with le2 have "affine hull (open_segment a b) = UNIV" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11167 |
apply simp |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11168 |
apply (rule affine_independent_span_gt) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11169 |
apply (simp_all add: affine_dependent_def insert_Diff_if) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11170 |
done |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11171 |
then show "interior (open_segment a b) = open_segment a b" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11172 |
using rel_interior_interior rel_interior_open_segment by blast |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11173 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11174 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11175 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11176 |
lemma interior_closed_segment: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11177 |
fixes a :: "'a::euclidean_space" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11178 |
shows "interior(closed_segment a b) = |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11179 |
(if 2 \<le> DIM('a) then {} else open_segment a b)"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11180 |
proof (cases "a = b") |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11181 |
case True then show ?thesis by simp |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11182 |
next |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11183 |
case False |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11184 |
then have "closure (open_segment a b) = closed_segment a b" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11185 |
by simp |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11186 |
then show ?thesis |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11187 |
by (metis (no_types) convex_interior_closure convex_open_segment interior_open_segment) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11188 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11189 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11190 |
lemmas interior_segment = interior_closed_segment interior_open_segment |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11191 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11192 |
lemma closed_segment_eq [simp]: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11193 |
fixes a :: "'a::euclidean_space" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11194 |
shows "closed_segment a b = closed_segment c d \<longleftrightarrow> {a,b} = {c,d}"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11195 |
proof |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11196 |
assume abcd: "closed_segment a b = closed_segment c d" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11197 |
show "{a,b} = {c,d}"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11198 |
proof (cases "a=b \<or> c=d") |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11199 |
case True with abcd show ?thesis by force |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11200 |
next |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11201 |
case False |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11202 |
then have neq: "a \<noteq> b \<and> c \<noteq> d" by force |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11203 |
have *: "closed_segment c d - {a, b} = rel_interior (closed_segment c d)"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11204 |
using neq abcd by (metis (no_types) open_segment_def rel_interior_closed_segment) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11205 |
have "b \<in> {c, d}"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11206 |
proof - |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11207 |
have "insert b (closed_segment c d) = closed_segment c d" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11208 |
using abcd by blast |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11209 |
then show ?thesis |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11210 |
by (metis DiffD2 Diff_insert2 False * insertI1 insert_Diff_if open_segment_def rel_interior_closed_segment) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11211 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11212 |
moreover have "a \<in> {c, d}"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11213 |
by (metis Diff_iff False * abcd ends_in_segment(1) insertI1 open_segment_def rel_interior_closed_segment) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11214 |
ultimately show "{a, b} = {c, d}"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11215 |
using neq by fastforce |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11216 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11217 |
next |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11218 |
assume "{a,b} = {c,d}"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11219 |
then show "closed_segment a b = closed_segment c d" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11220 |
by (simp add: segment_convex_hull) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11221 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11222 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11223 |
lemma closed_open_segment_eq [simp]: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11224 |
fixes a :: "'a::euclidean_space" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11225 |
shows "closed_segment a b \<noteq> open_segment c d" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11226 |
by (metis DiffE closed_segment_neq_empty closure_closed_segment closure_open_segment ends_in_segment(1) insertI1 open_segment_def) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11227 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11228 |
lemma open_closed_segment_eq [simp]: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11229 |
fixes a :: "'a::euclidean_space" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11230 |
shows "open_segment a b \<noteq> closed_segment c d" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11231 |
using closed_open_segment_eq by blast |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11232 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11233 |
lemma open_segment_eq [simp]: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11234 |
fixes a :: "'a::euclidean_space" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11235 |
shows "open_segment a b = open_segment c d \<longleftrightarrow> a = b \<and> c = d \<or> {a,b} = {c,d}"
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11236 |
(is "?lhs = ?rhs") |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11237 |
proof |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11238 |
assume abcd: ?lhs |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11239 |
show ?rhs |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11240 |
proof (cases "a=b \<or> c=d") |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11241 |
case True with abcd show ?thesis |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11242 |
using finite_open_segment by fastforce |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11243 |
next |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11244 |
case False |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11245 |
then have a2: "a \<noteq> b \<and> c \<noteq> d" by force |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11246 |
with abcd show ?rhs |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11247 |
unfolding open_segment_def |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11248 |
by (metis (no_types) abcd closed_segment_eq closure_open_segment) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11249 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11250 |
next |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11251 |
assume ?rhs |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11252 |
then show ?lhs |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11253 |
by (metis Diff_cancel convex_hull_singleton insert_absorb2 open_segment_def segment_convex_hull) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11254 |
qed |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11255 |
|
| 60420 | 11256 |
subsection\<open>Similar results for closure and (relative or absolute) frontier.\<close> |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11257 |
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11258 |
lemma closure_convex_hull [simp]: |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11259 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11260 |
shows "compact s ==> closure(convex hull s) = convex hull s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11261 |
by (simp add: compact_imp_closed compact_convex_hull) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11262 |
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11263 |
lemma rel_frontier_convex_hull_explicit: |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11264 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11265 |
assumes "~ affine_dependent s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11266 |
shows "rel_frontier(convex hull s) = |
| 64267 | 11267 |
{y. \<exists>u. (\<forall>x \<in> s. 0 \<le> u x) \<and> (\<exists>x \<in> s. u x = 0) \<and> sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}"
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11268 |
proof - |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11269 |
have fs: "finite s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11270 |
using assms by (simp add: aff_independent_finite) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11271 |
show ?thesis |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11272 |
apply (simp add: rel_frontier_def finite_imp_compact rel_interior_convex_hull_explicit assms fs) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11273 |
apply (auto simp: convex_hull_finite fs) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11274 |
apply (drule_tac x=u in spec) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11275 |
apply (rule_tac x=u in exI) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11276 |
apply force |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11277 |
apply (rename_tac v) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11278 |
apply (rule notE [OF assms]) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11279 |
apply (simp add: affine_dependent_explicit) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11280 |
apply (rule_tac x=s in exI) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11281 |
apply (auto simp: fs) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11282 |
apply (rule_tac x = "\<lambda>x. u x - v x" in exI) |
| 64267 | 11283 |
apply (force simp: sum_subtractf scaleR_diff_left) |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11284 |
done |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11285 |
qed |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11286 |
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11287 |
lemma frontier_convex_hull_explicit: |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11288 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11289 |
assumes "~ affine_dependent s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11290 |
shows "frontier(convex hull s) = |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11291 |
{y. \<exists>u. (\<forall>x \<in> s. 0 \<le> u x) \<and> (DIM ('a) < card s \<longrightarrow> (\<exists>x \<in> s. u x = 0)) \<and>
|
| 64267 | 11292 |
sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}" |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11293 |
proof - |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11294 |
have fs: "finite s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11295 |
using assms by (simp add: aff_independent_finite) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11296 |
show ?thesis |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11297 |
proof (cases "DIM ('a) < card s")
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11298 |
case True |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11299 |
with assms fs show ?thesis |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11300 |
by (simp add: rel_frontier_def frontier_def rel_frontier_convex_hull_explicit [symmetric] |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11301 |
interior_convex_hull_explicit_minimal rel_interior_convex_hull_explicit) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11302 |
next |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11303 |
case False |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11304 |
then have "card s \<le> DIM ('a)"
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11305 |
by linarith |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11306 |
then show ?thesis |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11307 |
using assms fs |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11308 |
apply (simp add: frontier_def interior_convex_hull_explicit finite_imp_compact) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11309 |
apply (simp add: convex_hull_finite) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11310 |
done |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11311 |
qed |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11312 |
qed |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11313 |
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11314 |
lemma rel_frontier_convex_hull_cases: |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11315 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11316 |
assumes "~ affine_dependent s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11317 |
shows "rel_frontier(convex hull s) = \<Union>{convex hull (s - {x}) |x. x \<in> s}"
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11318 |
proof - |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11319 |
have fs: "finite s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11320 |
using assms by (simp add: aff_independent_finite) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11321 |
{ fix u a
|
| 64267 | 11322 |
have "\<forall>x\<in>s. 0 \<le> u x \<Longrightarrow> a \<in> s \<Longrightarrow> u a = 0 \<Longrightarrow> sum u s = 1 \<Longrightarrow> |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11323 |
\<exists>x v. x \<in> s \<and> |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11324 |
(\<forall>x\<in>s - {x}. 0 \<le> v x) \<and>
|
| 64267 | 11325 |
sum v (s - {x}) = 1 \<and> (\<Sum>x\<in>s - {x}. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)"
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11326 |
apply (rule_tac x=a in exI) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11327 |
apply (rule_tac x=u in exI) |
| 64267 | 11328 |
apply (simp add: Groups_Big.sum_diff1 fs) |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11329 |
done } |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11330 |
moreover |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11331 |
{ fix a u
|
| 64267 | 11332 |
have "a \<in> s \<Longrightarrow> \<forall>x\<in>s - {a}. 0 \<le> u x \<Longrightarrow> sum u (s - {a}) = 1 \<Longrightarrow>
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11333 |
\<exists>v. (\<forall>x\<in>s. 0 \<le> v x) \<and> |
| 64267 | 11334 |
(\<exists>x\<in>s. v x = 0) \<and> sum v s = 1 \<and> (\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>x\<in>s - {a}. u x *\<^sub>R x)"
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11335 |
apply (rule_tac x="\<lambda>x. if x = a then 0 else u x" in exI) |
| 64267 | 11336 |
apply (auto simp: sum.If_cases Diff_eq if_smult fs) |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11337 |
done } |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11338 |
ultimately show ?thesis |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11339 |
using assms |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11340 |
apply (simp add: rel_frontier_convex_hull_explicit) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11341 |
apply (simp add: convex_hull_finite fs Union_SetCompr_eq, auto) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11342 |
done |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11343 |
qed |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11344 |
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11345 |
lemma frontier_convex_hull_eq_rel_frontier: |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11346 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11347 |
assumes "~ affine_dependent s" |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11348 |
shows "frontier(convex hull s) = |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11349 |
(if card s \<le> DIM ('a) then convex hull s else rel_frontier(convex hull s))"
|
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11350 |
using assms |
|
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11351 |
unfolding rel_frontier_def frontier_def |
|
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11352 |
by (simp add: affine_independent_span_gt rel_interior_interior |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11353 |
finite_imp_compact empty_interior_convex_hull aff_independent_finite) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11354 |
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11355 |
lemma frontier_convex_hull_cases: |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11356 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11357 |
assumes "~ affine_dependent s" |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11358 |
shows "frontier(convex hull s) = |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11359 |
(if card s \<le> DIM ('a) then convex hull s else \<Union>{convex hull (s - {x}) |x. x \<in> s})"
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11360 |
by (simp add: assms frontier_convex_hull_eq_rel_frontier rel_frontier_convex_hull_cases) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11361 |
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11362 |
lemma in_frontier_convex_hull: |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11363 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11364 |
assumes "finite s" "card s \<le> Suc (DIM ('a))" "x \<in> s"
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11365 |
shows "x \<in> frontier(convex hull s)" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11366 |
proof (cases "affine_dependent s") |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11367 |
case True |
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11368 |
with assms show ?thesis |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11369 |
apply (auto simp: affine_dependent_def frontier_def finite_imp_compact hull_inc) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11370 |
by (metis card.insert_remove convex_hull_subset_affine_hull empty_interior_affine_hull finite_Diff hull_redundant insert_Diff insert_Diff_single insert_not_empty interior_mono not_less_eq_eq subset_empty) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11371 |
next |
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11372 |
case False |
|
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11373 |
{ assume "card s = Suc (card Basis)"
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11374 |
then have cs: "Suc 0 < card s" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11375 |
by (simp add: DIM_positive) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11376 |
with subset_singletonD have "\<exists>y \<in> s. y \<noteq> x" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11377 |
by (cases "s \<le> {x}") fastforce+
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11378 |
} note [dest!] = this |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11379 |
show ?thesis using assms |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11380 |
unfolding frontier_convex_hull_cases [OF False] Union_SetCompr_eq |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11381 |
by (auto simp: le_Suc_eq hull_inc) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11382 |
qed |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11383 |
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11384 |
lemma not_in_interior_convex_hull: |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11385 |
fixes s :: "'a::euclidean_space set" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11386 |
assumes "finite s" "card s \<le> Suc (DIM ('a))" "x \<in> s"
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11387 |
shows "x \<notin> interior(convex hull s)" |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11388 |
using in_frontier_convex_hull [OF assms] |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11389 |
by (metis Diff_iff frontier_def) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11390 |
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11391 |
lemma interior_convex_hull_eq_empty: |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11392 |
fixes s :: "'a::euclidean_space set" |
| 60762 | 11393 |
assumes "card s = Suc (DIM ('a))"
|
|
60307
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11394 |
shows "interior(convex hull s) = {} \<longleftrightarrow> affine_dependent s"
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11395 |
proof - |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11396 |
{ fix a b
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11397 |
assume ab: "a \<in> interior (convex hull s)" "b \<in> s" "b \<in> affine hull (s - {b})"
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11398 |
then have "interior(affine hull s) = {}" using assms
|
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11399 |
by (metis DIM_positive One_nat_def Suc_mono card.remove card_infinite empty_interior_affine_hull eq_iff hull_redundant insert_Diff not_less zero_le_one) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11400 |
then have False using ab |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11401 |
by (metis convex_hull_subset_affine_hull equals0D interior_mono subset_eq) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11402 |
} then |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11403 |
show ?thesis |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11404 |
using assms |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11405 |
apply auto |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11406 |
apply (metis UNIV_I affine_hull_convex_hull affine_hull_empty affine_independent_span_eq convex_convex_hull empty_iff rel_interior_interior rel_interior_same_affine_hull) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11407 |
apply (auto simp: affine_dependent_def) |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11408 |
done |
|
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents:
60303
diff
changeset
|
11409 |
qed |
| 50104 | 11410 |
|
|
60800
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11411 |
|
| 60762 | 11412 |
subsection \<open>Coplanarity, and collinearity in terms of affine hull\<close> |
11413 |
||
11414 |
definition coplanar where |
|
11415 |
"coplanar s \<equiv> \<exists>u v w. s \<subseteq> affine hull {u,v,w}"
|
|
11416 |
||
11417 |
lemma collinear_affine_hull: |
|
11418 |
"collinear s \<longleftrightarrow> (\<exists>u v. s \<subseteq> affine hull {u,v})"
|
|
11419 |
proof (cases "s={}")
|
|
11420 |
case True then show ?thesis |
|
11421 |
by simp |
|
11422 |
next |
|
11423 |
case False |
|
11424 |
then obtain x where x: "x \<in> s" by auto |
|
11425 |
{ fix u
|
|
11426 |
assume *: "\<And>x y. \<lbrakk>x\<in>s; y\<in>s\<rbrakk> \<Longrightarrow> \<exists>c. x - y = c *\<^sub>R u" |
|
11427 |
have "\<exists>u v. s \<subseteq> {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}"
|
|
11428 |
apply (rule_tac x=x in exI) |
|
11429 |
apply (rule_tac x="x+u" in exI, clarify) |
|
11430 |
apply (erule exE [OF * [OF x]]) |
|
11431 |
apply (rename_tac c) |
|
11432 |
apply (rule_tac x="1+c" in exI) |
|
11433 |
apply (rule_tac x="-c" in exI) |
|
11434 |
apply (simp add: algebra_simps) |
|
11435 |
done |
|
11436 |
} moreover |
|
11437 |
{ fix u v x y
|
|
11438 |
assume *: "s \<subseteq> {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}"
|
|
11439 |
have "x\<in>s \<Longrightarrow> y\<in>s \<Longrightarrow> \<exists>c. x - y = c *\<^sub>R (v-u)" |
|
11440 |
apply (drule subsetD [OF *])+ |
|
11441 |
apply simp |
|
11442 |
apply clarify |
|
11443 |
apply (rename_tac r1 r2) |
|
11444 |
apply (rule_tac x="r1-r2" in exI) |
|
11445 |
apply (simp add: algebra_simps) |
|
11446 |
apply (metis scaleR_left.add) |
|
11447 |
done |
|
11448 |
} ultimately |
|
11449 |
show ?thesis |
|
11450 |
unfolding collinear_def affine_hull_2 |
|
11451 |
by blast |
|
11452 |
qed |
|
11453 |
||
|
62620
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11454 |
lemma collinear_closed_segment [simp]: "collinear (closed_segment a b)" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11455 |
by (metis affine_hull_convex_hull collinear_affine_hull hull_subset segment_convex_hull) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11456 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11457 |
lemma collinear_open_segment [simp]: "collinear (open_segment a b)" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11458 |
unfolding open_segment_def |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11459 |
by (metis convex_hull_subset_affine_hull segment_convex_hull dual_order.trans |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11460 |
convex_hull_subset_affine_hull Diff_subset collinear_affine_hull) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11461 |
|
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11462 |
lemma subset_continuous_image_segment_1: |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11463 |
fixes f :: "'a::euclidean_space \<Rightarrow> real" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11464 |
assumes "continuous_on (closed_segment a b) f" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11465 |
shows "closed_segment (f a) (f b) \<subseteq> image f (closed_segment a b)" |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11466 |
by (metis connected_segment convex_contains_segment ends_in_segment imageI |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11467 |
is_interval_connected_1 is_interval_convex connected_continuous_image [OF assms]) |
|
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents:
62618
diff
changeset
|
11468 |
|
|
64006
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11469 |
lemma continuous_injective_image_segment_1: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11470 |
fixes f :: "'a::euclidean_space \<Rightarrow> real" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11471 |
assumes contf: "continuous_on (closed_segment a b) f" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11472 |
and injf: "inj_on f (closed_segment a b)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11473 |
shows "f ` (closed_segment a b) = closed_segment (f a) (f b)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11474 |
proof |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11475 |
show "closed_segment (f a) (f b) \<subseteq> f ` closed_segment a b" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11476 |
by (metis subset_continuous_image_segment_1 contf) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11477 |
show "f ` closed_segment a b \<subseteq> closed_segment (f a) (f b)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11478 |
proof (cases "a = b") |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11479 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11480 |
then show ?thesis by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11481 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11482 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11483 |
then have fnot: "f a \<noteq> f b" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11484 |
using inj_onD injf by fastforce |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11485 |
moreover |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11486 |
have "f a \<notin> open_segment (f c) (f b)" if c: "c \<in> closed_segment a b" for c |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11487 |
proof (clarsimp simp add: open_segment_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11488 |
assume fa: "f a \<in> closed_segment (f c) (f b)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11489 |
moreover have "closed_segment (f c) (f b) \<subseteq> f ` closed_segment c b" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11490 |
by (meson closed_segment_subset contf continuous_on_subset convex_closed_segment ends_in_segment(2) subset_continuous_image_segment_1 that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11491 |
ultimately have "f a \<in> f ` closed_segment c b" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11492 |
by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11493 |
then have a: "a \<in> closed_segment c b" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11494 |
by (meson ends_in_segment inj_on_image_mem_iff_alt injf subset_closed_segment that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11495 |
have cb: "closed_segment c b \<subseteq> closed_segment a b" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11496 |
by (simp add: closed_segment_subset that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11497 |
show "f a = f c" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11498 |
proof (rule between_antisym) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11499 |
show "between (f c, f b) (f a)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11500 |
by (simp add: between_mem_segment fa) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11501 |
show "between (f a, f b) (f c)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11502 |
by (metis a cb between_antisym between_mem_segment between_triv1 subset_iff) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11503 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11504 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11505 |
moreover |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11506 |
have "f b \<notin> open_segment (f a) (f c)" if c: "c \<in> closed_segment a b" for c |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11507 |
proof (clarsimp simp add: open_segment_def fnot eq_commute) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11508 |
assume fb: "f b \<in> closed_segment (f a) (f c)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11509 |
moreover have "closed_segment (f a) (f c) \<subseteq> f ` closed_segment a c" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11510 |
by (meson contf continuous_on_subset ends_in_segment(1) subset_closed_segment subset_continuous_image_segment_1 that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11511 |
ultimately have "f b \<in> f ` closed_segment a c" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11512 |
by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11513 |
then have b: "b \<in> closed_segment a c" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11514 |
by (meson ends_in_segment inj_on_image_mem_iff_alt injf subset_closed_segment that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11515 |
have ca: "closed_segment a c \<subseteq> closed_segment a b" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11516 |
by (simp add: closed_segment_subset that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11517 |
show "f b = f c" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11518 |
proof (rule between_antisym) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11519 |
show "between (f c, f a) (f b)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11520 |
by (simp add: between_commute between_mem_segment fb) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11521 |
show "between (f b, f a) (f c)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11522 |
by (metis b between_antisym between_commute between_mem_segment between_triv2 that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11523 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11524 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11525 |
ultimately show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11526 |
by (force simp: closed_segment_eq_real_ivl open_segment_eq_real_ivl split: if_split_asm) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11527 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11528 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11529 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11530 |
lemma continuous_injective_image_open_segment_1: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11531 |
fixes f :: "'a::euclidean_space \<Rightarrow> real" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11532 |
assumes contf: "continuous_on (closed_segment a b) f" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11533 |
and injf: "inj_on f (closed_segment a b)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11534 |
shows "f ` (open_segment a b) = open_segment (f a) (f b)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11535 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11536 |
have "f ` (open_segment a b) = f ` (closed_segment a b) - {f a, f b}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11537 |
by (metis (no_types, hide_lams) empty_subsetI ends_in_segment image_insert image_is_empty inj_on_image_set_diff injf insert_subset open_segment_def segment_open_subset_closed) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11538 |
also have "... = open_segment (f a) (f b)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11539 |
using continuous_injective_image_segment_1 [OF assms] |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11540 |
by (simp add: open_segment_def inj_on_image_set_diff [OF injf]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11541 |
finally show ?thesis . |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11542 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
63971
diff
changeset
|
11543 |
|
| 60762 | 11544 |
lemma collinear_imp_coplanar: |
11545 |
"collinear s ==> coplanar s" |
|
11546 |
by (metis collinear_affine_hull coplanar_def insert_absorb2) |
|
11547 |
||
11548 |
lemma collinear_small: |
|
11549 |
assumes "finite s" "card s \<le> 2" |
|
11550 |
shows "collinear s" |
|
11551 |
proof - |
|
11552 |
have "card s = 0 \<or> card s = 1 \<or> card s = 2" |
|
11553 |
using assms by linarith |
|
11554 |
then show ?thesis using assms |
|
11555 |
using card_eq_SucD |
|
11556 |
by auto (metis collinear_2 numeral_2_eq_2) |
|
11557 |
qed |
|
11558 |
||
11559 |
lemma coplanar_small: |
|
11560 |
assumes "finite s" "card s \<le> 3" |
|
11561 |
shows "coplanar s" |
|
11562 |
proof - |
|
11563 |
have "card s \<le> 2 \<or> card s = Suc (Suc (Suc 0))" |
|
11564 |
using assms by linarith |
|
11565 |
then show ?thesis using assms |
|
11566 |
apply safe |
|
11567 |
apply (simp add: collinear_small collinear_imp_coplanar) |
|
11568 |
apply (safe dest!: card_eq_SucD) |
|
11569 |
apply (auto simp: coplanar_def) |
|
11570 |
apply (metis hull_subset insert_subset) |
|
11571 |
done |
|
11572 |
qed |
|
11573 |
||
11574 |
lemma coplanar_empty: "coplanar {}"
|
|
11575 |
by (simp add: coplanar_small) |
|
11576 |
||
11577 |
lemma coplanar_sing: "coplanar {a}"
|
|
11578 |
by (simp add: coplanar_small) |
|
11579 |
||
11580 |
lemma coplanar_2: "coplanar {a,b}"
|
|
11581 |
by (auto simp: card_insert_if coplanar_small) |
|
11582 |
||
11583 |
lemma coplanar_3: "coplanar {a,b,c}"
|
|
11584 |
by (auto simp: card_insert_if coplanar_small) |
|
11585 |
||
11586 |
lemma collinear_affine_hull_collinear: "collinear(affine hull s) \<longleftrightarrow> collinear s" |
|
11587 |
unfolding collinear_affine_hull |
|
11588 |
by (metis affine_affine_hull subset_hull hull_hull hull_mono) |
|
11589 |
||
11590 |
lemma coplanar_affine_hull_coplanar: "coplanar(affine hull s) \<longleftrightarrow> coplanar s" |
|
11591 |
unfolding coplanar_def |
|
11592 |
by (metis affine_affine_hull subset_hull hull_hull hull_mono) |
|
11593 |
||
11594 |
lemma coplanar_linear_image: |
|
11595 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" |
|
11596 |
assumes "coplanar s" "linear f" shows "coplanar(f ` s)" |
|
11597 |
proof - |
|
11598 |
{ fix u v w
|
|
11599 |
assume "s \<subseteq> affine hull {u, v, w}"
|
|
11600 |
then have "f ` s \<subseteq> f ` (affine hull {u, v, w})"
|
|
11601 |
by (simp add: image_mono) |
|
11602 |
then have "f ` s \<subseteq> affine hull (f ` {u, v, w})"
|
|
11603 |
by (metis assms(2) linear_conv_bounded_linear affine_hull_linear_image) |
|
11604 |
} then |
|
11605 |
show ?thesis |
|
11606 |
by auto (meson assms(1) coplanar_def) |
|
11607 |
qed |
|
11608 |
||
|
60800
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11609 |
lemma coplanar_translation_imp: "coplanar s \<Longrightarrow> coplanar ((\<lambda>x. a + x) ` s)" |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11610 |
unfolding coplanar_def |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11611 |
apply clarify |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11612 |
apply (rule_tac x="u+a" in exI) |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11613 |
apply (rule_tac x="v+a" in exI) |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11614 |
apply (rule_tac x="w+a" in exI) |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11615 |
using affine_hull_translation [of a "{u,v,w}" for u v w]
|
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11616 |
apply (force simp: add.commute) |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11617 |
done |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11618 |
|
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11619 |
lemma coplanar_translation_eq: "coplanar((\<lambda>x. a + x) ` s) \<longleftrightarrow> coplanar s" |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11620 |
by (metis (no_types) coplanar_translation_imp translation_galois) |
| 60762 | 11621 |
|
11622 |
lemma coplanar_linear_image_eq: |
|
11623 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
11624 |
assumes "linear f" "inj f" shows "coplanar(f ` s) = coplanar s" |
|
|
60809
457abb82fb9e
the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
60800
diff
changeset
|
11625 |
proof |
|
60800
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11626 |
assume "coplanar s" |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11627 |
then show "coplanar (f ` s)" |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11628 |
unfolding coplanar_def |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11629 |
using affine_hull_linear_image [of f "{u,v,w}" for u v w] assms
|
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11630 |
by (meson coplanar_def coplanar_linear_image) |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11631 |
next |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11632 |
obtain g where g: "linear g" "g \<circ> f = id" |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11633 |
using linear_injective_left_inverse [OF assms] |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11634 |
by blast |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11635 |
assume "coplanar (f ` s)" |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11636 |
then obtain u v w where "f ` s \<subseteq> affine hull {u, v, w}"
|
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11637 |
by (auto simp: coplanar_def) |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11638 |
then have "g ` f ` s \<subseteq> g ` (affine hull {u, v, w})"
|
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11639 |
by blast |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11640 |
then have "s \<subseteq> g ` (affine hull {u, v, w})"
|
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11641 |
using g by (simp add: Fun.image_comp) |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11642 |
then show "coplanar s" |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11643 |
unfolding coplanar_def |
| 61222 | 11644 |
using affine_hull_linear_image [of g "{u,v,w}" for u v w] \<open>linear g\<close> linear_conv_bounded_linear
|
|
60800
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11645 |
by fastforce |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11646 |
qed |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11647 |
(*The HOL Light proof is simply |
|
7d04351c795a
New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
11648 |
MATCH_ACCEPT_TAC(LINEAR_INVARIANT_RULE COPLANAR_LINEAR_IMAGE));; |
| 60762 | 11649 |
*) |
11650 |
||
11651 |
lemma coplanar_subset: "\<lbrakk>coplanar t; s \<subseteq> t\<rbrakk> \<Longrightarrow> coplanar s" |
|
11652 |
by (meson coplanar_def order_trans) |
|
11653 |
||
11654 |
lemma affine_hull_3_imp_collinear: "c \<in> affine hull {a,b} \<Longrightarrow> collinear {a,b,c}"
|
|
11655 |
by (metis collinear_2 collinear_affine_hull_collinear hull_redundant insert_commute) |
|
11656 |
||
11657 |
lemma collinear_3_imp_in_affine_hull: "\<lbrakk>collinear {a,b,c}; a \<noteq> b\<rbrakk> \<Longrightarrow> c \<in> affine hull {a,b}"
|
|
11658 |
unfolding collinear_def |
|
11659 |
apply clarify |
|
11660 |
apply (frule_tac x=b in bspec, blast, drule_tac x=a in bspec, blast, erule exE) |
|
11661 |
apply (drule_tac x=c in bspec, blast, drule_tac x=a in bspec, blast, erule exE) |
|
11662 |
apply (rename_tac y x) |
|
11663 |
apply (simp add: affine_hull_2) |
|
11664 |
apply (rule_tac x="1 - x/y" in exI) |
|
11665 |
apply (simp add: algebra_simps) |
|
11666 |
done |
|
11667 |
||
11668 |
lemma collinear_3_affine_hull: |
|
11669 |
assumes "a \<noteq> b" |
|
11670 |
shows "collinear {a,b,c} \<longleftrightarrow> c \<in> affine hull {a,b}"
|
|
11671 |
using affine_hull_3_imp_collinear assms collinear_3_imp_in_affine_hull by blast |
|
11672 |
||
11673 |
lemma collinear_3_eq_affine_dependent: |
|
11674 |
"collinear{a,b,c} \<longleftrightarrow> a = b \<or> a = c \<or> b = c \<or> affine_dependent {a,b,c}"
|
|
11675 |
apply (case_tac "a=b", simp) |
|
11676 |
apply (case_tac "a=c") |
|
11677 |
apply (simp add: insert_commute) |
|
11678 |
apply (case_tac "b=c") |
|
11679 |
apply (simp add: insert_commute) |
|
11680 |
apply (auto simp: affine_dependent_def collinear_3_affine_hull insert_Diff_if) |
|
11681 |
apply (metis collinear_3_affine_hull insert_commute)+ |
|
11682 |
done |
|
11683 |
||
11684 |
lemma affine_dependent_imp_collinear_3: |
|
11685 |
"affine_dependent {a,b,c} \<Longrightarrow> collinear{a,b,c}"
|
|
11686 |
by (simp add: collinear_3_eq_affine_dependent) |
|
11687 |
||
11688 |
lemma collinear_3: "NO_MATCH 0 x \<Longrightarrow> collinear {x,y,z} \<longleftrightarrow> collinear {0, x-y, z-y}"
|
|
11689 |
by (auto simp add: collinear_def) |
|
11690 |
||
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11691 |
lemma collinear_3_expand: |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11692 |
"collinear{a,b,c} \<longleftrightarrow> a = c \<or> (\<exists>u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)"
|
|
63957
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
11693 |
proof - |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11694 |
have "collinear{a,b,c} = collinear{a,c,b}"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11695 |
by (simp add: insert_commute) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11696 |
also have "... = collinear {0, a - c, b - c}"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11697 |
by (simp add: collinear_3) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11698 |
also have "... \<longleftrightarrow> (a = c \<or> b = c \<or> (\<exists>ca. b - c = ca *\<^sub>R (a - c)))" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11699 |
by (simp add: collinear_lemma) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11700 |
also have "... \<longleftrightarrow> a = c \<or> (\<exists>u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11701 |
by (cases "a = c \<or> b = c") (auto simp: algebra_simps) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11702 |
finally show ?thesis . |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11703 |
qed |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11704 |
|
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11705 |
lemma collinear_aff_dim: "collinear S \<longleftrightarrow> aff_dim S \<le> 1" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11706 |
proof |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11707 |
assume "collinear S" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11708 |
then obtain u and v :: "'a" where "aff_dim S \<le> aff_dim {u,v}"
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11709 |
by (metis \<open>collinear S\<close> aff_dim_affine_hull aff_dim_subset collinear_affine_hull) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11710 |
then show "aff_dim S \<le> 1" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11711 |
using order_trans by fastforce |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11712 |
next |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11713 |
assume "aff_dim S \<le> 1" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11714 |
then have le1: "aff_dim (affine hull S) \<le> 1" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11715 |
by simp |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11716 |
obtain B where "B \<subseteq> S" and B: "\<not> affine_dependent B" "affine hull S = affine hull B" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11717 |
using affine_basis_exists [of S] by auto |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11718 |
then have "finite B" "card B \<le> 2" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11719 |
using B le1 by (auto simp: affine_independent_iff_card) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11720 |
then have "collinear B" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11721 |
by (rule collinear_small) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11722 |
then show "collinear S" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11723 |
by (metis \<open>affine hull S = affine hull B\<close> collinear_affine_hull_collinear) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11724 |
qed |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11725 |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11726 |
lemma collinear_midpoint: "collinear{a,midpoint a b,b}"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11727 |
apply (auto simp: collinear_3 collinear_lemma) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11728 |
apply (drule_tac x="-1" in spec) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11729 |
apply (simp add: algebra_simps) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11730 |
done |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11731 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11732 |
lemma midpoint_collinear: |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11733 |
fixes a b c :: "'a::real_normed_vector" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11734 |
assumes "a \<noteq> c" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11735 |
shows "b = midpoint a c \<longleftrightarrow> collinear{a,b,c} \<and> dist a b = dist b c"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11736 |
proof - |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11737 |
have *: "a - (u *\<^sub>R a + (1 - u) *\<^sub>R c) = (1 - u) *\<^sub>R (a - c)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11738 |
"u *\<^sub>R a + (1 - u) *\<^sub>R c - c = u *\<^sub>R (a - c)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11739 |
"\<bar>1 - u\<bar> = \<bar>u\<bar> \<longleftrightarrow> u = 1/2" for u::real |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11740 |
by (auto simp: algebra_simps) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11741 |
have "b = midpoint a c \<Longrightarrow> collinear{a,b,c} "
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11742 |
using collinear_midpoint by blast |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11743 |
moreover have "collinear{a,b,c} \<Longrightarrow> b = midpoint a c \<longleftrightarrow> dist a b = dist b c"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11744 |
apply (auto simp: collinear_3_expand assms dist_midpoint) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11745 |
apply (simp add: dist_norm * assms midpoint_def del: divide_const_simps) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11746 |
apply (simp add: algebra_simps) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11747 |
done |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11748 |
ultimately show ?thesis by blast |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11749 |
qed |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11750 |
|
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11751 |
lemma between_imp_collinear: |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11752 |
fixes x :: "'a :: euclidean_space" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11753 |
assumes "between (a,b) x" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11754 |
shows "collinear {a,x,b}"
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11755 |
proof (cases "x = a \<or> x = b \<or> a = b") |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11756 |
case True with assms show ?thesis |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11757 |
by (auto simp: dist_commute) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11758 |
next |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11759 |
case False with assms show ?thesis |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11760 |
apply (auto simp: collinear_3 collinear_lemma between_norm) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11761 |
apply (drule_tac x="-(norm(b - x) / norm(x - a))" in spec) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11762 |
apply (simp add: vector_add_divide_simps eq_vector_fraction_iff real_vector.scale_minus_right [symmetric]) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11763 |
done |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11764 |
qed |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11765 |
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11766 |
lemma midpoint_between: |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11767 |
fixes a b :: "'a::euclidean_space" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11768 |
shows "b = midpoint a c \<longleftrightarrow> between (a,c) b \<and> dist a b = dist b c" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11769 |
proof (cases "a = c") |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11770 |
case True then show ?thesis |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11771 |
by (auto simp: dist_commute) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11772 |
next |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11773 |
case False |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11774 |
show ?thesis |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11775 |
apply (rule iffI) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11776 |
apply (simp add: between_midpoint(1) dist_midpoint) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11777 |
using False between_imp_collinear midpoint_collinear by blast |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11778 |
qed |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
11779 |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11780 |
lemma collinear_triples: |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11781 |
assumes "a \<noteq> b" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11782 |
shows "collinear(insert a (insert b S)) \<longleftrightarrow> (\<forall>x \<in> S. collinear{a,b,x})"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11783 |
(is "?lhs = ?rhs") |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11784 |
proof safe |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11785 |
fix x |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11786 |
assume ?lhs and "x \<in> S" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11787 |
then show "collinear {a, b, x}"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11788 |
using collinear_subset by force |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11789 |
next |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11790 |
assume ?rhs |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11791 |
then have "\<forall>x \<in> S. collinear{a,x,b}"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11792 |
by (simp add: insert_commute) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11793 |
then have *: "\<exists>u. x = u *\<^sub>R a + (1 - u) *\<^sub>R b" if "x \<in> (insert a (insert b S))" for x |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11794 |
using that assms collinear_3_expand by fastforce+ |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11795 |
show ?lhs |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11796 |
unfolding collinear_def |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11797 |
apply (rule_tac x="b-a" in exI) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11798 |
apply (clarify dest!: *) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11799 |
by (metis (no_types, hide_lams) add.commute diff_add_cancel diff_diff_eq2 real_vector.scale_right_diff_distrib scaleR_left.diff) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11800 |
qed |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11801 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11802 |
lemma collinear_4_3: |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11803 |
assumes "a \<noteq> b" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11804 |
shows "collinear {a,b,c,d} \<longleftrightarrow> collinear{a,b,c} \<and> collinear{a,b,d}"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11805 |
using collinear_triples [OF assms, of "{c,d}"] by (force simp:)
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11806 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11807 |
lemma collinear_3_trans: |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11808 |
assumes "collinear{a,b,c}" "collinear{b,c,d}" "b \<noteq> c"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11809 |
shows "collinear{a,b,d}"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11810 |
proof - |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11811 |
have "collinear{b,c,a,d}"
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11812 |
by (metis (full_types) assms collinear_4_3 insert_commute) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11813 |
then show ?thesis |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11814 |
by (simp add: collinear_subset) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11815 |
qed |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11816 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
11817 |
lemma affine_hull_eq_empty [simp]: "affine hull S = {} \<longleftrightarrow> S = {}"
|
| 60762 | 11818 |
using affine_hull_nonempty by blast |
11819 |
||
11820 |
lemma affine_hull_2_alt: |
|
11821 |
fixes a b :: "'a::real_vector" |
|
11822 |
shows "affine hull {a,b} = range (\<lambda>u. a + u *\<^sub>R (b - a))"
|
|
11823 |
apply (simp add: affine_hull_2, safe) |
|
11824 |
apply (rule_tac x=v in image_eqI) |
|
11825 |
apply (simp add: algebra_simps) |
|
11826 |
apply (metis scaleR_add_left scaleR_one, simp) |
|
11827 |
apply (rule_tac x="1-u" in exI) |
|
11828 |
apply (simp add: algebra_simps) |
|
11829 |
done |
|
11830 |
||
11831 |
lemma interior_convex_hull_3_minimal: |
|
11832 |
fixes a :: "'a::euclidean_space" |
|
11833 |
shows "\<lbrakk>~ collinear{a,b,c}; DIM('a) = 2\<rbrakk>
|
|
11834 |
\<Longrightarrow> interior(convex hull {a,b,c}) =
|
|
11835 |
{v. \<exists>x y z. 0 < x \<and> 0 < y \<and> 0 < z \<and> x + y + z = 1 \<and>
|
|
11836 |
x *\<^sub>R a + y *\<^sub>R b + z *\<^sub>R c = v}" |
|
11837 |
apply (simp add: collinear_3_eq_affine_dependent interior_convex_hull_explicit_minimal, safe) |
|
11838 |
apply (rule_tac x="u a" in exI, simp) |
|
11839 |
apply (rule_tac x="u b" in exI, simp) |
|
11840 |
apply (rule_tac x="u c" in exI, simp) |
|
11841 |
apply (rename_tac uu x y z) |
|
11842 |
apply (rule_tac x="\<lambda>r. (if r=a then x else if r=b then y else if r=c then z else 0)" in exI) |
|
11843 |
apply simp |
|
11844 |
done |
|
11845 |
||
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11846 |
subsection\<open>The infimum of the distance between two sets\<close> |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11847 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11848 |
definition setdist :: "'a::metric_space set \<Rightarrow> 'a set \<Rightarrow> real" where |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11849 |
"setdist s t \<equiv> |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11850 |
(if s = {} \<or> t = {} then 0
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11851 |
else Inf {dist x y| x y. x \<in> s \<and> y \<in> t})"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11852 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11853 |
lemma setdist_empty1 [simp]: "setdist {} t = 0"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11854 |
by (simp add: setdist_def) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11855 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11856 |
lemma setdist_empty2 [simp]: "setdist t {} = 0"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11857 |
by (simp add: setdist_def) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11858 |
|
| 63301 | 11859 |
lemma setdist_pos_le [simp]: "0 \<le> setdist s t" |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11860 |
by (auto simp: setdist_def ex_in_conv [symmetric] intro: cInf_greatest) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11861 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11862 |
lemma le_setdistI: |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11863 |
assumes "s \<noteq> {}" "t \<noteq> {}" "\<And>x y. \<lbrakk>x \<in> s; y \<in> t\<rbrakk> \<Longrightarrow> d \<le> dist x y"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11864 |
shows "d \<le> setdist s t" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11865 |
using assms |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11866 |
by (auto simp: setdist_def Set.ex_in_conv [symmetric] intro: cInf_greatest) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11867 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11868 |
lemma setdist_le_dist: "\<lbrakk>x \<in> s; y \<in> t\<rbrakk> \<Longrightarrow> setdist s t \<le> dist x y" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11869 |
unfolding setdist_def |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11870 |
by (auto intro!: bdd_belowI [where m=0] cInf_lower) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11871 |
|
|
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
|
11872 |
lemma le_setdist_iff: |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11873 |
"d \<le> setdist s t \<longleftrightarrow> |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11874 |
(\<forall>x \<in> s. \<forall>y \<in> t. d \<le> dist x y) \<and> (s = {} \<or> t = {} \<longrightarrow> d \<le> 0)"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11875 |
apply (cases "s = {} \<or> t = {}")
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11876 |
apply (force simp add: setdist_def) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11877 |
apply (intro iffI conjI) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11878 |
using setdist_le_dist apply fastforce |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11879 |
apply (auto simp: intro: le_setdistI) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11880 |
done |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11881 |
|
|
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
|
11882 |
lemma setdist_ltE: |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11883 |
assumes "setdist s t < b" "s \<noteq> {}" "t \<noteq> {}"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11884 |
obtains x y where "x \<in> s" "y \<in> t" "dist x y < b" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11885 |
using assms |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11886 |
by (auto simp: not_le [symmetric] le_setdist_iff) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11887 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11888 |
lemma setdist_refl: "setdist s s = 0" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11889 |
apply (cases "s = {}")
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11890 |
apply (force simp add: setdist_def) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11891 |
apply (rule antisym [OF _ setdist_pos_le]) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11892 |
apply (metis all_not_in_conv dist_self setdist_le_dist) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11893 |
done |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11894 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11895 |
lemma setdist_sym: "setdist s t = setdist t s" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11896 |
by (force simp: setdist_def dist_commute intro!: arg_cong [where f=Inf]) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11897 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11898 |
lemma setdist_triangle: "setdist s t \<le> setdist s {a} + setdist {a} t"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11899 |
proof (cases "s = {} \<or> t = {}")
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11900 |
case True then show ?thesis |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11901 |
using setdist_pos_le by fastforce |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11902 |
next |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11903 |
case False |
|
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
|
11904 |
have "\<And>x. x \<in> s \<Longrightarrow> setdist s t - dist x a \<le> setdist {a} t"
|
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11905 |
apply (rule le_setdistI, blast) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11906 |
using False apply (fastforce intro: le_setdistI) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11907 |
apply (simp add: algebra_simps) |
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62398
diff
changeset
|
11908 |
apply (metis dist_commute dist_triangle3 order_trans [OF setdist_le_dist]) |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11909 |
done |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11910 |
then have "setdist s t - setdist {a} t \<le> setdist s {a}"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11911 |
using False by (fastforce intro: le_setdistI) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11912 |
then show ?thesis |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11913 |
by (simp add: algebra_simps) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11914 |
qed |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11915 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11916 |
lemma setdist_singletons [simp]: "setdist {x} {y} = dist x y"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11917 |
by (simp add: setdist_def) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11918 |
|
| 61945 | 11919 |
lemma setdist_Lipschitz: "\<bar>setdist {x} s - setdist {y} s\<bar> \<le> dist x y"
|
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11920 |
apply (subst setdist_singletons [symmetric]) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11921 |
by (metis abs_diff_le_iff diff_le_eq setdist_triangle setdist_sym) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11922 |
|
| 63301 | 11923 |
lemma continuous_at_setdist [continuous_intros]: "continuous (at x) (\<lambda>y. (setdist {y} s))"
|
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11924 |
by (force simp: continuous_at_eps_delta dist_real_def intro: le_less_trans [OF setdist_Lipschitz]) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11925 |
|
| 63301 | 11926 |
lemma continuous_on_setdist [continuous_intros]: "continuous_on t (\<lambda>y. (setdist {y} s))"
|
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11927 |
by (metis continuous_at_setdist continuous_at_imp_continuous_on) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11928 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11929 |
lemma uniformly_continuous_on_setdist: "uniformly_continuous_on t (\<lambda>y. (setdist {y} s))"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11930 |
by (force simp: uniformly_continuous_on_def dist_real_def intro: le_less_trans [OF setdist_Lipschitz]) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11931 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11932 |
lemma setdist_subset_right: "\<lbrakk>t \<noteq> {}; t \<subseteq> u\<rbrakk> \<Longrightarrow> setdist s u \<le> setdist s t"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11933 |
apply (cases "s = {} \<or> u = {}", force)
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11934 |
apply (auto simp: setdist_def intro!: bdd_belowI [where m=0] cInf_superset_mono) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11935 |
done |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11936 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11937 |
lemma setdist_subset_left: "\<lbrakk>s \<noteq> {}; s \<subseteq> t\<rbrakk> \<Longrightarrow> setdist t u \<le> setdist s u"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11938 |
by (metis setdist_subset_right setdist_sym) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11939 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11940 |
lemma setdist_closure_1 [simp]: "setdist (closure s) t = setdist s t" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11941 |
proof (cases "s = {} \<or> t = {}")
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11942 |
case True then show ?thesis by force |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11943 |
next |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11944 |
case False |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11945 |
{ fix y
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11946 |
assume "y \<in> t" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11947 |
have "continuous_on (closure s) (\<lambda>a. dist a y)" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11948 |
by (auto simp: continuous_intros dist_norm) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11949 |
then have *: "\<And>x. x \<in> closure s \<Longrightarrow> setdist s t \<le> dist x y" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11950 |
apply (rule continuous_ge_on_closure) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11951 |
apply assumption |
| 61222 | 11952 |
apply (blast intro: setdist_le_dist \<open>y \<in> t\<close> ) |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11953 |
done |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11954 |
} note * = this |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11955 |
show ?thesis |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11956 |
apply (rule antisym) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11957 |
using False closure_subset apply (blast intro: setdist_subset_left) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11958 |
using False * |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11959 |
apply (force simp add: closure_eq_empty intro!: le_setdistI) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11960 |
done |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11961 |
qed |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11962 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11963 |
lemma setdist_closure_2 [simp]: "setdist t (closure s) = setdist t s" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11964 |
by (metis setdist_closure_1 setdist_sym) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11965 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11966 |
lemma setdist_compact_closed: |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11967 |
fixes s :: "'a::euclidean_space set" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11968 |
assumes s: "compact s" and t: "closed t" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11969 |
and "s \<noteq> {}" "t \<noteq> {}"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11970 |
shows "\<exists>x \<in> s. \<exists>y \<in> t. dist x y = setdist s t" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11971 |
proof - |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11972 |
have "{x - y |x y. x \<in> s \<and> y \<in> t} \<noteq> {}"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11973 |
using assms by blast |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11974 |
then |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11975 |
have "\<exists>x \<in> s. \<exists>y \<in> t. dist x y \<le> setdist s t" |
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
11976 |
apply (rule distance_attains_inf [where a=0, OF compact_closed_differences [OF s t]]) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
11977 |
apply (simp add: dist_norm le_setdist_iff) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62131
diff
changeset
|
11978 |
apply blast |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11979 |
done |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11980 |
then show ?thesis |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11981 |
by (blast intro!: antisym [OF _ setdist_le_dist] ) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11982 |
qed |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11983 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11984 |
lemma setdist_closed_compact: |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11985 |
fixes s :: "'a::euclidean_space set" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11986 |
assumes s: "closed s" and t: "compact t" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11987 |
and "s \<noteq> {}" "t \<noteq> {}"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11988 |
shows "\<exists>x \<in> s. \<exists>y \<in> t. dist x y = setdist s t" |
| 61222 | 11989 |
using setdist_compact_closed [OF t s \<open>t \<noteq> {}\<close> \<open>s \<noteq> {}\<close>]
|
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11990 |
by (metis dist_commute setdist_sym) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11991 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11992 |
lemma setdist_eq_0I: "\<lbrakk>x \<in> s; x \<in> t\<rbrakk> \<Longrightarrow> setdist s t = 0" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11993 |
by (metis antisym dist_self setdist_le_dist setdist_pos_le) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11994 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11995 |
lemma setdist_eq_0_compact_closed: |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11996 |
fixes s :: "'a::euclidean_space set" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11997 |
assumes s: "compact s" and t: "closed t" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11998 |
shows "setdist s t = 0 \<longleftrightarrow> s = {} \<or> t = {} \<or> s \<inter> t \<noteq> {}"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
11999 |
apply (cases "s = {} \<or> t = {}", force)
|
|
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
|
12000 |
using setdist_compact_closed [OF s t] |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12001 |
apply (force intro: setdist_eq_0I ) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12002 |
done |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12003 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12004 |
corollary setdist_gt_0_compact_closed: |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12005 |
fixes s :: "'a::euclidean_space set" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12006 |
assumes s: "compact s" and t: "closed t" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12007 |
shows "setdist s t > 0 \<longleftrightarrow> (s \<noteq> {} \<and> t \<noteq> {} \<and> s \<inter> t = {})"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12008 |
using setdist_pos_le [of s t] setdist_eq_0_compact_closed [OF assms] |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12009 |
by linarith |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12010 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12011 |
lemma setdist_eq_0_closed_compact: |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12012 |
fixes s :: "'a::euclidean_space set" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12013 |
assumes s: "closed s" and t: "compact t" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12014 |
shows "setdist s t = 0 \<longleftrightarrow> s = {} \<or> t = {} \<or> s \<inter> t \<noteq> {}"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12015 |
using setdist_eq_0_compact_closed [OF t s] |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12016 |
by (metis Int_commute setdist_sym) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12017 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12018 |
lemma setdist_eq_0_bounded: |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12019 |
fixes s :: "'a::euclidean_space set" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12020 |
assumes "bounded s \<or> bounded t" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12021 |
shows "setdist s t = 0 \<longleftrightarrow> s = {} \<or> t = {} \<or> closure s \<inter> closure t \<noteq> {}"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12022 |
apply (cases "s = {} \<or> t = {}", force)
|
|
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
|
12023 |
using setdist_eq_0_compact_closed [of "closure s" "closure t"] |
|
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
|
12024 |
setdist_eq_0_closed_compact [of "closure s" "closure t"] assms |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12025 |
apply (force simp add: bounded_closure compact_eq_bounded_closed) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12026 |
done |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12027 |
|
|
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
|
12028 |
lemma setdist_unique: |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12029 |
"\<lbrakk>a \<in> s; b \<in> t; \<And>x y. x \<in> s \<and> y \<in> t ==> dist a b \<le> dist x y\<rbrakk> |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12030 |
\<Longrightarrow> setdist s t = dist a b" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12031 |
by (force simp add: setdist_le_dist le_setdist_iff intro: antisym) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12032 |
|
|
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
|
12033 |
lemma setdist_closest_point: |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12034 |
"\<lbrakk>closed s; s \<noteq> {}\<rbrakk> \<Longrightarrow> setdist {a} s = dist a (closest_point s a)"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12035 |
apply (rule setdist_unique) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12036 |
using closest_point_le |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12037 |
apply (auto simp: closest_point_in_set) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12038 |
done |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12039 |
|
| 63301 | 12040 |
lemma setdist_eq_0_sing_1: |
12041 |
fixes s :: "'a::euclidean_space set" |
|
12042 |
shows "setdist {x} s = 0 \<longleftrightarrow> s = {} \<or> x \<in> closure s"
|
|
12043 |
by (auto simp: setdist_eq_0_bounded) |
|
12044 |
||
12045 |
lemma setdist_eq_0_sing_2: |
|
12046 |
fixes s :: "'a::euclidean_space set" |
|
12047 |
shows "setdist s {x} = 0 \<longleftrightarrow> s = {} \<or> x \<in> closure s"
|
|
12048 |
by (auto simp: setdist_eq_0_bounded) |
|
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12049 |
|
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
12050 |
lemma setdist_neq_0_sing_1: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
12051 |
fixes s :: "'a::euclidean_space set" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
12052 |
shows "\<lbrakk>setdist {x} s = a; a \<noteq> 0\<rbrakk> \<Longrightarrow> s \<noteq> {} \<and> x \<notin> closure s"
|
| 63301 | 12053 |
by (auto simp: setdist_eq_0_sing_1) |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
12054 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
12055 |
lemma setdist_neq_0_sing_2: |
| 63301 | 12056 |
fixes s :: "'a::euclidean_space set" |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
12057 |
shows "\<lbrakk>setdist s {x} = a; a \<noteq> 0\<rbrakk> \<Longrightarrow> s \<noteq> {} \<and> x \<notin> closure s"
|
| 63301 | 12058 |
by (auto simp: setdist_eq_0_sing_2) |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
12059 |
|
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12060 |
lemma setdist_sing_in_set: |
| 63301 | 12061 |
fixes s :: "'a::euclidean_space set" |
12062 |
shows "x \<in> s \<Longrightarrow> setdist {x} s = 0"
|
|
12063 |
using closure_subset by (auto simp: setdist_eq_0_sing_1) |
|
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12064 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12065 |
lemma setdist_le_sing: "x \<in> s ==> setdist s t \<le> setdist {x} t"
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12066 |
using setdist_subset_left by auto |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60809
diff
changeset
|
12067 |
|
| 63301 | 12068 |
lemma setdist_eq_0_closed: |
12069 |
fixes s :: "'a::euclidean_space set" |
|
12070 |
shows "closed s \<Longrightarrow> (setdist {x} s = 0 \<longleftrightarrow> s = {} \<or> x \<in> s)"
|
|
12071 |
by (simp add: setdist_eq_0_sing_1) |
|
12072 |
||
12073 |
lemma setdist_eq_0_closedin: |
|
12074 |
fixes s :: "'a::euclidean_space set" |
|
12075 |
shows "\<lbrakk>closedin (subtopology euclidean u) s; x \<in> u\<rbrakk> |
|
12076 |
\<Longrightarrow> (setdist {x} s = 0 \<longleftrightarrow> s = {} \<or> x \<in> s)"
|
|
12077 |
by (auto simp: closedin_limpt setdist_eq_0_sing_1 closure_def) |
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12078 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12079 |
subsection\<open>Basic lemmas about hyperplanes and halfspaces\<close> |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12080 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12081 |
lemma hyperplane_eq_Ex: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12082 |
assumes "a \<noteq> 0" obtains x where "a \<bullet> x = b" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12083 |
by (rule_tac x = "(b / (a \<bullet> a)) *\<^sub>R a" in that) (simp add: assms) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12084 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12085 |
lemma hyperplane_eq_empty: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12086 |
"{x. a \<bullet> x = b} = {} \<longleftrightarrow> a = 0 \<and> b \<noteq> 0"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12087 |
using hyperplane_eq_Ex apply auto[1] |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12088 |
using inner_zero_right by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12089 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12090 |
lemma hyperplane_eq_UNIV: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12091 |
"{x. a \<bullet> x = b} = UNIV \<longleftrightarrow> a = 0 \<and> b = 0"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12092 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12093 |
have "UNIV \<subseteq> {x. a \<bullet> x = b} \<Longrightarrow> a = 0 \<and> b = 0"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12094 |
apply (drule_tac c = "((b+1) / (a \<bullet> a)) *\<^sub>R a" in subsetD) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12095 |
apply simp_all |
| 64240 | 12096 |
by (metis add_cancel_right_right div_by_1 zero_neq_one) |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12097 |
then show ?thesis by force |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12098 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12099 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12100 |
lemma halfspace_eq_empty_lt: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12101 |
"{x. a \<bullet> x < b} = {} \<longleftrightarrow> a = 0 \<and> b \<le> 0"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12102 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12103 |
have "{x. a \<bullet> x < b} \<subseteq> {} \<Longrightarrow> a = 0 \<and> b \<le> 0"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12104 |
apply (rule ccontr) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12105 |
apply (drule_tac c = "((b-1) / (a \<bullet> a)) *\<^sub>R a" in subsetD) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12106 |
apply force+ |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12107 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12108 |
then show ?thesis by force |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12109 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12110 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12111 |
lemma halfspace_eq_empty_gt: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12112 |
"{x. a \<bullet> x > b} = {} \<longleftrightarrow> a = 0 \<and> b \<ge> 0"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12113 |
using halfspace_eq_empty_lt [of "-a" "-b"] |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12114 |
by simp |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12115 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12116 |
lemma halfspace_eq_empty_le: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12117 |
"{x. a \<bullet> x \<le> b} = {} \<longleftrightarrow> a = 0 \<and> b < 0"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12118 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12119 |
have "{x. a \<bullet> x \<le> b} \<subseteq> {} \<Longrightarrow> a = 0 \<and> b < 0"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12120 |
apply (rule ccontr) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12121 |
apply (drule_tac c = "((b-1) / (a \<bullet> a)) *\<^sub>R a" in subsetD) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12122 |
apply force+ |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12123 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12124 |
then show ?thesis by force |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12125 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12126 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12127 |
lemma halfspace_eq_empty_ge: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12128 |
"{x. a \<bullet> x \<ge> b} = {} \<longleftrightarrow> a = 0 \<and> b > 0"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12129 |
using halfspace_eq_empty_le [of "-a" "-b"] |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12130 |
by simp |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12131 |
|
| 63301 | 12132 |
subsection\<open>Use set distance for an easy proof of separation properties\<close> |
12133 |
||
12134 |
proposition separation_closures: |
|
12135 |
fixes S :: "'a::euclidean_space set" |
|
12136 |
assumes "S \<inter> closure T = {}" "T \<inter> closure S = {}"
|
|
12137 |
obtains U V where "U \<inter> V = {}" "open U" "open V" "S \<subseteq> U" "T \<subseteq> V"
|
|
12138 |
proof (cases "S = {} \<or> T = {}")
|
|
12139 |
case True with that show ?thesis by auto |
|
12140 |
next |
|
12141 |
case False |
|
12142 |
define f where "f \<equiv> \<lambda>x. setdist {x} T - setdist {x} S"
|
|
| 63332 | 12143 |
have contf: "continuous_on UNIV f" |
12144 |
unfolding f_def by (intro continuous_intros continuous_on_setdist) |
|
| 63301 | 12145 |
show ?thesis |
12146 |
proof (rule_tac U = "{x. f x > 0}" and V = "{x. f x < 0}" in that)
|
|
12147 |
show "{x. 0 < f x} \<inter> {x. f x < 0} = {}"
|
|
12148 |
by auto |
|
12149 |
show "open {x. 0 < f x}"
|
|
| 63332 | 12150 |
by (simp add: open_Collect_less contf continuous_on_const) |
| 63301 | 12151 |
show "open {x. f x < 0}"
|
| 63332 | 12152 |
by (simp add: open_Collect_less contf continuous_on_const) |
| 63301 | 12153 |
show "S \<subseteq> {x. 0 < f x}"
|
12154 |
apply (clarsimp simp add: f_def setdist_sing_in_set) |
|
12155 |
using assms |
|
12156 |
by (metis False IntI empty_iff le_less setdist_eq_0_sing_2 setdist_pos_le setdist_sym) |
|
12157 |
show "T \<subseteq> {x. f x < 0}"
|
|
12158 |
apply (clarsimp simp add: f_def setdist_sing_in_set) |
|
12159 |
using assms |
|
12160 |
by (metis False IntI empty_iff le_less setdist_eq_0_sing_2 setdist_pos_le setdist_sym) |
|
12161 |
qed |
|
12162 |
qed |
|
12163 |
||
12164 |
lemma separation_normal: |
|
12165 |
fixes S :: "'a::euclidean_space set" |
|
12166 |
assumes "closed S" "closed T" "S \<inter> T = {}"
|
|
12167 |
obtains U V where "open U" "open V" "S \<subseteq> U" "T \<subseteq> V" "U \<inter> V = {}"
|
|
12168 |
using separation_closures [of S T] |
|
12169 |
by (metis assms closure_closed disjnt_def inf_commute) |
|
12170 |
||
12171 |
||
12172 |
lemma separation_normal_local: |
|
12173 |
fixes S :: "'a::euclidean_space set" |
|
12174 |
assumes US: "closedin (subtopology euclidean U) S" |
|
12175 |
and UT: "closedin (subtopology euclidean U) T" |
|
12176 |
and "S \<inter> T = {}"
|
|
12177 |
obtains S' T' where "openin (subtopology euclidean U) S'" |
|
12178 |
"openin (subtopology euclidean U) T'" |
|
12179 |
"S \<subseteq> S'" "T \<subseteq> T'" "S' \<inter> T' = {}"
|
|
12180 |
proof (cases "S = {} \<or> T = {}")
|
|
12181 |
case True with that show ?thesis |
|
12182 |
apply safe |
|
12183 |
using UT closedin_subset apply blast |
|
12184 |
using US closedin_subset apply blast |
|
12185 |
done |
|
12186 |
next |
|
12187 |
case False |
|
12188 |
define f where "f \<equiv> \<lambda>x. setdist {x} T - setdist {x} S"
|
|
12189 |
have contf: "continuous_on U f" |
|
12190 |
unfolding f_def by (intro continuous_intros) |
|
12191 |
show ?thesis |
|
12192 |
proof (rule_tac S' = "{x\<in>U. f x > 0}" and T' = "{x\<in>U. f x < 0}" in that)
|
|
12193 |
show "{x \<in> U. 0 < f x} \<inter> {x \<in> U. f x < 0} = {}"
|
|
12194 |
by auto |
|
12195 |
have "openin (subtopology euclidean U) {x \<in> U. f x \<in> {0<..}}"
|
|
| 64122 | 12196 |
by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) |
| 63301 | 12197 |
then show "openin (subtopology euclidean U) {x \<in> U. 0 < f x}" by simp
|
12198 |
next |
|
12199 |
have "openin (subtopology euclidean U) {x \<in> U. f x \<in> {..<0}}"
|
|
| 64122 | 12200 |
by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) |
| 63301 | 12201 |
then show "openin (subtopology euclidean U) {x \<in> U. f x < 0}" by simp
|
12202 |
next |
|
12203 |
show "S \<subseteq> {x \<in> U. 0 < f x}"
|
|
12204 |
apply (clarsimp simp add: f_def setdist_sing_in_set) |
|
12205 |
using assms |
|
12206 |
by (metis False Int_insert_right closedin_imp_subset empty_not_insert insert_absorb insert_subset linorder_neqE_linordered_idom not_le setdist_eq_0_closedin setdist_pos_le) |
|
12207 |
show "T \<subseteq> {x \<in> U. f x < 0}"
|
|
12208 |
apply (clarsimp simp add: f_def setdist_sing_in_set) |
|
12209 |
using assms |
|
12210 |
by (metis False closedin_subset disjoint_iff_not_equal insert_absorb insert_subset linorder_neqE_linordered_idom not_less setdist_eq_0_closedin setdist_pos_le topspace_euclidean_subtopology) |
|
12211 |
qed |
|
12212 |
qed |
|
12213 |
||
12214 |
lemma separation_normal_compact: |
|
12215 |
fixes S :: "'a::euclidean_space set" |
|
12216 |
assumes "compact S" "closed T" "S \<inter> T = {}"
|
|
12217 |
obtains U V where "open U" "compact(closure U)" "open V" "S \<subseteq> U" "T \<subseteq> V" "U \<inter> V = {}"
|
|
12218 |
proof - |
|
12219 |
have "closed S" "bounded S" |
|
12220 |
using assms by (auto simp: compact_eq_bounded_closed) |
|
12221 |
then obtain r where "r>0" and r: "S \<subseteq> ball 0 r" |
|
12222 |
by (auto dest!: bounded_subset_ballD) |
|
12223 |
have **: "closed (T \<union> - ball 0 r)" "S \<inter> (T \<union> - ball 0 r) = {}"
|
|
12224 |
using assms r by blast+ |
|
12225 |
then show ?thesis |
|
12226 |
apply (rule separation_normal [OF \<open>closed S\<close>]) |
|
12227 |
apply (rule_tac U=U and V=V in that) |
|
12228 |
by auto (meson bounded_ball bounded_subset compl_le_swap2 disjoint_eq_subset_Compl) |
|
12229 |
qed |
|
12230 |
||
12231 |
subsection\<open>Proper maps, including projections out of compact sets\<close> |
|
12232 |
||
12233 |
lemma finite_indexed_bound: |
|
12234 |
assumes A: "finite A" "\<And>x. x \<in> A \<Longrightarrow> \<exists>n::'a::linorder. P x n" |
|
12235 |
shows "\<exists>m. \<forall>x \<in> A. \<exists>k\<le>m. P x k" |
|
12236 |
using A |
|
12237 |
proof (induction A) |
|
12238 |
case empty then show ?case by force |
|
12239 |
next |
|
12240 |
case (insert a A) |
|
12241 |
then obtain m n where "\<forall>x \<in> A. \<exists>k\<le>m. P x k" "P a n" |
|
12242 |
by force |
|
12243 |
then show ?case |
|
12244 |
apply (rule_tac x="max m n" in exI, safe) |
|
12245 |
using max.cobounded2 apply blast |
|
12246 |
by (meson le_max_iff_disj) |
|
12247 |
qed |
|
12248 |
||
12249 |
proposition proper_map: |
|
12250 |
fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel" |
|
12251 |
assumes "closedin (subtopology euclidean S) K" |
|
12252 |
and com: "\<And>U. \<lbrakk>U \<subseteq> T; compact U\<rbrakk> \<Longrightarrow> compact {x \<in> S. f x \<in> U}"
|
|
12253 |
and "f ` S \<subseteq> T" |
|
12254 |
shows "closedin (subtopology euclidean T) (f ` K)" |
|
12255 |
proof - |
|
12256 |
have "K \<subseteq> S" |
|
12257 |
using assms closedin_imp_subset by metis |
|
12258 |
obtain C where "closed C" and Keq: "K = S \<inter> C" |
|
12259 |
using assms by (auto simp: closedin_closed) |
|
12260 |
have *: "y \<in> f ` K" if "y \<in> T" and y: "y islimpt f ` K" for y |
|
12261 |
proof - |
|
12262 |
obtain h where "\<forall>n. (\<exists>x\<in>K. h n = f x) \<and> h n \<noteq> y" "inj h" and hlim: "(h \<longlongrightarrow> y) sequentially" |
|
12263 |
using \<open>y \<in> T\<close> y by (force simp: limpt_sequential_inj) |
|
12264 |
then obtain X where X: "\<And>n. X n \<in> K \<and> h n = f (X n) \<and> h n \<noteq> y" |
|
12265 |
by metis |
|
12266 |
then have fX: "\<And>n. f (X n) = h n" |
|
12267 |
by metis |
|
12268 |
have "compact (C \<inter> {a \<in> S. f a \<in> insert y (range (\<lambda>i. f(X(n + i))))})" for n
|
|
12269 |
apply (rule closed_Int_compact [OF \<open>closed C\<close>]) |
|
12270 |
apply (rule com) |
|
12271 |
using X \<open>K \<subseteq> S\<close> \<open>f ` S \<subseteq> T\<close> \<open>y \<in> T\<close> apply blast |
|
12272 |
apply (rule compact_sequence_with_limit) |
|
12273 |
apply (simp add: fX add.commute [of n] LIMSEQ_ignore_initial_segment [OF hlim]) |
|
12274 |
done |
|
12275 |
then have comf: "compact {a \<in> K. f a \<in> insert y (range (\<lambda>i. f(X(n + i))))}" for n
|
|
12276 |
by (simp add: Keq Int_def conj_commute) |
|
12277 |
have ne: "\<Inter>\<F> \<noteq> {}"
|
|
12278 |
if "finite \<F>" |
|
12279 |
and \<F>: "\<And>t. t \<in> \<F> \<Longrightarrow> |
|
12280 |
(\<exists>n. t = {a \<in> K. f a \<in> insert y (range (\<lambda>i. f (X (n + i))))})"
|
|
12281 |
for \<F> |
|
12282 |
proof - |
|
12283 |
obtain m where m: "\<And>t. t \<in> \<F> \<Longrightarrow> \<exists>k\<le>m. t = {a \<in> K. f a \<in> insert y (range (\<lambda>i. f (X (k + i))))}"
|
|
12284 |
apply (rule exE) |
|
12285 |
apply (rule finite_indexed_bound [OF \<open>finite \<F>\<close> \<F>], assumption, force) |
|
12286 |
done |
|
12287 |
have "X m \<in> \<Inter>\<F>" |
|
12288 |
using X le_Suc_ex by (fastforce dest: m) |
|
12289 |
then show ?thesis by blast |
|
12290 |
qed |
|
12291 |
have "\<Inter>{{a. a \<in> K \<and> f a \<in> insert y (range (\<lambda>i. f(X(n + i))))} |n. n \<in> UNIV}
|
|
12292 |
\<noteq> {}"
|
|
12293 |
apply (rule compact_fip_heine_borel) |
|
12294 |
using comf apply force |
|
12295 |
using ne apply (simp add: subset_iff del: insert_iff) |
|
12296 |
done |
|
12297 |
then have "\<exists>x. x \<in> (\<Inter>n. {a \<in> K. f a \<in> insert y (range (\<lambda>i. f (X (n + i))))})"
|
|
12298 |
by blast |
|
12299 |
then show ?thesis |
|
12300 |
apply (simp add: image_iff fX) |
|
12301 |
by (metis \<open>inj h\<close> le_add1 not_less_eq_eq rangeI range_ex1_eq) |
|
12302 |
qed |
|
12303 |
with assms closedin_subset show ?thesis |
|
12304 |
by (force simp: closedin_limpt) |
|
12305 |
qed |
|
12306 |
||
12307 |
||
12308 |
lemma compact_continuous_image_eq: |
|
12309 |
fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel" |
|
12310 |
assumes f: "inj_on f S" |
|
12311 |
shows "continuous_on S f \<longleftrightarrow> (\<forall>T. compact T \<and> T \<subseteq> S \<longrightarrow> compact(f ` T))" |
|
12312 |
(is "?lhs = ?rhs") |
|
12313 |
proof |
|
12314 |
assume ?lhs then show ?rhs |
|
12315 |
by (metis continuous_on_subset compact_continuous_image) |
|
12316 |
next |
|
12317 |
assume RHS: ?rhs |
|
12318 |
obtain g where gf: "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x" |
|
12319 |
by (metis inv_into_f_f f) |
|
12320 |
then have *: "{x \<in> S. f x \<in> U} = g ` U" if "U \<subseteq> f ` S" for U
|
|
12321 |
using that by fastforce |
|
12322 |
have gfim: "g ` f ` S \<subseteq> S" using gf by auto |
|
12323 |
have **: "compact {x \<in> f ` S. g x \<in> C}" if C: "C \<subseteq> S" "compact C" for C
|
|
12324 |
proof - |
|
12325 |
obtain h :: "'a set \<Rightarrow> 'a" where "h C \<in> C \<and> h C \<notin> S \<or> compact (f ` C)" |
|
12326 |
by (force simp: C RHS) |
|
12327 |
moreover have "f ` C = {b \<in> f ` S. g b \<in> C}"
|
|
12328 |
using C gf by auto |
|
12329 |
ultimately show "compact {b \<in> f ` S. g b \<in> C}"
|
|
12330 |
using C by auto |
|
12331 |
qed |
|
12332 |
show ?lhs |
|
12333 |
using proper_map [OF _ _ gfim] ** |
|
12334 |
by (simp add: continuous_on_closed * closedin_imp_subset) |
|
12335 |
qed |
|
12336 |
||
12337 |
lemma proper_map_from_compact: |
|
12338 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
12339 |
assumes contf: "continuous_on S f" and imf: "f ` S \<subseteq> T" and "compact S" |
|
12340 |
"closedin (subtopology euclidean T) K" |
|
12341 |
shows "compact {x. x \<in> S \<and> f x \<in> K}"
|
|
12342 |
by (rule closedin_compact [OF \<open>compact S\<close>] continuous_closedin_preimage_gen assms)+ |
|
12343 |
||
12344 |
lemma proper_map_fst: |
|
12345 |
assumes "compact T" "K \<subseteq> S" "compact K" |
|
12346 |
shows "compact {z \<in> S \<times> T. fst z \<in> K}"
|
|
12347 |
proof - |
|
12348 |
have "{z \<in> S \<times> T. fst z \<in> K} = K \<times> T"
|
|
12349 |
using assms by auto |
|
12350 |
then show ?thesis by (simp add: assms compact_Times) |
|
12351 |
qed |
|
12352 |
||
12353 |
lemma closed_map_fst: |
|
12354 |
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" |
|
12355 |
assumes "compact T" "closedin (subtopology euclidean (S \<times> T)) c" |
|
12356 |
shows "closedin (subtopology euclidean S) (fst ` c)" |
|
12357 |
proof - |
|
12358 |
have *: "fst ` (S \<times> T) \<subseteq> S" |
|
12359 |
by auto |
|
12360 |
show ?thesis |
|
12361 |
using proper_map [OF _ _ *] by (simp add: proper_map_fst assms) |
|
12362 |
qed |
|
12363 |
||
12364 |
lemma proper_map_snd: |
|
12365 |
assumes "compact S" "K \<subseteq> T" "compact K" |
|
12366 |
shows "compact {z \<in> S \<times> T. snd z \<in> K}"
|
|
12367 |
proof - |
|
12368 |
have "{z \<in> S \<times> T. snd z \<in> K} = S \<times> K"
|
|
12369 |
using assms by auto |
|
12370 |
then show ?thesis by (simp add: assms compact_Times) |
|
12371 |
qed |
|
12372 |
||
12373 |
lemma closed_map_snd: |
|
12374 |
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" |
|
12375 |
assumes "compact S" "closedin (subtopology euclidean (S \<times> T)) c" |
|
12376 |
shows "closedin (subtopology euclidean T) (snd ` c)" |
|
12377 |
proof - |
|
12378 |
have *: "snd ` (S \<times> T) \<subseteq> T" |
|
12379 |
by auto |
|
12380 |
show ?thesis |
|
12381 |
using proper_map [OF _ _ *] by (simp add: proper_map_snd assms) |
|
12382 |
qed |
|
12383 |
||
12384 |
lemma closedin_compact_projection: |
|
12385 |
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" |
|
12386 |
assumes "compact S" and clo: "closedin (subtopology euclidean (S \<times> T)) U" |
|
12387 |
shows "closedin (subtopology euclidean T) {y. \<exists>x. x \<in> S \<and> (x, y) \<in> U}"
|
|
12388 |
proof - |
|
12389 |
have "U \<subseteq> S \<times> T" |
|
12390 |
by (metis clo closedin_imp_subset) |
|
12391 |
then have "{y. \<exists>x. x \<in> S \<and> (x, y) \<in> U} = snd ` U"
|
|
12392 |
by force |
|
12393 |
moreover have "closedin (subtopology euclidean T) (snd ` U)" |
|
12394 |
by (rule closed_map_snd [OF assms]) |
|
12395 |
ultimately show ?thesis |
|
12396 |
by simp |
|
12397 |
qed |
|
12398 |
||
12399 |
||
12400 |
lemma closed_compact_projection: |
|
12401 |
fixes S :: "'a::euclidean_space set" |
|
12402 |
and T :: "('a * 'b::euclidean_space) set"
|
|
12403 |
assumes "compact S" and clo: "closed T" |
|
12404 |
shows "closed {y. \<exists>x. x \<in> S \<and> (x, y) \<in> T}"
|
|
12405 |
proof - |
|
12406 |
have *: "{y. \<exists>x. x \<in> S \<and> Pair x y \<in> T} =
|
|
12407 |
{y. \<exists>x. x \<in> S \<and> Pair x y \<in> ((S \<times> UNIV) \<inter> T)}"
|
|
12408 |
by auto |
|
12409 |
show ?thesis |
|
12410 |
apply (subst *) |
|
12411 |
apply (rule closedin_closed_trans [OF _ closed_UNIV]) |
|
12412 |
apply (rule closedin_compact_projection [OF \<open>compact S\<close>]) |
|
12413 |
by (simp add: clo closedin_closed_Int) |
|
12414 |
qed |
|
12415 |
||
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12416 |
subsubsection\<open>Representing affine hull as a finite intersection of hyperplanes\<close> |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12417 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12418 |
proposition affine_hull_convex_Int_nonempty_interior: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12419 |
fixes S :: "'a::real_normed_vector set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12420 |
assumes "convex S" "S \<inter> interior T \<noteq> {}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12421 |
shows "affine hull (S \<inter> T) = affine hull S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12422 |
proof |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12423 |
show "affine hull (S \<inter> T) \<subseteq> affine hull S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12424 |
by (simp add: hull_mono) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12425 |
next |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12426 |
obtain a where "a \<in> S" "a \<in> T" and at: "a \<in> interior T" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12427 |
using assms interior_subset by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12428 |
then obtain e where "e > 0" and e: "cball a e \<subseteq> T" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12429 |
using mem_interior_cball by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12430 |
have *: "x \<in> op + a ` span ((\<lambda>x. x - a) ` (S \<inter> T))" if "x \<in> S" for x |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12431 |
proof (cases "x = a") |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12432 |
case True with that span_0 eq_add_iff image_def mem_Collect_eq show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12433 |
by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12434 |
next |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12435 |
case False |
| 63040 | 12436 |
define k where "k = min (1/2) (e / norm (x-a))" |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12437 |
have k: "0 < k" "k < 1" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12438 |
using \<open>e > 0\<close> False by (auto simp: k_def) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12439 |
then have xa: "(x-a) = inverse k *\<^sub>R k *\<^sub>R (x-a)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12440 |
by simp |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12441 |
have "e / norm (x - a) \<ge> k" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12442 |
using k_def by linarith |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12443 |
then have "a + k *\<^sub>R (x - a) \<in> cball a e" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12444 |
using \<open>0 < k\<close> False by (simp add: dist_norm field_simps) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12445 |
then have T: "a + k *\<^sub>R (x - a) \<in> T" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12446 |
using e by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12447 |
have S: "a + k *\<^sub>R (x - a) \<in> S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12448 |
using k \<open>a \<in> S\<close> convexD [OF \<open>convex S\<close> \<open>a \<in> S\<close> \<open>x \<in> S\<close>, of "1-k" k] |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12449 |
by (simp add: algebra_simps) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12450 |
have "inverse k *\<^sub>R k *\<^sub>R (x-a) \<in> span ((\<lambda>x. x - a) ` (S \<inter> T))" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12451 |
apply (rule span_mul) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12452 |
apply (rule span_superset) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12453 |
apply (rule image_eqI [where x = "a + k *\<^sub>R (x - a)"]) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12454 |
apply (auto simp: S T) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12455 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12456 |
with xa image_iff show ?thesis by fastforce |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12457 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12458 |
show "affine hull S \<subseteq> affine hull (S \<inter> T)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12459 |
apply (simp add: subset_hull) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12460 |
apply (simp add: \<open>a \<in> S\<close> \<open>a \<in> T\<close> hull_inc affine_hull_span_gen [of a]) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12461 |
apply (force simp: *) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12462 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12463 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12464 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12465 |
corollary affine_hull_convex_Int_open: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12466 |
fixes S :: "'a::real_normed_vector set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12467 |
assumes "convex S" "open T" "S \<inter> T \<noteq> {}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12468 |
shows "affine hull (S \<inter> T) = affine hull S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12469 |
using affine_hull_convex_Int_nonempty_interior assms interior_eq by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12470 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12471 |
corollary affine_hull_affine_Int_nonempty_interior: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12472 |
fixes S :: "'a::real_normed_vector set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12473 |
assumes "affine S" "S \<inter> interior T \<noteq> {}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12474 |
shows "affine hull (S \<inter> T) = affine hull S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12475 |
by (simp add: affine_hull_convex_Int_nonempty_interior affine_imp_convex assms) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12476 |
|
|
63018
ae2ec7d86ad4
tidying some proofs; getting rid of "nonempty_witness"
paulson <lp15@cam.ac.uk>
parents:
63016
diff
changeset
|
12477 |
corollary affine_hull_affine_Int_open: |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12478 |
fixes S :: "'a::real_normed_vector set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12479 |
assumes "affine S" "open T" "S \<inter> T \<noteq> {}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12480 |
shows "affine hull (S \<inter> T) = affine hull S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12481 |
by (simp add: affine_hull_convex_Int_open affine_imp_convex assms) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12482 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12483 |
corollary affine_hull_convex_Int_openin: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12484 |
fixes S :: "'a::real_normed_vector set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12485 |
assumes "convex S" "openin (subtopology euclidean (affine hull S)) T" "S \<inter> T \<noteq> {}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12486 |
shows "affine hull (S \<inter> T) = affine hull S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12487 |
using assms unfolding openin_open |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12488 |
by (metis affine_hull_convex_Int_open hull_subset inf.orderE inf_assoc) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12489 |
|
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
12490 |
corollary affine_hull_openin: |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12491 |
fixes S :: "'a::real_normed_vector set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12492 |
assumes "openin (subtopology euclidean (affine hull T)) S" "S \<noteq> {}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12493 |
shows "affine hull S = affine hull T" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12494 |
using assms unfolding openin_open |
|
63018
ae2ec7d86ad4
tidying some proofs; getting rid of "nonempty_witness"
paulson <lp15@cam.ac.uk>
parents:
63016
diff
changeset
|
12495 |
by (metis affine_affine_hull affine_hull_affine_Int_open hull_hull) |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12496 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12497 |
corollary affine_hull_open: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12498 |
fixes S :: "'a::real_normed_vector set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12499 |
assumes "open S" "S \<noteq> {}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12500 |
shows "affine hull S = UNIV" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12501 |
by (metis affine_hull_convex_Int_nonempty_interior assms convex_UNIV hull_UNIV inf_top.left_neutral interior_open) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12502 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12503 |
lemma aff_dim_convex_Int_nonempty_interior: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12504 |
fixes S :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12505 |
shows "\<lbrakk>convex S; S \<inter> interior T \<noteq> {}\<rbrakk> \<Longrightarrow> aff_dim(S \<inter> T) = aff_dim S"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12506 |
using aff_dim_affine_hull2 affine_hull_convex_Int_nonempty_interior by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12507 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12508 |
lemma aff_dim_convex_Int_open: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12509 |
fixes S :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12510 |
shows "\<lbrakk>convex S; open T; S \<inter> T \<noteq> {}\<rbrakk> \<Longrightarrow> aff_dim(S \<inter> T) = aff_dim S"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12511 |
using aff_dim_convex_Int_nonempty_interior interior_eq by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12512 |
|
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
12513 |
lemma affine_hull_Diff: |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
12514 |
fixes S:: "'a::real_normed_vector set" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
12515 |
assumes ope: "openin (subtopology euclidean (affine hull S)) S" and "finite F" "F \<subset> S" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
12516 |
shows "affine hull (S - F) = affine hull S" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
12517 |
proof - |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
12518 |
have clo: "closedin (subtopology euclidean S) F" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
12519 |
using assms finite_imp_closedin by auto |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
12520 |
moreover have "S - F \<noteq> {}"
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
12521 |
using assms by auto |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
12522 |
ultimately show ?thesis |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
12523 |
by (metis ope closedin_def topspace_euclidean_subtopology affine_hull_openin openin_trans) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
12524 |
qed |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
12525 |
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12526 |
lemma affine_hull_halfspace_lt: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12527 |
fixes a :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12528 |
shows "affine hull {x. a \<bullet> x < r} = (if a = 0 \<and> r \<le> 0 then {} else UNIV)"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12529 |
using halfspace_eq_empty_lt [of a r] |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12530 |
by (simp add: open_halfspace_lt affine_hull_open) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12531 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12532 |
lemma affine_hull_halfspace_le: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12533 |
fixes a :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12534 |
shows "affine hull {x. a \<bullet> x \<le> r} = (if a = 0 \<and> r < 0 then {} else UNIV)"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12535 |
proof (cases "a = 0") |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12536 |
case True then show ?thesis by simp |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12537 |
next |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12538 |
case False |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12539 |
then have "affine hull closure {x. a \<bullet> x < r} = UNIV"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12540 |
using affine_hull_halfspace_lt closure_same_affine_hull by fastforce |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12541 |
moreover have "{x. a \<bullet> x < r} \<subseteq> {x. a \<bullet> x \<le> r}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12542 |
by (simp add: Collect_mono) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12543 |
ultimately show ?thesis using False antisym_conv hull_mono top_greatest |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12544 |
by (metis affine_hull_halfspace_lt) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12545 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12546 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12547 |
lemma affine_hull_halfspace_gt: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12548 |
fixes a :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12549 |
shows "affine hull {x. a \<bullet> x > r} = (if a = 0 \<and> r \<ge> 0 then {} else UNIV)"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12550 |
using halfspace_eq_empty_gt [of r a] |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12551 |
by (simp add: open_halfspace_gt affine_hull_open) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12552 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12553 |
lemma affine_hull_halfspace_ge: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12554 |
fixes a :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12555 |
shows "affine hull {x. a \<bullet> x \<ge> r} = (if a = 0 \<and> r > 0 then {} else UNIV)"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12556 |
using affine_hull_halfspace_le [of "-a" "-r"] by simp |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12557 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12558 |
lemma aff_dim_halfspace_lt: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12559 |
fixes a :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12560 |
shows "aff_dim {x. a \<bullet> x < r} =
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12561 |
(if a = 0 \<and> r \<le> 0 then -1 else DIM('a))"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12562 |
by simp (metis aff_dim_open halfspace_eq_empty_lt open_halfspace_lt) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12563 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12564 |
lemma aff_dim_halfspace_le: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12565 |
fixes a :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12566 |
shows "aff_dim {x. a \<bullet> x \<le> r} =
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12567 |
(if a = 0 \<and> r < 0 then -1 else DIM('a))"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12568 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12569 |
have "int (DIM('a)) = aff_dim (UNIV::'a set)"
|
| 63072 | 12570 |
by (simp add: aff_dim_UNIV) |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12571 |
then have "aff_dim (affine hull {x. a \<bullet> x \<le> r}) = DIM('a)" if "(a = 0 \<longrightarrow> r \<ge> 0)"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12572 |
using that by (simp add: affine_hull_halfspace_le not_less) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12573 |
then show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12574 |
by (force simp: aff_dim_affine_hull) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12575 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12576 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12577 |
lemma aff_dim_halfspace_gt: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12578 |
fixes a :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12579 |
shows "aff_dim {x. a \<bullet> x > r} =
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12580 |
(if a = 0 \<and> r \<ge> 0 then -1 else DIM('a))"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12581 |
by simp (metis aff_dim_open halfspace_eq_empty_gt open_halfspace_gt) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12582 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12583 |
lemma aff_dim_halfspace_ge: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12584 |
fixes a :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12585 |
shows "aff_dim {x. a \<bullet> x \<ge> r} =
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12586 |
(if a = 0 \<and> r > 0 then -1 else DIM('a))"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12587 |
using aff_dim_halfspace_le [of "-a" "-r"] by simp |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12588 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12589 |
subsection\<open>Properties of special hyperplanes\<close> |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12590 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12591 |
lemma subspace_hyperplane: "subspace {x. a \<bullet> x = 0}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12592 |
by (simp add: subspace_def inner_right_distrib) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12593 |
|
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
12594 |
lemma subspace_hyperplane2: "subspace {x. x \<bullet> a = 0}"
|
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
12595 |
by (simp add: inner_commute inner_right_distrib subspace_def) |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
12596 |
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12597 |
lemma special_hyperplane_span: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12598 |
fixes S :: "'n::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12599 |
assumes "k \<in> Basis" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12600 |
shows "{x. k \<bullet> x = 0} = span (Basis - {k})"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12601 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12602 |
have *: "x \<in> span (Basis - {k})" if "k \<bullet> x = 0" for x
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12603 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12604 |
have "x = (\<Sum>b\<in>Basis. (x \<bullet> b) *\<^sub>R b)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12605 |
by (simp add: euclidean_representation) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12606 |
also have "... = (\<Sum>b \<in> Basis - {k}. (x \<bullet> b) *\<^sub>R b)"
|
| 64267 | 12607 |
by (auto simp: sum.remove [of _ k] inner_commute assms that) |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12608 |
finally have "x = (\<Sum>b\<in>Basis - {k}. (x \<bullet> b) *\<^sub>R b)" .
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12609 |
then show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12610 |
by (simp add: Linear_Algebra.span_finite) metis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12611 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12612 |
show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12613 |
apply (rule span_subspace [symmetric]) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12614 |
using assms |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12615 |
apply (auto simp: inner_not_same_Basis intro: * subspace_hyperplane) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12616 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12617 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12618 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12619 |
lemma dim_special_hyperplane: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12620 |
fixes k :: "'n::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12621 |
shows "k \<in> Basis \<Longrightarrow> dim {x. k \<bullet> x = 0} = DIM('n) - 1"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12622 |
apply (simp add: special_hyperplane_span) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12623 |
apply (rule Linear_Algebra.dim_unique [OF subset_refl]) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12624 |
apply (auto simp: Diff_subset independent_substdbasis) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12625 |
apply (metis member_remove remove_def span_clauses(1)) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12626 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12627 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12628 |
proposition dim_hyperplane: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12629 |
fixes a :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12630 |
assumes "a \<noteq> 0" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12631 |
shows "dim {x. a \<bullet> x = 0} = DIM('a) - 1"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12632 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12633 |
have span0: "span {x. a \<bullet> x = 0} = {x. a \<bullet> x = 0}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12634 |
by (rule span_unique) (auto simp: subspace_hyperplane) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12635 |
then obtain B where "independent B" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12636 |
and Bsub: "B \<subseteq> {x. a \<bullet> x = 0}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12637 |
and subspB: "{x. a \<bullet> x = 0} \<subseteq> span B"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12638 |
and card0: "(card B = dim {x. a \<bullet> x = 0})"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12639 |
and ortho: "pairwise orthogonal B" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12640 |
using orthogonal_basis_exists by metis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12641 |
with assms have "a \<notin> span B" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12642 |
by (metis (mono_tags, lifting) span_eq inner_eq_zero_iff mem_Collect_eq span0 span_subspace) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12643 |
then have ind: "independent (insert a B)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12644 |
by (simp add: \<open>independent B\<close> independent_insert) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12645 |
have "finite B" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12646 |
using \<open>independent B\<close> independent_bound by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12647 |
have "UNIV \<subseteq> span (insert a B)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12648 |
proof fix y::'a |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12649 |
obtain r z where z: "y = r *\<^sub>R a + z" "a \<bullet> z = 0" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12650 |
apply (rule_tac r="(a \<bullet> y) / (a \<bullet> a)" and z = "y - ((a \<bullet> y) / (a \<bullet> a)) *\<^sub>R a" in that) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12651 |
using assms |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12652 |
by (auto simp: algebra_simps) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12653 |
show "y \<in> span (insert a B)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12654 |
by (metis (mono_tags, lifting) z Bsub Convex_Euclidean_Space.span_eq |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12655 |
add_diff_cancel_left' mem_Collect_eq span0 span_breakdown_eq span_subspace subspB) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12656 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12657 |
then have dima: "DIM('a) = dim(insert a B)"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12658 |
by (metis antisym dim_UNIV dim_subset_UNIV subset_le_dim) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12659 |
then show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12660 |
by (metis (mono_tags, lifting) Bsub Diff_insert_absorb \<open>a \<notin> span B\<close> ind card0 card_Diff_singleton dim_span indep_card_eq_dim_span insertI1 subsetCE subspB) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12661 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12662 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12663 |
lemma lowdim_eq_hyperplane: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12664 |
fixes S :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12665 |
assumes "dim S = DIM('a) - 1"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12666 |
obtains a where "a \<noteq> 0" and "span S = {x. a \<bullet> x = 0}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12667 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12668 |
have [simp]: "dim S < DIM('a)"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12669 |
by (simp add: DIM_positive assms) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12670 |
then obtain b where b: "b \<noteq> 0" "span S \<subseteq> {a. b \<bullet> a = 0}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12671 |
using lowdim_subset_hyperplane [of S] by fastforce |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12672 |
show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12673 |
using b that real_vector_class.subspace_span [of S] |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12674 |
by (simp add: assms dim_hyperplane subspace_dim_equal subspace_hyperplane) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12675 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12676 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12677 |
lemma dim_eq_hyperplane: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12678 |
fixes S :: "'n::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12679 |
shows "dim S = DIM('n) - 1 \<longleftrightarrow> (\<exists>a. a \<noteq> 0 \<and> span S = {x. a \<bullet> x = 0})"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12680 |
by (metis One_nat_def dim_hyperplane dim_span lowdim_eq_hyperplane) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12681 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12682 |
proposition aff_dim_eq_hyperplane: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12683 |
fixes S :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12684 |
shows "aff_dim S = DIM('a) - 1 \<longleftrightarrow> (\<exists>a b. a \<noteq> 0 \<and> affine hull S = {x. a \<bullet> x = b})"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12685 |
proof (cases "S = {}")
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12686 |
case True then show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12687 |
by (auto simp: dest: hyperplane_eq_Ex) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12688 |
next |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12689 |
case False |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12690 |
then obtain c where "c \<in> S" by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12691 |
show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12692 |
proof (cases "c = 0") |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12693 |
case True show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12694 |
apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \<open>c \<in> S\<close> hull_inc dim_eq_hyperplane |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12695 |
del: One_nat_def) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12696 |
apply (rule ex_cong) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
12697 |
apply (metis (mono_tags) span_0 \<open>c = 0\<close> image_add_0 inner_zero_right mem_Collect_eq) |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12698 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12699 |
next |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12700 |
case False |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12701 |
have xc_im: "x \<in> op + c ` {y. a \<bullet> y = 0}" if "a \<bullet> x = a \<bullet> c" for a x
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12702 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12703 |
have "\<exists>y. a \<bullet> y = 0 \<and> c + y = x" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12704 |
by (metis that add.commute diff_add_cancel inner_commute inner_diff_left right_minus_eq) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12705 |
then show "x \<in> op + c ` {y. a \<bullet> y = 0}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12706 |
by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12707 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12708 |
have 2: "span ((\<lambda>x. x - c) ` S) = {x. a \<bullet> x = 0}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12709 |
if "op + c ` span ((\<lambda>x. x - c) ` S) = {x. a \<bullet> x = b}" for a b
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12710 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12711 |
have "b = a \<bullet> c" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12712 |
using span_0 that by fastforce |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12713 |
with that have "op + c ` span ((\<lambda>x. x - c) ` S) = {x. a \<bullet> x = a \<bullet> c}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12714 |
by simp |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12715 |
then have "span ((\<lambda>x. x - c) ` S) = (\<lambda>x. x - c) ` {x. a \<bullet> x = a \<bullet> c}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12716 |
by (metis (no_types) image_cong translation_galois uminus_add_conv_diff) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12717 |
also have "... = {x. a \<bullet> x = 0}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12718 |
by (force simp: inner_distrib inner_diff_right |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12719 |
intro: image_eqI [where x="x+c" for x]) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12720 |
finally show ?thesis . |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12721 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12722 |
show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12723 |
apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \<open>c \<in> S\<close> hull_inc dim_eq_hyperplane |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12724 |
del: One_nat_def, safe) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12725 |
apply (fastforce simp add: inner_distrib intro: xc_im) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12726 |
apply (force simp: intro!: 2) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12727 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12728 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12729 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12730 |
|
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
12731 |
corollary aff_dim_hyperplane [simp]: |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12732 |
fixes a :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12733 |
shows "a \<noteq> 0 \<Longrightarrow> aff_dim {x. a \<bullet> x = r} = DIM('a) - 1"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12734 |
by (metis aff_dim_eq_hyperplane affine_hull_eq affine_hyperplane) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12735 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12736 |
subsection\<open>Some stepping theorems\<close> |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12737 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12738 |
lemma dim_empty [simp]: "dim ({} :: 'a::euclidean_space set) = 0"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12739 |
by (force intro!: dim_unique) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12740 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12741 |
lemma dim_insert: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12742 |
fixes x :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12743 |
shows "dim (insert x S) = (if x \<in> span S then dim S else dim S + 1)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12744 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12745 |
show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12746 |
proof (cases "x \<in> span S") |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12747 |
case True then show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12748 |
by (metis dim_span span_redundant) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12749 |
next |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12750 |
case False |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12751 |
obtain B where B: "B \<subseteq> span S" "independent B" "span S \<subseteq> span B" "card B = dim (span S)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12752 |
using basis_exists [of "span S"] by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12753 |
have 1: "insert x B \<subseteq> span (insert x S)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12754 |
by (meson \<open>B \<subseteq> span S\<close> dual_order.trans insertI1 insert_subsetI span_mono span_superset subset_insertI) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12755 |
have 2: "span (insert x S) \<subseteq> span (insert x B)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12756 |
by (metis \<open>B \<subseteq> span S\<close> \<open>span S \<subseteq> span B\<close> span_breakdown_eq span_subspace subsetI subspace_span) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12757 |
have 3: "independent (insert x B)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12758 |
by (metis B independent_insert span_subspace subspace_span False) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12759 |
have "dim (span (insert x S)) = Suc (dim S)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12760 |
apply (rule dim_unique [OF 1 2 3]) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12761 |
by (metis B False card_insert_disjoint dim_span independent_imp_finite subsetCE) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12762 |
then show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12763 |
by (simp add: False) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12764 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12765 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12766 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12767 |
lemma dim_singleton [simp]: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12768 |
fixes x :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12769 |
shows "dim{x} = (if x = 0 then 0 else 1)"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12770 |
by (simp add: dim_insert) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12771 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12772 |
lemma dim_eq_0 [simp]: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12773 |
fixes S :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12774 |
shows "dim S = 0 \<longleftrightarrow> S \<subseteq> {0}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12775 |
apply safe |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12776 |
apply (metis DIM_positive DIM_real card_ge_dim_independent contra_subsetD dim_empty dim_insert dim_singleton empty_subsetI independent_empty less_not_refl zero_le) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12777 |
by (metis dim_singleton dim_subset le_0_eq) |
| 64122 | 12778 |
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12779 |
lemma aff_dim_insert: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12780 |
fixes a :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12781 |
shows "aff_dim (insert a S) = (if a \<in> affine hull S then aff_dim S else aff_dim S + 1)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12782 |
proof (cases "S = {}")
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12783 |
case True then show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12784 |
by simp |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12785 |
next |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12786 |
case False |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12787 |
then obtain x s' where S: "S = insert x s'" "x \<notin> s'" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12788 |
by (meson Set.set_insert all_not_in_conv) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12789 |
show ?thesis using S |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12790 |
apply (simp add: hull_redundant cong: aff_dim_affine_hull2) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12791 |
apply (simp add: affine_hull_insert_span_gen hull_inc) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12792 |
apply (simp add: insert_commute [of a] hull_inc aff_dim_eq_dim [of x] dim_insert span_0) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12793 |
apply (metis (no_types, lifting) add_minus_cancel image_iff uminus_add_conv_diff) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12794 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12795 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12796 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12797 |
lemma subspace_bounded_eq_trivial: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12798 |
fixes S :: "'a::real_normed_vector set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12799 |
assumes "subspace S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12800 |
shows "bounded S \<longleftrightarrow> S = {0}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12801 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12802 |
have "False" if "bounded S" "x \<in> S" "x \<noteq> 0" for x |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12803 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12804 |
obtain B where B: "\<And>y. y \<in> S \<Longrightarrow> norm y < B" "B > 0" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12805 |
using \<open>bounded S\<close> by (force simp: bounded_pos_less) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12806 |
have "(B / norm x) *\<^sub>R x \<in> S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12807 |
using assms subspace_mul \<open>x \<in> S\<close> by auto |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12808 |
moreover have "norm ((B / norm x) *\<^sub>R x) = B" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12809 |
using that B by (simp add: algebra_simps) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12810 |
ultimately show False using B by force |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12811 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12812 |
then have "bounded S \<Longrightarrow> S = {0}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12813 |
using assms subspace_0 by fastforce |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12814 |
then show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12815 |
by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12816 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12817 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12818 |
lemma affine_bounded_eq_trivial: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12819 |
fixes S :: "'a::real_normed_vector set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12820 |
assumes "affine S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12821 |
shows "bounded S \<longleftrightarrow> S = {} \<or> (\<exists>a. S = {a})"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12822 |
proof (cases "S = {}")
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12823 |
case True then show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12824 |
by simp |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12825 |
next |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12826 |
case False |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12827 |
then obtain b where "b \<in> S" by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12828 |
with False assms show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12829 |
apply safe |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12830 |
using affine_diffs_subspace [OF assms \<open>b \<in> S\<close>] |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12831 |
apply (metis (no_types, lifting) subspace_bounded_eq_trivial ab_left_minus bounded_translation |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12832 |
image_empty image_insert translation_invert) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12833 |
apply force |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12834 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12835 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12836 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12837 |
lemma affine_bounded_eq_lowdim: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12838 |
fixes S :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12839 |
assumes "affine S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12840 |
shows "bounded S \<longleftrightarrow> aff_dim S \<le> 0" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12841 |
apply safe |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12842 |
using affine_bounded_eq_trivial assms apply fastforce |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12843 |
by (metis aff_dim_sing aff_dim_subset affine_dim_equal affine_sing all_not_in_conv assms bounded_empty bounded_insert dual_order.antisym empty_subsetI insert_subset) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12844 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12845 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12846 |
lemma bounded_hyperplane_eq_trivial_0: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12847 |
fixes a :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12848 |
assumes "a \<noteq> 0" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12849 |
shows "bounded {x. a \<bullet> x = 0} \<longleftrightarrow> DIM('a) = 1"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12850 |
proof |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12851 |
assume "bounded {x. a \<bullet> x = 0}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12852 |
then have "aff_dim {x. a \<bullet> x = 0} \<le> 0"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12853 |
by (simp add: affine_bounded_eq_lowdim affine_hyperplane) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12854 |
with assms show "DIM('a) = 1"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12855 |
by (simp add: le_Suc_eq aff_dim_hyperplane) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12856 |
next |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12857 |
assume "DIM('a) = 1"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12858 |
then show "bounded {x. a \<bullet> x = 0}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12859 |
by (simp add: aff_dim_hyperplane affine_bounded_eq_lowdim affine_hyperplane assms) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12860 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12861 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12862 |
lemma bounded_hyperplane_eq_trivial: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12863 |
fixes a :: "'a::euclidean_space" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12864 |
shows "bounded {x. a \<bullet> x = r} \<longleftrightarrow> (if a = 0 then r \<noteq> 0 else DIM('a) = 1)"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12865 |
proof (simp add: bounded_hyperplane_eq_trivial_0, clarify) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12866 |
assume "r \<noteq> 0" "a \<noteq> 0" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12867 |
have "aff_dim {x. y \<bullet> x = 0} = aff_dim {x. a \<bullet> x = r}" if "y \<noteq> 0" for y::'a
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12868 |
by (metis that \<open>a \<noteq> 0\<close> aff_dim_hyperplane) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12869 |
then show "bounded {x. a \<bullet> x = r} = (DIM('a) = Suc 0)"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12870 |
by (metis One_nat_def \<open>a \<noteq> 0\<close> affine_bounded_eq_lowdim affine_hyperplane bounded_hyperplane_eq_trivial_0) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12871 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12872 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12873 |
subsection\<open>General case without assuming closure and getting non-strict separation\<close> |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12874 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12875 |
proposition separating_hyperplane_closed_point_inset: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12876 |
fixes S :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12877 |
assumes "convex S" "closed S" "S \<noteq> {}" "z \<notin> S"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12878 |
obtains a b where "a \<in> S" "(a - z) \<bullet> z < b" "\<And>x. x \<in> S \<Longrightarrow> b < (a - z) \<bullet> x" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12879 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12880 |
obtain y where "y \<in> S" and y: "\<And>u. u \<in> S \<Longrightarrow> dist z y \<le> dist z u" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12881 |
using distance_attains_inf [of S z] assms by auto |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12882 |
then have *: "(y - z) \<bullet> z < (y - z) \<bullet> z + (norm (y - z))\<^sup>2 / 2" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12883 |
using \<open>y \<in> S\<close> \<open>z \<notin> S\<close> by auto |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12884 |
show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12885 |
proof (rule that [OF \<open>y \<in> S\<close> *]) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12886 |
fix x |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12887 |
assume "x \<in> S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12888 |
have yz: "0 < (y - z) \<bullet> (y - z)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12889 |
using \<open>y \<in> S\<close> \<open>z \<notin> S\<close> by auto |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12890 |
{ assume 0: "0 < ((z - y) \<bullet> (x - y))"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12891 |
with any_closest_point_dot [OF \<open>convex S\<close> \<open>closed S\<close>] |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12892 |
have False |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12893 |
using y \<open>x \<in> S\<close> \<open>y \<in> S\<close> not_less by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12894 |
} |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12895 |
then have "0 \<le> ((y - z) \<bullet> (x - y))" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12896 |
by (force simp: not_less inner_diff_left) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12897 |
with yz have "0 < 2 * ((y - z) \<bullet> (x - y)) + (y - z) \<bullet> (y - z)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12898 |
by (simp add: algebra_simps) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12899 |
then show "(y - z) \<bullet> z + (norm (y - z))\<^sup>2 / 2 < (y - z) \<bullet> x" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12900 |
by (simp add: field_simps inner_diff_left inner_diff_right dot_square_norm [symmetric]) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12901 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12902 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12903 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12904 |
lemma separating_hyperplane_closed_0_inset: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12905 |
fixes S :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12906 |
assumes "convex S" "closed S" "S \<noteq> {}" "0 \<notin> S"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12907 |
obtains a b where "a \<in> S" "a \<noteq> 0" "0 < b" "\<And>x. x \<in> S \<Longrightarrow> a \<bullet> x > b" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12908 |
using separating_hyperplane_closed_point_inset [OF assms] |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12909 |
by simp (metis \<open>0 \<notin> S\<close>) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12910 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12911 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12912 |
proposition separating_hyperplane_set_0_inspan: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12913 |
fixes S :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12914 |
assumes "convex S" "S \<noteq> {}" "0 \<notin> S"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12915 |
obtains a where "a \<in> span S" "a \<noteq> 0" "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> a \<bullet> x" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12916 |
proof - |
| 63040 | 12917 |
define k where [abs_def]: "k c = {x. 0 \<le> c \<bullet> x}" for c :: 'a
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12918 |
have *: "span S \<inter> frontier (cball 0 1) \<inter> \<Inter>f' \<noteq> {}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12919 |
if f': "finite f'" "f' \<subseteq> k ` S" for f' |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12920 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12921 |
obtain C where "C \<subseteq> S" "finite C" and C: "f' = k ` C" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12922 |
using finite_subset_image [OF f'] by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12923 |
obtain a where "a \<in> S" "a \<noteq> 0" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12924 |
using \<open>S \<noteq> {}\<close> \<open>0 \<notin> S\<close> ex_in_conv by blast
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12925 |
then have "norm (a /\<^sub>R (norm a)) = 1" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12926 |
by simp |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12927 |
moreover have "a /\<^sub>R (norm a) \<in> span S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12928 |
by (simp add: \<open>a \<in> S\<close> span_mul span_superset) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12929 |
ultimately have ass: "a /\<^sub>R (norm a) \<in> span S \<inter> sphere 0 1" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12930 |
by simp |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12931 |
show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12932 |
proof (cases "C = {}")
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12933 |
case True with C ass show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12934 |
by auto |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12935 |
next |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12936 |
case False |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12937 |
have "closed (convex hull C)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12938 |
using \<open>finite C\<close> compact_eq_bounded_closed finite_imp_compact_convex_hull by auto |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12939 |
moreover have "convex hull C \<noteq> {}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12940 |
by (simp add: False) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12941 |
moreover have "0 \<notin> convex hull C" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12942 |
by (metis \<open>C \<subseteq> S\<close> \<open>convex S\<close> \<open>0 \<notin> S\<close> convex_hull_subset hull_same insert_absorb insert_subset) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12943 |
ultimately obtain a b |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12944 |
where "a \<in> convex hull C" "a \<noteq> 0" "0 < b" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12945 |
and ab: "\<And>x. x \<in> convex hull C \<Longrightarrow> a \<bullet> x > b" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12946 |
using separating_hyperplane_closed_0_inset by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12947 |
then have "a \<in> S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12948 |
by (metis \<open>C \<subseteq> S\<close> assms(1) subsetCE subset_hull) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12949 |
moreover have "norm (a /\<^sub>R (norm a)) = 1" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12950 |
using \<open>a \<noteq> 0\<close> by simp |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12951 |
moreover have "a /\<^sub>R (norm a) \<in> span S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12952 |
by (simp add: \<open>a \<in> S\<close> span_mul span_superset) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12953 |
ultimately have ass: "a /\<^sub>R (norm a) \<in> span S \<inter> sphere 0 1" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12954 |
by simp |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12955 |
have aa: "a /\<^sub>R (norm a) \<in> (\<Inter>c\<in>C. {x. 0 \<le> c \<bullet> x})"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12956 |
apply (clarsimp simp add: divide_simps) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12957 |
using ab \<open>0 < b\<close> |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12958 |
by (metis hull_inc inner_commute less_eq_real_def less_trans) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12959 |
show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12960 |
apply (simp add: C k_def) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12961 |
using ass aa Int_iff empty_iff by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12962 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12963 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12964 |
have "(span S \<inter> frontier(cball 0 1)) \<inter> (\<Inter> (k ` S)) \<noteq> {}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12965 |
apply (rule compact_imp_fip) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12966 |
apply (blast intro: compact_cball) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12967 |
using closed_halfspace_ge k_def apply blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12968 |
apply (metis *) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12969 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12970 |
then show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12971 |
unfolding set_eq_iff k_def |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12972 |
by simp (metis inner_commute norm_eq_zero that zero_neq_one) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12973 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12974 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12975 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12976 |
lemma separating_hyperplane_set_point_inaff: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12977 |
fixes S :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12978 |
assumes "convex S" "S \<noteq> {}" and zno: "z \<notin> S"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12979 |
obtains a b where "(z + a) \<in> affine hull (insert z S)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12980 |
and "a \<noteq> 0" and "a \<bullet> z \<le> b" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12981 |
and "\<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12982 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12983 |
from separating_hyperplane_set_0_inspan [of "image (\<lambda>x. -z + x) S"] |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12984 |
have "convex (op + (- z) ` S)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12985 |
by (simp add: \<open>convex S\<close>) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12986 |
moreover have "op + (- z) ` S \<noteq> {}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12987 |
by (simp add: \<open>S \<noteq> {}\<close>)
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12988 |
moreover have "0 \<notin> op + (- z) ` S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12989 |
using zno by auto |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12990 |
ultimately obtain a where "a \<in> span (op + (- z) ` S)" "a \<noteq> 0" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12991 |
and a: "\<And>x. x \<in> (op + (- z) ` S) \<Longrightarrow> 0 \<le> a \<bullet> x" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12992 |
using separating_hyperplane_set_0_inspan [of "image (\<lambda>x. -z + x) S"] |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12993 |
by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12994 |
then have szx: "\<And>x. x \<in> S \<Longrightarrow> a \<bullet> z \<le> a \<bullet> x" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12995 |
by (metis (no_types, lifting) imageI inner_minus_right inner_right_distrib minus_add neg_le_0_iff_le neg_le_iff_le real_add_le_0_iff) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12996 |
show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12997 |
apply (rule_tac a=a and b = "a \<bullet> z" in that, simp_all) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12998 |
using \<open>a \<in> span (op + (- z) ` S)\<close> affine_hull_insert_span_gen apply blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
12999 |
apply (simp_all add: \<open>a \<noteq> 0\<close> szx) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13000 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13001 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13002 |
|
| 63072 | 13003 |
proposition supporting_hyperplane_rel_boundary: |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13004 |
fixes S :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13005 |
assumes "convex S" "x \<in> S" and xno: "x \<notin> rel_interior S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13006 |
obtains a where "a \<noteq> 0" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13007 |
and "\<And>y. y \<in> S \<Longrightarrow> a \<bullet> x \<le> a \<bullet> y" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13008 |
and "\<And>y. y \<in> rel_interior S \<Longrightarrow> a \<bullet> x < a \<bullet> y" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13009 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13010 |
obtain a b where aff: "(x + a) \<in> affine hull (insert x (rel_interior S))" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13011 |
and "a \<noteq> 0" and "a \<bullet> x \<le> b" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13012 |
and ageb: "\<And>u. u \<in> (rel_interior S) \<Longrightarrow> a \<bullet> u \<ge> b" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13013 |
using separating_hyperplane_set_point_inaff [of "rel_interior S" x] assms |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13014 |
by (auto simp: rel_interior_eq_empty convex_rel_interior) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13015 |
have le_ay: "a \<bullet> x \<le> a \<bullet> y" if "y \<in> S" for y |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13016 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13017 |
have con: "continuous_on (closure (rel_interior S)) (op \<bullet> a)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13018 |
by (rule continuous_intros continuous_on_subset | blast)+ |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13019 |
have y: "y \<in> closure (rel_interior S)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13020 |
using \<open>convex S\<close> closure_def convex_closure_rel_interior \<open>y \<in> S\<close> |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13021 |
by fastforce |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13022 |
show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13023 |
using continuous_ge_on_closure [OF con y] ageb \<open>a \<bullet> x \<le> b\<close> |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13024 |
by fastforce |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13025 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13026 |
have 3: "a \<bullet> x < a \<bullet> y" if "y \<in> rel_interior S" for y |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13027 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13028 |
obtain e where "0 < e" "y \<in> S" and e: "cball y e \<inter> affine hull S \<subseteq> S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13029 |
using \<open>y \<in> rel_interior S\<close> by (force simp: rel_interior_cball) |
| 63040 | 13030 |
define y' where "y' = y - (e / norm a) *\<^sub>R ((x + a) - x)" |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13031 |
have "y' \<in> cball y e" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13032 |
unfolding y'_def using \<open>0 < e\<close> by force |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13033 |
moreover have "y' \<in> affine hull S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13034 |
unfolding y'_def |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13035 |
by (metis \<open>x \<in> S\<close> \<open>y \<in> S\<close> \<open>convex S\<close> aff affine_affine_hull hull_redundant |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13036 |
rel_interior_same_affine_hull hull_inc mem_affine_3_minus2) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13037 |
ultimately have "y' \<in> S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13038 |
using e by auto |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13039 |
have "a \<bullet> x \<le> a \<bullet> y" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13040 |
using le_ay \<open>a \<noteq> 0\<close> \<open>y \<in> S\<close> by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13041 |
moreover have "a \<bullet> x \<noteq> a \<bullet> y" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13042 |
using le_ay [OF \<open>y' \<in> S\<close>] \<open>a \<noteq> 0\<close> |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13043 |
apply (simp add: y'_def inner_diff dot_square_norm power2_eq_square) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13044 |
by (metis \<open>0 < e\<close> add_le_same_cancel1 inner_commute inner_real_def inner_zero_left le_diff_eq norm_le_zero_iff real_mult_le_cancel_iff2) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13045 |
ultimately show ?thesis by force |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13046 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13047 |
show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13048 |
by (rule that [OF \<open>a \<noteq> 0\<close> le_ay 3]) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13049 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13050 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13051 |
lemma supporting_hyperplane_relative_frontier: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13052 |
fixes S :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13053 |
assumes "convex S" "x \<in> closure S" "x \<notin> rel_interior S" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13054 |
obtains a where "a \<noteq> 0" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13055 |
and "\<And>y. y \<in> closure S \<Longrightarrow> a \<bullet> x \<le> a \<bullet> y" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13056 |
and "\<And>y. y \<in> rel_interior S \<Longrightarrow> a \<bullet> x < a \<bullet> y" |
| 63072 | 13057 |
using supporting_hyperplane_rel_boundary [of "closure S" x] |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13058 |
by (metis assms convex_closure convex_rel_interior_closure) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13059 |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13060 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13061 |
subsection\<open> Some results on decomposing convex hulls: intersections, simplicial subdivision\<close> |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13062 |
|
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13063 |
lemma |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13064 |
fixes s :: "'a::euclidean_space set" |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13065 |
assumes "~ (affine_dependent(s \<union> t))" |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13066 |
shows convex_hull_Int_subset: "convex hull s \<inter> convex hull t \<subseteq> convex hull (s \<inter> t)" (is ?C) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13067 |
and affine_hull_Int_subset: "affine hull s \<inter> affine hull t \<subseteq> affine hull (s \<inter> t)" (is ?A) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13068 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13069 |
have [simp]: "finite s" "finite t" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13070 |
using aff_independent_finite assms by blast+ |
| 64267 | 13071 |
have "sum u (s \<inter> t) = 1 \<and> |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13072 |
(\<Sum>v\<in>s \<inter> t. u v *\<^sub>R v) = (\<Sum>v\<in>s. u v *\<^sub>R v)" |
| 64267 | 13073 |
if [simp]: "sum u s = 1" |
13074 |
"sum v t = 1" |
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13075 |
and eq: "(\<Sum>x\<in>t. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)" for u v |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13076 |
proof - |
| 63040 | 13077 |
define f where "f x = (if x \<in> s then u x else 0) - (if x \<in> t then v x else 0)" for x |
| 64267 | 13078 |
have "sum f (s \<union> t) = 0" |
13079 |
apply (simp add: f_def sum_Un sum_subtractf) |
|
13080 |
apply (simp add: sum.inter_restrict [symmetric] Int_commute) |
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13081 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13082 |
moreover have "(\<Sum>x\<in>(s \<union> t). f x *\<^sub>R x) = 0" |
| 64267 | 13083 |
apply (simp add: f_def sum_Un scaleR_left_diff_distrib sum_subtractf) |
13084 |
apply (simp add: if_smult sum.inter_restrict [symmetric] Int_commute eq |
|
| 63566 | 13085 |
cong del: if_weak_cong) |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13086 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13087 |
ultimately have "\<And>v. v \<in> s \<union> t \<Longrightarrow> f v = 0" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13088 |
using aff_independent_finite assms unfolding affine_dependent_explicit |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13089 |
by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13090 |
then have u [simp]: "\<And>x. x \<in> s \<Longrightarrow> u x = (if x \<in> t then v x else 0)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13091 |
by (simp add: f_def) presburger |
| 64267 | 13092 |
have "sum u (s \<inter> t) = sum u s" |
13093 |
by (simp add: sum.inter_restrict) |
|
13094 |
then have "sum u (s \<inter> t) = 1" |
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13095 |
using that by linarith |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13096 |
moreover have "(\<Sum>v\<in>s \<inter> t. u v *\<^sub>R v) = (\<Sum>v\<in>s. u v *\<^sub>R v)" |
| 64267 | 13097 |
by (auto simp: if_smult sum.inter_restrict intro: sum.cong) |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13098 |
ultimately show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13099 |
by force |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13100 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13101 |
then show ?A ?C |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13102 |
by (auto simp: convex_hull_finite affine_hull_finite) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13103 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13104 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13105 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13106 |
proposition affine_hull_Int: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13107 |
fixes s :: "'a::euclidean_space set" |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13108 |
assumes "~ (affine_dependent(s \<union> t))" |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13109 |
shows "affine hull (s \<inter> t) = affine hull s \<inter> affine hull t" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13110 |
apply (rule subset_antisym) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13111 |
apply (simp add: hull_mono) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13112 |
by (simp add: affine_hull_Int_subset assms) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13113 |
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13114 |
proposition convex_hull_Int: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13115 |
fixes s :: "'a::euclidean_space set" |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13116 |
assumes "~ (affine_dependent(s \<union> t))" |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13117 |
shows "convex hull (s \<inter> t) = convex hull s \<inter> convex hull t" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13118 |
apply (rule subset_antisym) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13119 |
apply (simp add: hull_mono) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13120 |
by (simp add: convex_hull_Int_subset assms) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13121 |
|
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13122 |
proposition |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13123 |
fixes s :: "'a::euclidean_space set set" |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13124 |
assumes "~ (affine_dependent (\<Union>s))" |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13125 |
shows affine_hull_Inter: "affine hull (\<Inter>s) = (\<Inter>t\<in>s. affine hull t)" (is "?A") |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13126 |
and convex_hull_Inter: "convex hull (\<Inter>s) = (\<Inter>t\<in>s. convex hull t)" (is "?C") |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13127 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13128 |
have "finite s" |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13129 |
using aff_independent_finite assms finite_UnionD by blast |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13130 |
then have "?A \<and> ?C" using assms |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13131 |
proof (induction s rule: finite_induct) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13132 |
case empty then show ?case by auto |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13133 |
next |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13134 |
case (insert t F) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13135 |
then show ?case |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13136 |
proof (cases "F={}")
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13137 |
case True then show ?thesis by simp |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13138 |
next |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13139 |
case False |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13140 |
with "insert.prems" have [simp]: "\<not> affine_dependent (t \<union> \<Inter>F)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13141 |
by (auto intro: affine_dependent_subset) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13142 |
have [simp]: "\<not> affine_dependent (\<Union>F)" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13143 |
using affine_independent_subset insert.prems by fastforce |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13144 |
show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13145 |
by (simp add: affine_hull_Int convex_hull_Int insert.IH) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13146 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13147 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13148 |
then show "?A" "?C" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13149 |
by auto |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13150 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13151 |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13152 |
proposition in_convex_hull_exchange_unique: |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13153 |
fixes S :: "'a::euclidean_space set" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13154 |
assumes naff: "~ affine_dependent S" and a: "a \<in> convex hull S" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13155 |
and S: "T \<subseteq> S" "T' \<subseteq> S" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13156 |
and x: "x \<in> convex hull (insert a T)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13157 |
and x': "x \<in> convex hull (insert a T')" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13158 |
shows "x \<in> convex hull (insert a (T \<inter> T'))" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13159 |
proof (cases "a \<in> S") |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13160 |
case True |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13161 |
then have "\<not> affine_dependent (insert a T \<union> insert a T')" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13162 |
using affine_dependent_subset assms by auto |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13163 |
then have "x \<in> convex hull (insert a T \<inter> insert a T')" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13164 |
by (metis IntI convex_hull_Int x x') |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13165 |
then show ?thesis |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13166 |
by simp |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13167 |
next |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13168 |
case False |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13169 |
then have anot: "a \<notin> T" "a \<notin> T'" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13170 |
using assms by auto |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13171 |
have [simp]: "finite S" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13172 |
by (simp add: aff_independent_finite assms) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13173 |
then obtain b where b0: "\<And>s. s \<in> S \<Longrightarrow> 0 \<le> b s" |
| 64267 | 13174 |
and b1: "sum b S = 1" and aeq: "a = (\<Sum>s\<in>S. b s *\<^sub>R s)" |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13175 |
using a by (auto simp: convex_hull_finite) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13176 |
have fin [simp]: "finite T" "finite T'" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13177 |
using assms infinite_super \<open>finite S\<close> by blast+ |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13178 |
then obtain c c' where c0: "\<And>t. t \<in> insert a T \<Longrightarrow> 0 \<le> c t" |
| 64267 | 13179 |
and c1: "sum c (insert a T) = 1" |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13180 |
and xeq: "x = (\<Sum>t \<in> insert a T. c t *\<^sub>R t)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13181 |
and c'0: "\<And>t. t \<in> insert a T' \<Longrightarrow> 0 \<le> c' t" |
| 64267 | 13182 |
and c'1: "sum c' (insert a T') = 1" |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13183 |
and x'eq: "x = (\<Sum>t \<in> insert a T'. c' t *\<^sub>R t)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13184 |
using x x' by (auto simp: convex_hull_finite) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13185 |
with fin anot |
| 64267 | 13186 |
have sumTT': "sum c T = 1 - c a" "sum c' T' = 1 - c' a" |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13187 |
and wsumT: "(\<Sum>t \<in> T. c t *\<^sub>R t) = x - c a *\<^sub>R a" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13188 |
by simp_all |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13189 |
have wsumT': "(\<Sum>t \<in> T'. c' t *\<^sub>R t) = x - c' a *\<^sub>R a" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13190 |
using x'eq fin anot by simp |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13191 |
define cc where "cc \<equiv> \<lambda>x. if x \<in> T then c x else 0" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13192 |
define cc' where "cc' \<equiv> \<lambda>x. if x \<in> T' then c' x else 0" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13193 |
define dd where "dd \<equiv> \<lambda>x. cc x - cc' x + (c a - c' a) * b x" |
| 64267 | 13194 |
have sumSS': "sum cc S = 1 - c a" "sum cc' S = 1 - c' a" |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13195 |
unfolding cc_def cc'_def using S |
| 64267 | 13196 |
by (simp_all add: Int_absorb1 Int_absorb2 sum_subtractf sum.inter_restrict [symmetric] sumTT') |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13197 |
have wsumSS: "(\<Sum>t \<in> S. cc t *\<^sub>R t) = x - c a *\<^sub>R a" "(\<Sum>t \<in> S. cc' t *\<^sub>R t) = x - c' a *\<^sub>R a" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13198 |
unfolding cc_def cc'_def using S |
| 64267 | 13199 |
by (simp_all add: Int_absorb1 Int_absorb2 if_smult sum.inter_restrict [symmetric] wsumT wsumT' cong: if_cong) |
13200 |
have sum_dd0: "sum dd S = 0" |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13201 |
unfolding dd_def using S |
| 64267 | 13202 |
by (simp add: sumSS' comm_monoid_add_class.sum.distrib sum_subtractf |
13203 |
algebra_simps sum_distrib_right [symmetric] b1) |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13204 |
have "(\<Sum>v\<in>S. (b v * x) *\<^sub>R v) = x *\<^sub>R (\<Sum>v\<in>S. b v *\<^sub>R v)" for x |
| 64267 | 13205 |
by (simp add: pth_5 real_vector.scale_sum_right mult.commute) |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13206 |
then have *: "(\<Sum>v\<in>S. (b v * x) *\<^sub>R v) = x *\<^sub>R a" for x |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13207 |
using aeq by blast |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13208 |
have "(\<Sum>v \<in> S. dd v *\<^sub>R v) = 0" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13209 |
unfolding dd_def using S |
| 64267 | 13210 |
by (simp add: * wsumSS sum.distrib sum_subtractf algebra_simps) |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13211 |
then have dd0: "dd v = 0" if "v \<in> S" for v |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13212 |
using naff that \<open>finite S\<close> sum_dd0 unfolding affine_dependent_explicit |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13213 |
apply (simp only: not_ex) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13214 |
apply (drule_tac x=S in spec) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13215 |
apply (drule_tac x=dd in spec, simp) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13216 |
done |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13217 |
consider "c' a \<le> c a" | "c a \<le> c' a" by linarith |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13218 |
then show ?thesis |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13219 |
proof cases |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13220 |
case 1 |
| 64267 | 13221 |
then have "sum cc S \<le> sum cc' S" |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13222 |
by (simp add: sumSS') |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13223 |
then have le: "cc x \<le> cc' x" if "x \<in> S" for x |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13224 |
using dd0 [OF that] 1 b0 mult_left_mono that |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13225 |
by (fastforce simp add: dd_def algebra_simps) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13226 |
have cc0: "cc x = 0" if "x \<in> S" "x \<notin> T \<inter> T'" for x |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13227 |
using le [OF \<open>x \<in> S\<close>] that c0 |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13228 |
by (force simp: cc_def cc'_def split: if_split_asm) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13229 |
show ?thesis |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13230 |
proof (simp add: convex_hull_finite, intro exI conjI) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13231 |
show "\<forall>x\<in>T \<inter> T'. 0 \<le> (cc(a := c a)) x" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13232 |
by (simp add: c0 cc_def) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13233 |
show "0 \<le> (cc(a := c a)) a" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13234 |
by (simp add: c0) |
| 64267 | 13235 |
have "sum (cc(a := c a)) (insert a (T \<inter> T')) = c a + sum (cc(a := c a)) (T \<inter> T')" |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13236 |
by (simp add: anot) |
| 64267 | 13237 |
also have "... = c a + sum (cc(a := c a)) S" |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13238 |
apply simp |
| 64267 | 13239 |
apply (rule sum.mono_neutral_left) |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13240 |
using \<open>T \<subseteq> S\<close> apply (auto simp: \<open>a \<notin> S\<close> cc0) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13241 |
done |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13242 |
also have "... = c a + (1 - c a)" |
| 64267 | 13243 |
by (metis \<open>a \<notin> S\<close> fun_upd_other sum.cong sumSS') |
13244 |
finally show "sum (cc(a := c a)) (insert a (T \<inter> T')) = 1" |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13245 |
by simp |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13246 |
have "(\<Sum>x\<in>insert a (T \<inter> T'). (cc(a := c a)) x *\<^sub>R x) = c a *\<^sub>R a + (\<Sum>x \<in> T \<inter> T'. (cc(a := c a)) x *\<^sub>R x)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13247 |
by (simp add: anot) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13248 |
also have "... = c a *\<^sub>R a + (\<Sum>x \<in> S. (cc(a := c a)) x *\<^sub>R x)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13249 |
apply simp |
| 64267 | 13250 |
apply (rule sum.mono_neutral_left) |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13251 |
using \<open>T \<subseteq> S\<close> apply (auto simp: \<open>a \<notin> S\<close> cc0) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13252 |
done |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13253 |
also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" |
| 64267 | 13254 |
by (simp add: wsumSS \<open>a \<notin> S\<close> if_smult sum_delta_notmem) |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13255 |
finally show "(\<Sum>x\<in>insert a (T \<inter> T'). (cc(a := c a)) x *\<^sub>R x) = x" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13256 |
by simp |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13257 |
qed |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13258 |
next |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13259 |
case 2 |
| 64267 | 13260 |
then have "sum cc' S \<le> sum cc S" |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13261 |
by (simp add: sumSS') |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13262 |
then have le: "cc' x \<le> cc x" if "x \<in> S" for x |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13263 |
using dd0 [OF that] 2 b0 mult_left_mono that |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13264 |
by (fastforce simp add: dd_def algebra_simps) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13265 |
have cc0: "cc' x = 0" if "x \<in> S" "x \<notin> T \<inter> T'" for x |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13266 |
using le [OF \<open>x \<in> S\<close>] that c'0 |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13267 |
by (force simp: cc_def cc'_def split: if_split_asm) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13268 |
show ?thesis |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13269 |
proof (simp add: convex_hull_finite, intro exI conjI) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13270 |
show "\<forall>x\<in>T \<inter> T'. 0 \<le> (cc'(a := c' a)) x" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13271 |
by (simp add: c'0 cc'_def) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13272 |
show "0 \<le> (cc'(a := c' a)) a" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13273 |
by (simp add: c'0) |
| 64267 | 13274 |
have "sum (cc'(a := c' a)) (insert a (T \<inter> T')) = c' a + sum (cc'(a := c' a)) (T \<inter> T')" |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13275 |
by (simp add: anot) |
| 64267 | 13276 |
also have "... = c' a + sum (cc'(a := c' a)) S" |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13277 |
apply simp |
| 64267 | 13278 |
apply (rule sum.mono_neutral_left) |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13279 |
using \<open>T \<subseteq> S\<close> apply (auto simp: \<open>a \<notin> S\<close> cc0) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13280 |
done |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13281 |
also have "... = c' a + (1 - c' a)" |
| 64267 | 13282 |
by (metis \<open>a \<notin> S\<close> fun_upd_other sum.cong sumSS') |
13283 |
finally show "sum (cc'(a := c' a)) (insert a (T \<inter> T')) = 1" |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13284 |
by simp |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13285 |
have "(\<Sum>x\<in>insert a (T \<inter> T'). (cc'(a := c' a)) x *\<^sub>R x) = c' a *\<^sub>R a + (\<Sum>x \<in> T \<inter> T'. (cc'(a := c' a)) x *\<^sub>R x)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13286 |
by (simp add: anot) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13287 |
also have "... = c' a *\<^sub>R a + (\<Sum>x \<in> S. (cc'(a := c' a)) x *\<^sub>R x)" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13288 |
apply simp |
| 64267 | 13289 |
apply (rule sum.mono_neutral_left) |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13290 |
using \<open>T \<subseteq> S\<close> apply (auto simp: \<open>a \<notin> S\<close> cc0) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13291 |
done |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13292 |
also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" |
| 64267 | 13293 |
by (simp add: wsumSS \<open>a \<notin> S\<close> if_smult sum_delta_notmem) |
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13294 |
finally show "(\<Sum>x\<in>insert a (T \<inter> T'). (cc'(a := c' a)) x *\<^sub>R x) = x" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13295 |
by simp |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13296 |
qed |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13297 |
qed |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13298 |
qed |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13299 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13300 |
corollary convex_hull_exchange_Int: |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13301 |
fixes a :: "'a::euclidean_space" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13302 |
assumes "~ affine_dependent S" "a \<in> convex hull S" "T \<subseteq> S" "T' \<subseteq> S" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13303 |
shows "(convex hull (insert a T)) \<inter> (convex hull (insert a T')) = |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13304 |
convex hull (insert a (T \<inter> T'))" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13305 |
apply (rule subset_antisym) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13306 |
using in_convex_hull_exchange_unique assms apply blast |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13307 |
by (metis hull_mono inf_le1 inf_le2 insert_inter_insert le_inf_iff) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
13308 |
|
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13309 |
lemma Int_closed_segment: |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13310 |
fixes b :: "'a::euclidean_space" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13311 |
assumes "b \<in> closed_segment a c \<or> ~collinear{a,b,c}"
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13312 |
shows "closed_segment a b \<inter> closed_segment b c = {b}"
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13313 |
proof (cases "c = a") |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13314 |
case True |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13315 |
then show ?thesis |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13316 |
using assms collinear_3_eq_affine_dependent by fastforce |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13317 |
next |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13318 |
case False |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13319 |
from assms show ?thesis |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13320 |
proof |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13321 |
assume "b \<in> closed_segment a c" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13322 |
moreover have "\<not> affine_dependent {a, c}"
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13323 |
by (simp add: affine_independent_2) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13324 |
ultimately show ?thesis |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13325 |
using False convex_hull_exchange_Int [of "{a,c}" b "{a}" "{c}"]
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13326 |
by (simp add: segment_convex_hull insert_commute) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13327 |
next |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13328 |
assume ncoll: "\<not> collinear {a, b, c}"
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13329 |
have False if "closed_segment a b \<inter> closed_segment b c \<noteq> {b}"
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13330 |
proof - |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13331 |
have "b \<in> closed_segment a b" and "b \<in> closed_segment b c" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13332 |
by auto |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13333 |
with that obtain d where "b \<noteq> d" "d \<in> closed_segment a b" "d \<in> closed_segment b c" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13334 |
by force |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13335 |
then have d: "collinear {a, d, b}" "collinear {b, d, c}"
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13336 |
by (auto simp: between_mem_segment between_imp_collinear) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13337 |
have "collinear {a, b, c}"
|
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13338 |
apply (rule collinear_3_trans [OF _ _ \<open>b \<noteq> d\<close>]) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13339 |
using d by (auto simp: insert_commute) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13340 |
with ncoll show False .. |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13341 |
qed |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13342 |
then show ?thesis |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13343 |
by blast |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13344 |
qed |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13345 |
qed |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
13346 |
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13347 |
lemma affine_hull_finite_intersection_hyperplanes: |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13348 |
fixes s :: "'a::euclidean_space set" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13349 |
obtains f where |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13350 |
"finite f" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13351 |
"of_nat (card f) + aff_dim s = DIM('a)"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13352 |
"affine hull s = \<Inter>f" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13353 |
"\<And>h. h \<in> f \<Longrightarrow> \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x = b}"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13354 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13355 |
obtain b where "b \<subseteq> s" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13356 |
and indb: "\<not> affine_dependent b" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13357 |
and eq: "affine hull s = affine hull b" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13358 |
using affine_basis_exists by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13359 |
obtain c where indc: "\<not> affine_dependent c" and "b \<subseteq> c" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13360 |
and affc: "affine hull c = UNIV" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13361 |
by (metis extend_to_affine_basis affine_UNIV hull_same indb subset_UNIV) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13362 |
then have "finite c" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13363 |
by (simp add: aff_independent_finite) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13364 |
then have fbc: "finite b" "card b \<le> card c" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13365 |
using \<open>b \<subseteq> c\<close> infinite_super by (auto simp: card_mono) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13366 |
have imeq: "(\<lambda>x. affine hull x) ` ((\<lambda>a. c - {a}) ` (c - b)) = ((\<lambda>a. affine hull (c - {a})) ` (c - b))"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13367 |
by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13368 |
have card1: "card ((\<lambda>a. affine hull (c - {a})) ` (c - b)) = card (c - b)"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13369 |
apply (rule card_image [OF inj_onI]) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13370 |
by (metis Diff_eq_empty_iff Diff_iff indc affine_dependent_def hull_subset insert_iff) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13371 |
have card2: "(card (c - b)) + aff_dim s = DIM('a)"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13372 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13373 |
have aff: "aff_dim (UNIV::'a set) = aff_dim c" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13374 |
by (metis aff_dim_affine_hull affc) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13375 |
have "aff_dim b = aff_dim s" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13376 |
by (metis (no_types) aff_dim_affine_hull eq) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13377 |
then have "int (card b) = 1 + aff_dim s" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13378 |
by (simp add: aff_dim_affine_independent indb) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13379 |
then show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13380 |
using fbc aff |
| 63072 | 13381 |
by (simp add: \<open>\<not> affine_dependent c\<close> \<open>b \<subseteq> c\<close> aff_dim_affine_independent aff_dim_UNIV card_Diff_subset of_nat_diff) |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13382 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13383 |
show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13384 |
proof (cases "c = b") |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13385 |
case True show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13386 |
apply (rule_tac f="{}" in that)
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13387 |
using True affc |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13388 |
apply (simp_all add: eq [symmetric]) |
| 63072 | 13389 |
by (metis aff_dim_UNIV aff_dim_affine_hull) |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13390 |
next |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13391 |
case False |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13392 |
have ind: "\<not> affine_dependent (\<Union>a\<in>c - b. c - {a})"
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13393 |
by (rule affine_independent_subset [OF indc]) auto |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13394 |
have affeq: "affine hull s = (\<Inter>x\<in>(\<lambda>a. c - {a}) ` (c - b). affine hull x)"
|
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13395 |
using \<open>b \<subseteq> c\<close> False |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13396 |
apply (subst affine_hull_Inter [OF ind, symmetric]) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13397 |
apply (simp add: eq double_diff) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13398 |
done |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13399 |
have *: "1 + aff_dim (c - {t}) = int (DIM('a))"
|
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13400 |
if t: "t \<in> c" for t |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13401 |
proof - |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13402 |
have "insert t c = c" |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13403 |
using t by blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13404 |
then show ?thesis |
| 63072 | 13405 |
by (metis (full_types) add.commute aff_dim_affine_hull aff_dim_insert aff_dim_UNIV affc affine_dependent_def indc insert_Diff_single t) |
|
63016
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13406 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13407 |
show ?thesis |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13408 |
apply (rule_tac f = "(\<lambda>x. affine hull x) ` ((\<lambda>a. c - {a}) ` (c - b))" in that)
|
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13409 |
using \<open>finite c\<close> apply blast |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13410 |
apply (simp add: imeq card1 card2) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13411 |
apply (simp add: affeq, clarify) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13412 |
apply (metis DIM_positive One_nat_def Suc_leI add_diff_cancel_left' of_nat_1 aff_dim_eq_hyperplane of_nat_diff *) |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13413 |
done |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13414 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13415 |
qed |
|
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents:
63007
diff
changeset
|
13416 |
|
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13417 |
subsection\<open>Misc results about span\<close> |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13418 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13419 |
lemma eq_span_insert_eq: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13420 |
assumes "(x - y) \<in> span S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13421 |
shows "span(insert x S) = span(insert y S)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13422 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13423 |
have *: "span(insert x S) \<subseteq> span(insert y S)" if "(x - y) \<in> span S" for x y |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13424 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13425 |
have 1: "(r *\<^sub>R x - r *\<^sub>R y) \<in> span S" for r |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13426 |
by (metis real_vector.scale_right_diff_distrib span_mul that) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13427 |
have 2: "(z - k *\<^sub>R y) - k *\<^sub>R (x - y) = z - k *\<^sub>R x" for z k |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13428 |
by (simp add: real_vector.scale_right_diff_distrib) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13429 |
show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13430 |
apply (clarsimp simp add: span_breakdown_eq) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13431 |
by (metis 1 2 diff_add_cancel real_vector.scale_right_diff_distrib span_add_eq) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13432 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13433 |
show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13434 |
apply (intro subset_antisym * assms) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13435 |
using assms subspace_neg subspace_span minus_diff_eq by force |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13436 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13437 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13438 |
lemma dim_psubset: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13439 |
fixes S :: "'a :: euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13440 |
shows "span S \<subset> span T \<Longrightarrow> dim S < dim T" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13441 |
by (metis (no_types, hide_lams) dim_span less_le not_le subspace_dim_equal subspace_span) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13442 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13443 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13444 |
lemma basis_subspace_exists: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13445 |
fixes S :: "'a::euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13446 |
shows |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13447 |
"subspace S |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13448 |
\<Longrightarrow> \<exists>b. finite b \<and> b \<subseteq> S \<and> |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13449 |
independent b \<and> span b = S \<and> card b = dim S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13450 |
by (metis span_subspace basis_exists independent_imp_finite) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13451 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13452 |
lemma affine_hyperplane_sums_eq_UNIV_0: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13453 |
fixes S :: "'a :: euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13454 |
assumes "affine S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13455 |
and "0 \<in> S" and "w \<in> S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13456 |
and "a \<bullet> w \<noteq> 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13457 |
shows "{x + y| x y. x \<in> S \<and> a \<bullet> y = 0} = UNIV"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13458 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13459 |
have "subspace S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13460 |
by (simp add: assms subspace_affine) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13461 |
have span1: "span {y. a \<bullet> y = 0} \<subseteq> span {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13462 |
apply (rule span_mono) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13463 |
using \<open>0 \<in> S\<close> add.left_neutral by force |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13464 |
have "w \<notin> span {y. a \<bullet> y = 0}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13465 |
using \<open>a \<bullet> w \<noteq> 0\<close> span_induct subspace_hyperplane by auto |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13466 |
moreover have "w \<in> span {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13467 |
using \<open>w \<in> S\<close> |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13468 |
by (metis (mono_tags, lifting) inner_zero_right mem_Collect_eq pth_d span_superset) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13469 |
ultimately have span2: "span {y. a \<bullet> y = 0} \<noteq> span {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13470 |
by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13471 |
have "a \<noteq> 0" using assms inner_zero_left by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13472 |
then have "DIM('a) - 1 = dim {y. a \<bullet> y = 0}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13473 |
by (simp add: dim_hyperplane) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13474 |
also have "... < dim {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13475 |
using span1 span2 by (blast intro: dim_psubset) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13476 |
finally have DIM_lt: "DIM('a) - 1 < dim {x + y |x y. x \<in> S \<and> a \<bullet> y = 0}" .
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13477 |
have subs: "subspace {x + y| x y. x \<in> S \<and> a \<bullet> y = 0}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13478 |
using subspace_sums [OF \<open>subspace S\<close> subspace_hyperplane] by simp |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13479 |
moreover have "span {x + y| x y. x \<in> S \<and> a \<bullet> y = 0} = UNIV"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13480 |
apply (rule dim_eq_full [THEN iffD1]) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13481 |
apply (rule antisym [OF dim_subset_UNIV]) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13482 |
using DIM_lt apply simp |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13483 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13484 |
ultimately show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13485 |
by (simp add: subs) (metis (lifting) span_eq subs) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13486 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13487 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13488 |
proposition affine_hyperplane_sums_eq_UNIV: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13489 |
fixes S :: "'a :: euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13490 |
assumes "affine S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13491 |
and "S \<inter> {v. a \<bullet> v = b} \<noteq> {}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13492 |
and "S - {v. a \<bullet> v = b} \<noteq> {}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13493 |
shows "{x + y| x y. x \<in> S \<and> a \<bullet> y = b} = UNIV"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13494 |
proof (cases "a = 0") |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13495 |
case True with assms show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13496 |
by (auto simp: if_splits) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13497 |
next |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13498 |
case False |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13499 |
obtain c where "c \<in> S" and c: "a \<bullet> c = b" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13500 |
using assms by force |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13501 |
with affine_diffs_subspace [OF \<open>affine S\<close>] |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13502 |
have "subspace (op + (- c) ` S)" by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13503 |
then have aff: "affine (op + (- c) ` S)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13504 |
by (simp add: subspace_imp_affine) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13505 |
have 0: "0 \<in> op + (- c) ` S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13506 |
by (simp add: \<open>c \<in> S\<close>) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13507 |
obtain d where "d \<in> S" and "a \<bullet> d \<noteq> b" and dc: "d-c \<in> op + (- c) ` S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13508 |
using assms by auto |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13509 |
then have adc: "a \<bullet> (d - c) \<noteq> 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13510 |
by (simp add: c inner_diff_right) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13511 |
let ?U = "op + (c+c) ` {x + y |x y. x \<in> op + (- c) ` S \<and> a \<bullet> y = 0}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13512 |
have "u + v \<in> op + (c + c) ` {x + v |x v. x \<in> op + (- c) ` S \<and> a \<bullet> v = 0}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13513 |
if "u \<in> S" "b = a \<bullet> v" for u v |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13514 |
apply (rule_tac x="u+v-c-c" in image_eqI) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13515 |
apply (simp_all add: algebra_simps) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13516 |
apply (rule_tac x="u-c" in exI) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13517 |
apply (rule_tac x="v-c" in exI) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13518 |
apply (simp add: algebra_simps that c) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13519 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13520 |
moreover have "\<lbrakk>a \<bullet> v = 0; u \<in> S\<rbrakk> |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13521 |
\<Longrightarrow> \<exists>x ya. v + (u + c) = x + ya \<and> x \<in> S \<and> a \<bullet> ya = b" for v u |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13522 |
by (metis add.left_commute c inner_right_distrib pth_d) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13523 |
ultimately have "{x + y |x y. x \<in> S \<and> a \<bullet> y = b} = ?U"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13524 |
by (fastforce simp: algebra_simps) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13525 |
also have "... = op + (c+c) ` UNIV" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13526 |
by (simp add: affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc]) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13527 |
also have "... = UNIV" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13528 |
by (simp add: translation_UNIV) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13529 |
finally show ?thesis . |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13530 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13531 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13532 |
proposition dim_sums_Int: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13533 |
fixes S :: "'a :: euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13534 |
assumes "subspace S" "subspace T" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13535 |
shows "dim {x + y |x y. x \<in> S \<and> y \<in> T} + dim(S \<inter> T) = dim S + dim T"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13536 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13537 |
obtain B where B: "B \<subseteq> S \<inter> T" "S \<inter> T \<subseteq> span B" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13538 |
and indB: "independent B" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13539 |
and cardB: "card B = dim (S \<inter> T)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13540 |
using basis_exists by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13541 |
then obtain C D where "B \<subseteq> C" "C \<subseteq> S" "independent C" "S \<subseteq> span C" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13542 |
and "B \<subseteq> D" "D \<subseteq> T" "independent D" "T \<subseteq> span D" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13543 |
using maximal_independent_subset_extend |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13544 |
by (metis Int_subset_iff \<open>B \<subseteq> S \<inter> T\<close> indB) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13545 |
then have "finite B" "finite C" "finite D" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13546 |
by (simp_all add: independent_imp_finite indB independent_bound) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13547 |
have Beq: "B = C \<inter> D" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13548 |
apply (rule sym) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13549 |
apply (rule spanning_subset_independent) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13550 |
using \<open>B \<subseteq> C\<close> \<open>B \<subseteq> D\<close> apply blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13551 |
apply (meson \<open>independent C\<close> independent_mono inf.cobounded1) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13552 |
using B \<open>C \<subseteq> S\<close> \<open>D \<subseteq> T\<close> apply auto |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13553 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13554 |
then have Deq: "D = B \<union> (D - C)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13555 |
by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13556 |
have CUD: "C \<union> D \<subseteq> {x + y |x y. x \<in> S \<and> y \<in> T}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13557 |
apply safe |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13558 |
apply (metis add.right_neutral subsetCE \<open>C \<subseteq> S\<close> \<open>subspace T\<close> set_eq_subset span_0 span_minimal) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13559 |
apply (metis add.left_neutral subsetCE \<open>D \<subseteq> T\<close> \<open>subspace S\<close> set_eq_subset span_0 span_minimal) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13560 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13561 |
have "a v = 0" if 0: "(\<Sum>v\<in>C. a v *\<^sub>R v) + (\<Sum>v\<in>D - C. a v *\<^sub>R v) = 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13562 |
and v: "v \<in> C \<union> (D-C)" for a v |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13563 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13564 |
have eq: "(\<Sum>v\<in>D - C. a v *\<^sub>R v) = - (\<Sum>v\<in>C. a v *\<^sub>R v)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13565 |
using that add_eq_0_iff by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13566 |
have "(\<Sum>v\<in>D - C. a v *\<^sub>R v) \<in> S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13567 |
apply (subst eq) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13568 |
apply (rule subspace_neg [OF \<open>subspace S\<close>]) |
| 64267 | 13569 |
apply (rule subspace_sum [OF \<open>subspace S\<close>]) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13570 |
by (meson subsetCE subspace_mul \<open>C \<subseteq> S\<close> \<open>subspace S\<close>) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13571 |
moreover have "(\<Sum>v\<in>D - C. a v *\<^sub>R v) \<in> T" |
| 64267 | 13572 |
apply (rule subspace_sum [OF \<open>subspace T\<close>]) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13573 |
by (meson DiffD1 \<open>D \<subseteq> T\<close> \<open>subspace T\<close> subset_eq subspace_def) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13574 |
ultimately have "(\<Sum>v \<in> D-C. a v *\<^sub>R v) \<in> span B" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13575 |
using B by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13576 |
then obtain e where e: "(\<Sum>v\<in>B. e v *\<^sub>R v) = (\<Sum>v \<in> D-C. a v *\<^sub>R v)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13577 |
using span_finite [OF \<open>finite B\<close>] by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13578 |
have "\<And>c v. \<lbrakk>(\<Sum>v\<in>C. c v *\<^sub>R v) = 0; v \<in> C\<rbrakk> \<Longrightarrow> c v = 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13579 |
using independent_explicit \<open>independent C\<close> by blast |
| 63148 | 13580 |
define cc where "cc x = (if x \<in> B then a x + e x else a x)" for x |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13581 |
have [simp]: "C \<inter> B = B" "D \<inter> B = B" "C \<inter> - B = C-D" "B \<inter> (D - C) = {}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13582 |
using \<open>B \<subseteq> C\<close> \<open>B \<subseteq> D\<close> Beq by blast+ |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13583 |
have f2: "(\<Sum>v\<in>C \<inter> D. e v *\<^sub>R v) = (\<Sum>v\<in>D - C. a v *\<^sub>R v)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13584 |
using Beq e by presburger |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13585 |
have f3: "(\<Sum>v\<in>C \<union> D. a v *\<^sub>R v) = (\<Sum>v\<in>C - D. a v *\<^sub>R v) + (\<Sum>v\<in>D - C. a v *\<^sub>R v) + (\<Sum>v\<in>C \<inter> D. a v *\<^sub>R v)" |
| 64267 | 13586 |
using \<open>finite C\<close> \<open>finite D\<close> sum.union_diff2 by blast |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13587 |
have f4: "(\<Sum>v\<in>C \<union> (D - C). a v *\<^sub>R v) = (\<Sum>v\<in>C. a v *\<^sub>R v) + (\<Sum>v\<in>D - C. a v *\<^sub>R v)" |
| 64267 | 13588 |
by (meson Diff_disjoint \<open>finite C\<close> \<open>finite D\<close> finite_Diff sum.union_disjoint) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13589 |
have "(\<Sum>v\<in>C. cc v *\<^sub>R v) = 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13590 |
using 0 f2 f3 f4 |
| 64267 | 13591 |
apply (simp add: cc_def Beq if_smult \<open>finite C\<close> sum.If_cases algebra_simps sum.distrib) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13592 |
apply (simp add: add.commute add.left_commute diff_eq) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13593 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13594 |
then have "\<And>v. v \<in> C \<Longrightarrow> cc v = 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13595 |
using independent_explicit \<open>independent C\<close> by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13596 |
then have C0: "\<And>v. v \<in> C - B \<Longrightarrow> a v = 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13597 |
by (simp add: cc_def Beq) meson |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13598 |
then have [simp]: "(\<Sum>x\<in>C - B. a x *\<^sub>R x) = 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13599 |
by simp |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13600 |
have "(\<Sum>x\<in>C. a x *\<^sub>R x) = (\<Sum>x\<in>B. a x *\<^sub>R x)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13601 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13602 |
have "C - D = C - B" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13603 |
using Beq by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13604 |
then show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13605 |
using Beq \<open>(\<Sum>x\<in>C - B. a x *\<^sub>R x) = 0\<close> f3 f4 by auto |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13606 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13607 |
with 0 have Dcc0: "(\<Sum>v\<in>D. a v *\<^sub>R v) = 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13608 |
apply (subst Deq) |
| 64267 | 13609 |
by (simp add: \<open>finite B\<close> \<open>finite D\<close> sum_Un) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13610 |
then have D0: "\<And>v. v \<in> D \<Longrightarrow> a v = 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13611 |
using independent_explicit \<open>independent D\<close> by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13612 |
show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13613 |
using v C0 D0 Beq by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13614 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13615 |
then have "independent (C \<union> (D - C))" |
| 64267 | 13616 |
by (simp add: independent_explicit \<open>finite C\<close> \<open>finite D\<close> sum_Un del: Un_Diff_cancel) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13617 |
then have indCUD: "independent (C \<union> D)" by simp |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13618 |
have "dim (S \<inter> T) = card B" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13619 |
by (rule dim_unique [OF B indB refl]) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13620 |
moreover have "dim S = card C" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13621 |
by (metis \<open>C \<subseteq> S\<close> \<open>independent C\<close> \<open>S \<subseteq> span C\<close> basis_card_eq_dim) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13622 |
moreover have "dim T = card D" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13623 |
by (metis \<open>D \<subseteq> T\<close> \<open>independent D\<close> \<open>T \<subseteq> span D\<close> basis_card_eq_dim) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13624 |
moreover have "dim {x + y |x y. x \<in> S \<and> y \<in> T} = card(C \<union> D)"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13625 |
apply (rule dim_unique [OF CUD _ indCUD refl], clarify) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13626 |
apply (meson \<open>S \<subseteq> span C\<close> \<open>T \<subseteq> span D\<close> span_add span_inc span_minimal subsetCE subspace_span sup.bounded_iff) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13627 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13628 |
ultimately show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13629 |
using \<open>B = C \<inter> D\<close> [symmetric] |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13630 |
by (simp add: \<open>independent C\<close> \<open>independent D\<close> card_Un_Int independent_finite) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13631 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13632 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13633 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13634 |
lemma aff_dim_sums_Int_0: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13635 |
assumes "affine S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13636 |
and "affine T" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13637 |
and "0 \<in> S" "0 \<in> T" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13638 |
shows "aff_dim {x + y| x y. x \<in> S \<and> y \<in> T} = (aff_dim S + aff_dim T) - aff_dim(S \<inter> T)"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13639 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13640 |
have "0 \<in> {x + y |x y. x \<in> S \<and> y \<in> T}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13641 |
using assms by force |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13642 |
then have 0: "0 \<in> affine hull {x + y |x y. x \<in> S \<and> y \<in> T}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13643 |
by (metis (lifting) hull_inc) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13644 |
have sub: "subspace S" "subspace T" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13645 |
using assms by (auto simp: subspace_affine) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13646 |
show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13647 |
using dim_sums_Int [OF sub] by (simp add: aff_dim_zero assms 0 hull_inc) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13648 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13649 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13650 |
proposition aff_dim_sums_Int: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13651 |
assumes "affine S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13652 |
and "affine T" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13653 |
and "S \<inter> T \<noteq> {}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13654 |
shows "aff_dim {x + y| x y. x \<in> S \<and> y \<in> T} = (aff_dim S + aff_dim T) - aff_dim(S \<inter> T)"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13655 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13656 |
obtain a where a: "a \<in> S" "a \<in> T" using assms by force |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13657 |
have aff: "affine (op+ (-a) ` S)" "affine (op+ (-a) ` T)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13658 |
using assms by (auto simp: affine_translation [symmetric]) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13659 |
have zero: "0 \<in> (op+ (-a) ` S)" "0 \<in> (op+ (-a) ` T)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13660 |
using a assms by auto |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13661 |
have [simp]: "{x + y |x y. x \<in> op + (- a) ` S \<and> y \<in> op + (- a) ` T} =
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13662 |
op + (- 2 *\<^sub>R a) ` {x + y| x y. x \<in> S \<and> y \<in> T}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13663 |
by (force simp: algebra_simps scaleR_2) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13664 |
have [simp]: "op + (- a) ` S \<inter> op + (- a) ` T = op + (- a) ` (S \<inter> T)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13665 |
by auto |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13666 |
show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13667 |
using aff_dim_sums_Int_0 [OF aff zero] |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13668 |
by (auto simp: aff_dim_translation_eq) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13669 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13670 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13671 |
lemma aff_dim_affine_Int_hyperplane: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13672 |
fixes a :: "'a::euclidean_space" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13673 |
assumes "affine S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13674 |
shows "aff_dim(S \<inter> {x. a \<bullet> x = b}) =
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13675 |
(if S \<inter> {v. a \<bullet> v = b} = {} then - 1
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13676 |
else if S \<subseteq> {v. a \<bullet> v = b} then aff_dim S
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13677 |
else aff_dim S - 1)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13678 |
proof (cases "a = 0") |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13679 |
case True with assms show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13680 |
by auto |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13681 |
next |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13682 |
case False |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13683 |
then have "aff_dim (S \<inter> {x. a \<bullet> x = b}) = aff_dim S - 1"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13684 |
if "x \<in> S" "a \<bullet> x \<noteq> b" and non: "S \<inter> {v. a \<bullet> v = b} \<noteq> {}" for x
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13685 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13686 |
have [simp]: "{x + y| x y. x \<in> S \<and> a \<bullet> y = b} = UNIV"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13687 |
using affine_hyperplane_sums_eq_UNIV [OF assms non] that by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13688 |
show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13689 |
using aff_dim_sums_Int [OF assms affine_hyperplane non] |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13690 |
by (simp add: of_nat_diff False) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13691 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13692 |
then show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13693 |
by (metis (mono_tags, lifting) inf.orderE aff_dim_empty_eq mem_Collect_eq subsetI) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13694 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13695 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13696 |
lemma aff_dim_lt_full: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13697 |
fixes S :: "'a::euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13698 |
shows "aff_dim S < DIM('a) \<longleftrightarrow> (affine hull S \<noteq> UNIV)"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13699 |
by (metis (no_types) aff_dim_affine_hull aff_dim_le_DIM aff_dim_UNIV affine_hull_UNIV less_le) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13700 |
|
| 64122 | 13701 |
|
13702 |
lemma dim_Times: |
|
13703 |
fixes S :: "'a :: euclidean_space set" and T :: "'a set" |
|
13704 |
assumes "subspace S" "subspace T" |
|
13705 |
shows "dim(S \<times> T) = dim S + dim T" |
|
13706 |
proof - |
|
13707 |
have ss: "subspace ((\<lambda>x. (x, 0)) ` S)" "subspace (Pair 0 ` T)" |
|
13708 |
by (rule subspace_linear_image, unfold_locales, auto simp: assms)+ |
|
13709 |
have "dim(S \<times> T) = dim({u + v |u v. u \<in> (\<lambda>x. (x, 0)) ` S \<and> v \<in> Pair 0 ` T})"
|
|
13710 |
by (simp add: Times_eq_image_sum) |
|
13711 |
moreover have "dim ((\<lambda>x. (x, 0::'a)) ` S) = dim S" "dim (Pair (0::'a) ` T) = dim T" |
|
13712 |
by (auto simp: additive.intro linear.intro linear_axioms.intro inj_on_def intro: dim_image_eq) |
|
13713 |
moreover have "dim ((\<lambda>x. (x, 0)) ` S \<inter> Pair 0 ` T) = 0" |
|
13714 |
by (subst dim_eq_0) (force simp: zero_prod_def) |
|
13715 |
ultimately show ?thesis |
|
13716 |
using dim_sums_Int [OF ss] by linarith |
|
13717 |
qed |
|
13718 |
||
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13719 |
subsection\<open> Orthogonal bases, Gram-Schmidt process, and related theorems\<close> |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13720 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13721 |
lemma span_delete_0 [simp]: "span(S - {0}) = span S"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13722 |
proof |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13723 |
show "span (S - {0}) \<subseteq> span S"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13724 |
by (blast intro!: span_mono) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13725 |
next |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13726 |
have "span S \<subseteq> span(insert 0 (S - {0}))"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13727 |
by (blast intro!: span_mono) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13728 |
also have "... \<subseteq> span(S - {0})"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13729 |
using span_insert_0 by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13730 |
finally show "span S \<subseteq> span (S - {0})" .
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13731 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13732 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13733 |
lemma span_image_scale: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13734 |
assumes "finite S" and nz: "\<And>x. x \<in> S \<Longrightarrow> c x \<noteq> 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13735 |
shows "span ((\<lambda>x. c x *\<^sub>R x) ` S) = span S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13736 |
using assms |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13737 |
proof (induction S arbitrary: c) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13738 |
case (empty c) show ?case by simp |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13739 |
next |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13740 |
case (insert x F c) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13741 |
show ?case |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13742 |
proof (intro set_eqI iffI) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13743 |
fix y |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13744 |
assume "y \<in> span ((\<lambda>x. c x *\<^sub>R x) ` insert x F)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13745 |
then show "y \<in> span (insert x F)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13746 |
using insert by (force simp: span_breakdown_eq) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13747 |
next |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13748 |
fix y |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13749 |
assume "y \<in> span (insert x F)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13750 |
then show "y \<in> span ((\<lambda>x. c x *\<^sub>R x) ` insert x F)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13751 |
using insert |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13752 |
apply (clarsimp simp: span_breakdown_eq) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13753 |
apply (rule_tac x="k / c x" in exI) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13754 |
by simp |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13755 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13756 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13757 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13758 |
lemma pairwise_orthogonal_independent: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13759 |
assumes "pairwise orthogonal S" and "0 \<notin> S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13760 |
shows "independent S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13761 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13762 |
have 0: "\<And>x y. \<lbrakk>x \<noteq> y; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13763 |
using assms by (simp add: pairwise_def orthogonal_def) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13764 |
have "False" if "a \<in> S" and a: "a \<in> span (S - {a})" for a
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13765 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13766 |
obtain T U where "T \<subseteq> S - {a}" "a = (\<Sum>v\<in>T. U v *\<^sub>R v)"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13767 |
using a by (force simp: span_explicit) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13768 |
then have "a \<bullet> a = a \<bullet> (\<Sum>v\<in>T. U v *\<^sub>R v)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13769 |
by simp |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13770 |
also have "... = 0" |
| 64267 | 13771 |
apply (simp add: inner_sum_right) |
13772 |
apply (rule comm_monoid_add_class.sum.neutral) |
|
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13773 |
by (metis "0" DiffE \<open>T \<subseteq> S - {a}\<close> mult_not_zero singletonI subsetCE \<open>a \<in> S\<close>)
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13774 |
finally show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13775 |
using \<open>0 \<notin> S\<close> \<open>a \<in> S\<close> by auto |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13776 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13777 |
then show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13778 |
by (force simp: dependent_def) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13779 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13780 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13781 |
lemma pairwise_orthogonal_imp_finite: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13782 |
fixes S :: "'a::euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13783 |
assumes "pairwise orthogonal S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13784 |
shows "finite S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13785 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13786 |
have "independent (S - {0})"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13787 |
apply (rule pairwise_orthogonal_independent) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13788 |
apply (metis Diff_iff assms pairwise_def) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13789 |
by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13790 |
then show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13791 |
by (meson independent_imp_finite infinite_remove) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13792 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13793 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13794 |
lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13795 |
by (simp add: subspace_def orthogonal_clauses) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13796 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13797 |
lemma subspace_orthogonal_to_vectors: "subspace {y. \<forall>x \<in> S. orthogonal x y}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13798 |
by (simp add: subspace_def orthogonal_clauses) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13799 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13800 |
lemma orthogonal_to_span: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13801 |
assumes a: "a \<in> span S" and x: "\<And>y. y \<in> S \<Longrightarrow> orthogonal x y" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13802 |
shows "orthogonal x a" |
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
13803 |
apply (rule span_induct [OF a subspace_orthogonal_to_vector]) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13804 |
apply (simp add: x) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13805 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13806 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13807 |
proposition Gram_Schmidt_step: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13808 |
fixes S :: "'a::euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13809 |
assumes S: "pairwise orthogonal S" and x: "x \<in> span S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13810 |
shows "orthogonal x (a - (\<Sum>b\<in>S. (b \<bullet> a / (b \<bullet> b)) *\<^sub>R b))" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13811 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13812 |
have "finite S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13813 |
by (simp add: S pairwise_orthogonal_imp_finite) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13814 |
have "orthogonal (a - (\<Sum>b\<in>S. (b \<bullet> a / (b \<bullet> b)) *\<^sub>R b)) x" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13815 |
if "x \<in> S" for x |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13816 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13817 |
have "a \<bullet> x = (\<Sum>y\<in>S. if y = x then y \<bullet> a else 0)" |
| 64267 | 13818 |
by (simp add: \<open>finite S\<close> inner_commute sum.delta that) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13819 |
also have "... = (\<Sum>b\<in>S. b \<bullet> a * (b \<bullet> x) / (b \<bullet> b))" |
| 64267 | 13820 |
apply (rule sum.cong [OF refl], simp) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13821 |
by (meson S orthogonal_def pairwise_def that) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13822 |
finally show ?thesis |
| 64267 | 13823 |
by (simp add: orthogonal_def algebra_simps inner_sum_left) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13824 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13825 |
then show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13826 |
using orthogonal_to_span orthogonal_commute x by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13827 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13828 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13829 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13830 |
lemma orthogonal_extension_aux: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13831 |
fixes S :: "'a::euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13832 |
assumes "finite T" "finite S" "pairwise orthogonal S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13833 |
shows "\<exists>U. pairwise orthogonal (S \<union> U) \<and> span (S \<union> U) = span (S \<union> T)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13834 |
using assms |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13835 |
proof (induction arbitrary: S) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13836 |
case empty then show ?case |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13837 |
by simp (metis sup_bot_right) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13838 |
next |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13839 |
case (insert a T) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13840 |
have 0: "\<And>x y. \<lbrakk>x \<noteq> y; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13841 |
using insert by (simp add: pairwise_def orthogonal_def) |
| 63148 | 13842 |
define a' where "a' = a - (\<Sum>b\<in>S. (b \<bullet> a / (b \<bullet> b)) *\<^sub>R b)" |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13843 |
obtain U where orthU: "pairwise orthogonal (S \<union> insert a' U)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13844 |
and spanU: "span (insert a' S \<union> U) = span (insert a' S \<union> T)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13845 |
apply (rule exE [OF insert.IH [of "insert a' S"]]) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13846 |
apply (auto simp: Gram_Schmidt_step a'_def insert.prems orthogonal_commute pairwise_orthogonal_insert span_clauses) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13847 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13848 |
have orthS: "\<And>x. x \<in> S \<Longrightarrow> a' \<bullet> x = 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13849 |
apply (simp add: a'_def) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13850 |
using Gram_Schmidt_step [OF \<open>pairwise orthogonal S\<close>] |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13851 |
apply (force simp: orthogonal_def inner_commute span_inc [THEN subsetD]) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13852 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13853 |
have "span (S \<union> insert a' U) = span (insert a' (S \<union> T))" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13854 |
using spanU by simp |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13855 |
also have "... = span (insert a (S \<union> T))" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13856 |
apply (rule eq_span_insert_eq) |
| 64267 | 13857 |
apply (simp add: a'_def span_neg span_sum span_clauses(1) span_mul) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13858 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13859 |
also have "... = span (S \<union> insert a T)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13860 |
by simp |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13861 |
finally show ?case |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13862 |
apply (rule_tac x="insert a' U" in exI) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13863 |
using orthU apply auto |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13864 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13865 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13866 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13867 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13868 |
proposition orthogonal_extension: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13869 |
fixes S :: "'a::euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13870 |
assumes S: "pairwise orthogonal S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13871 |
obtains U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> T)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13872 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13873 |
obtain B where "finite B" "span B = span T" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13874 |
using basis_subspace_exists [of "span T"] subspace_span by auto |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13875 |
with orthogonal_extension_aux [of B S] |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13876 |
obtain U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> B)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13877 |
using assms pairwise_orthogonal_imp_finite by auto |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13878 |
show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13879 |
apply (rule_tac U=U in that) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13880 |
apply (simp add: \<open>pairwise orthogonal (S \<union> U)\<close>) |
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
13881 |
by (metis \<open>span (S \<union> U) = span (S \<union> B)\<close> \<open>span B = span T\<close> span_Un) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13882 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13883 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13884 |
corollary orthogonal_extension_strong: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13885 |
fixes S :: "'a::euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13886 |
assumes S: "pairwise orthogonal S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13887 |
obtains U where "U \<inter> (insert 0 S) = {}" "pairwise orthogonal (S \<union> U)"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13888 |
"span (S \<union> U) = span (S \<union> T)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13889 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13890 |
obtain U where "pairwise orthogonal (S \<union> U)" "span (S \<union> U) = span (S \<union> T)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13891 |
using orthogonal_extension assms by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13892 |
then show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13893 |
apply (rule_tac U = "U - (insert 0 S)" in that) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13894 |
apply blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13895 |
apply (force simp: pairwise_def) |
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
13896 |
apply (metis (no_types, lifting) Un_Diff_cancel span_insert_0 span_Un) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13897 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13898 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13899 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13900 |
subsection\<open>Decomposing a vector into parts in orthogonal subspaces.\<close> |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13901 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13902 |
text\<open>existence of orthonormal basis for a subspace.\<close> |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13903 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13904 |
lemma orthogonal_spanningset_subspace: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13905 |
fixes S :: "'a :: euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13906 |
assumes "subspace S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13907 |
obtains B where "B \<subseteq> S" "pairwise orthogonal B" "span B = S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13908 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13909 |
obtain B where "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13910 |
using basis_exists by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13911 |
with orthogonal_extension [of "{}" B]
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13912 |
show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13913 |
by (metis Un_empty_left assms pairwise_empty span_inc span_subspace that) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13914 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13915 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13916 |
lemma orthogonal_basis_subspace: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13917 |
fixes S :: "'a :: euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13918 |
assumes "subspace S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13919 |
obtains B where "0 \<notin> B" "B \<subseteq> S" "pairwise orthogonal B" "independent B" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13920 |
"card B = dim S" "span B = S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13921 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13922 |
obtain B where "B \<subseteq> S" "pairwise orthogonal B" "span B = S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13923 |
using assms orthogonal_spanningset_subspace by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13924 |
then show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13925 |
apply (rule_tac B = "B - {0}" in that)
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13926 |
apply (auto simp: indep_card_eq_dim_span pairwise_subset Diff_subset pairwise_orthogonal_independent elim: pairwise_subset) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13927 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13928 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13929 |
|
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13930 |
proposition orthonormal_basis_subspace: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13931 |
fixes S :: "'a :: euclidean_space set" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13932 |
assumes "subspace S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13933 |
obtains B where "B \<subseteq> S" "pairwise orthogonal B" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13934 |
and "\<And>x. x \<in> B \<Longrightarrow> norm x = 1" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13935 |
and "independent B" "card B = dim S" "span B = S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13936 |
proof - |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13937 |
obtain B where "0 \<notin> B" "B \<subseteq> S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13938 |
and orth: "pairwise orthogonal B" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13939 |
and "independent B" "card B = dim S" "span B = S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13940 |
by (blast intro: orthogonal_basis_subspace [OF assms]) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13941 |
have 1: "(\<lambda>x. x /\<^sub>R norm x) ` B \<subseteq> S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13942 |
using \<open>span B = S\<close> span_clauses(1) span_mul by fastforce |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13943 |
have 2: "pairwise orthogonal ((\<lambda>x. x /\<^sub>R norm x) ` B)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13944 |
using orth by (force simp: pairwise_def orthogonal_clauses) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13945 |
have 3: "\<And>x. x \<in> (\<lambda>x. x /\<^sub>R norm x) ` B \<Longrightarrow> norm x = 1" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13946 |
by (metis (no_types, lifting) \<open>0 \<notin> B\<close> image_iff norm_sgn sgn_div_norm) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13947 |
have 4: "independent ((\<lambda>x. x /\<^sub>R norm x) ` B)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13948 |
by (metis "2" "3" norm_zero pairwise_orthogonal_independent zero_neq_one) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13949 |
have "inj_on (\<lambda>x. x /\<^sub>R norm x) B" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13950 |
proof |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13951 |
fix x y |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13952 |
assume "x \<in> B" "y \<in> B" "x /\<^sub>R norm x = y /\<^sub>R norm y" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13953 |
moreover have "\<And>i. i \<in> B \<Longrightarrow> norm (i /\<^sub>R norm i) = 1" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13954 |
using 3 by blast |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13955 |
ultimately show "x = y" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13956 |
by (metis norm_eq_1 orth orthogonal_clauses(7) orthogonal_commute orthogonal_def pairwise_def zero_neq_one) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13957 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13958 |
then have 5: "card ((\<lambda>x. x /\<^sub>R norm x) ` B) = dim S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13959 |
by (metis \<open>card B = dim S\<close> card_image) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13960 |
have 6: "span ((\<lambda>x. x /\<^sub>R norm x) ` B) = S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13961 |
by (metis "1" "4" "5" assms card_eq_dim independent_finite span_subspace) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13962 |
show ?thesis |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13963 |
by (rule that [OF 1 2 3 4 5 6]) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13964 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
13965 |
|
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13966 |
proposition orthogonal_subspace_decomp_exists: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13967 |
fixes S :: "'a :: euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13968 |
obtains y z where "y \<in> span S" "\<And>w. w \<in> span S \<Longrightarrow> orthogonal z w" "x = y + z" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13969 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13970 |
obtain T where "0 \<notin> T" "T \<subseteq> span S" "pairwise orthogonal T" "independent T" "card T = dim (span S)" "span T = span S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13971 |
using orthogonal_basis_subspace subspace_span by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13972 |
let ?a = "\<Sum>b\<in>T. (b \<bullet> x / (b \<bullet> b)) *\<^sub>R b" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13973 |
have orth: "orthogonal (x - ?a) w" if "w \<in> span S" for w |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13974 |
by (simp add: Gram_Schmidt_step \<open>pairwise orthogonal T\<close> \<open>span T = span S\<close> orthogonal_commute that) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13975 |
show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13976 |
apply (rule_tac y = "?a" and z = "x - ?a" in that) |
| 64267 | 13977 |
apply (meson \<open>T \<subseteq> span S\<close> span_mul span_sum subsetCE) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13978 |
apply (fact orth, simp) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13979 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13980 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13981 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13982 |
lemma orthogonal_subspace_decomp_unique: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13983 |
fixes S :: "'a :: euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13984 |
assumes "x + y = x' + y'" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13985 |
and ST: "x \<in> span S" "x' \<in> span S" "y \<in> span T" "y' \<in> span T" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13986 |
and orth: "\<And>a b. \<lbrakk>a \<in> S; b \<in> T\<rbrakk> \<Longrightarrow> orthogonal a b" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13987 |
shows "x = x' \<and> y = y'" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13988 |
proof - |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13989 |
have "x + y - y' = x'" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13990 |
by (simp add: assms) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13991 |
moreover have "\<And>a b. \<lbrakk>a \<in> span S; b \<in> span T\<rbrakk> \<Longrightarrow> orthogonal a b" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13992 |
by (meson orth orthogonal_commute orthogonal_to_span) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13993 |
ultimately have "0 = x' - x" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13994 |
by (metis (full_types) add_diff_cancel_left' ST diff_right_commute orthogonal_clauses(10) orthogonal_clauses(5) orthogonal_self) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13995 |
with assms show ?thesis by auto |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13996 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
13997 |
|
|
63492
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
13998 |
proposition dim_orthogonal_sum: |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
13999 |
fixes A :: "'a::euclidean_space set" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14000 |
assumes "\<And>x y. \<lbrakk>x \<in> A; y \<in> B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14001 |
shows "dim(A \<union> B) = dim A + dim B" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14002 |
proof - |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14003 |
have 1: "\<And>x y. \<lbrakk>x \<in> span A; y \<in> B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14004 |
by (erule span_induct [OF _ subspace_hyperplane2]; simp add: assms) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14005 |
have "\<And>x y. \<lbrakk>x \<in> span A; y \<in> span B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14006 |
apply (erule span_induct [OF _ subspace_hyperplane]) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14007 |
using 1 by (simp add: ) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14008 |
then have 0: "\<And>x y. \<lbrakk>x \<in> span A; y \<in> span B\<rbrakk> \<Longrightarrow> x \<bullet> y = 0" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14009 |
by simp |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14010 |
have "dim(A \<union> B) = dim (span (A \<union> B))" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14011 |
by simp |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14012 |
also have "... = dim ((\<lambda>(a, b). a + b) ` (span A \<times> span B))" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14013 |
by (simp add: span_Un) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14014 |
also have "... = dim {x + y |x y. x \<in> span A \<and> y \<in> span B}"
|
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14015 |
by (auto intro!: arg_cong [where f=dim]) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14016 |
also have "... = dim {x + y |x y. x \<in> span A \<and> y \<in> span B} + dim(span A \<inter> span B)"
|
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14017 |
by (auto simp: dest: 0) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14018 |
also have "... = dim (span A) + dim (span B)" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14019 |
by (rule dim_sums_Int) auto |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14020 |
also have "... = dim A + dim B" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14021 |
by simp |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14022 |
finally show ?thesis . |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14023 |
qed |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14024 |
|
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14025 |
lemma dim_subspace_orthogonal_to_vectors: |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14026 |
fixes A :: "'a::euclidean_space set" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14027 |
assumes "subspace A" "subspace B" "A \<subseteq> B" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14028 |
shows "dim {y \<in> B. \<forall>x \<in> A. orthogonal x y} + dim A = dim B"
|
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14029 |
proof - |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14030 |
have "dim (span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A)) = dim (span B)"
|
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14031 |
proof (rule arg_cong [where f=dim, OF subset_antisym]) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14032 |
show "span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A) \<subseteq> span B"
|
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14033 |
by (simp add: \<open>A \<subseteq> B\<close> Collect_restrict span_mono) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14034 |
next |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14035 |
have *: "x \<in> span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A)"
|
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14036 |
if "x \<in> B" for x |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14037 |
proof - |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14038 |
obtain y z where "x = y + z" "y \<in> span A" and orth: "\<And>w. w \<in> span A \<Longrightarrow> orthogonal z w" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14039 |
using orthogonal_subspace_decomp_exists [of A x] that by auto |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14040 |
have "y \<in> span B" |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14041 |
by (metis span_eq \<open>y \<in> span A\<close> assms subset_iff) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14042 |
then have "z \<in> {a \<in> B. \<forall>x. x \<in> A \<longrightarrow> orthogonal x a}"
|
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14043 |
by simp (metis (no_types) span_eq \<open>x = y + z\<close> \<open>subspace A\<close> \<open>subspace B\<close> orth orthogonal_commute span_add_eq that) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14044 |
then have z: "z \<in> span {y \<in> B. \<forall>x\<in>A. orthogonal x y}"
|
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14045 |
by (meson span_inc subset_iff) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14046 |
then show ?thesis |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14047 |
apply (simp add: span_Un image_def) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14048 |
apply (rule bexI [OF _ z]) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14049 |
apply (simp add: \<open>x = y + z\<close> \<open>y \<in> span A\<close>) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14050 |
done |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14051 |
qed |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14052 |
show "span B \<subseteq> span ({y \<in> B. \<forall>x\<in>A. orthogonal x y} \<union> A)"
|
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14053 |
by (rule span_minimal) (auto intro: * span_minimal elim: ) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14054 |
qed |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14055 |
then show ?thesis |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14056 |
by (metis (no_types, lifting) dim_orthogonal_sum dim_span mem_Collect_eq orthogonal_commute orthogonal_def) |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14057 |
qed |
|
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
14058 |
|
| 64122 | 14059 |
lemma aff_dim_openin: |
14060 |
fixes S :: "'a::euclidean_space set" |
|
14061 |
assumes ope: "openin (subtopology euclidean T) S" and "affine T" "S \<noteq> {}"
|
|
14062 |
shows "aff_dim S = aff_dim T" |
|
14063 |
proof - |
|
14064 |
show ?thesis |
|
14065 |
proof (rule order_antisym) |
|
14066 |
show "aff_dim S \<le> aff_dim T" |
|
14067 |
by (blast intro: aff_dim_subset [OF openin_imp_subset] ope) |
|
14068 |
next |
|
14069 |
obtain a where "a \<in> S" |
|
14070 |
using \<open>S \<noteq> {}\<close> by blast
|
|
14071 |
have "S \<subseteq> T" |
|
14072 |
using ope openin_imp_subset by auto |
|
14073 |
then have "a \<in> T" |
|
14074 |
using \<open>a \<in> S\<close> by auto |
|
14075 |
then have subT': "subspace ((\<lambda>x. - a + x) ` T)" |
|
14076 |
using affine_diffs_subspace \<open>affine T\<close> by auto |
|
14077 |
then obtain B where Bsub: "B \<subseteq> ((\<lambda>x. - a + x) ` T)" and po: "pairwise orthogonal B" |
|
14078 |
and eq1: "\<And>x. x \<in> B \<Longrightarrow> norm x = 1" and "independent B" |
|
14079 |
and cardB: "card B = dim ((\<lambda>x. - a + x) ` T)" |
|
14080 |
and spanB: "span B = ((\<lambda>x. - a + x) ` T)" |
|
14081 |
by (rule orthonormal_basis_subspace) auto |
|
14082 |
obtain e where "0 < e" and e: "cball a e \<inter> T \<subseteq> S" |
|
14083 |
by (meson \<open>a \<in> S\<close> openin_contains_cball ope) |
|
14084 |
have "aff_dim T = aff_dim ((\<lambda>x. - a + x) ` T)" |
|
14085 |
by (metis aff_dim_translation_eq) |
|
14086 |
also have "... = dim ((\<lambda>x. - a + x) ` T)" |
|
14087 |
using aff_dim_subspace subT' by blast |
|
14088 |
also have "... = card B" |
|
14089 |
by (simp add: cardB) |
|
14090 |
also have "... = card ((\<lambda>x. e *\<^sub>R x) ` B)" |
|
14091 |
using \<open>0 < e\<close> by (force simp: inj_on_def card_image) |
|
14092 |
also have "... \<le> dim ((\<lambda>x. - a + x) ` S)" |
|
14093 |
proof (simp, rule independent_card_le_dim) |
|
14094 |
have e': "cball 0 e \<inter> (\<lambda>x. x - a) ` T \<subseteq> (\<lambda>x. x - a) ` S" |
|
14095 |
using e by (auto simp: dist_norm norm_minus_commute subset_eq) |
|
14096 |
have "(\<lambda>x. e *\<^sub>R x) ` B \<subseteq> cball 0 e \<inter> (\<lambda>x. x - a) ` T" |
|
14097 |
using Bsub \<open>0 < e\<close> eq1 subT' \<open>a \<in> T\<close> by (auto simp: subspace_def) |
|
14098 |
then show "(\<lambda>x. e *\<^sub>R x) ` B \<subseteq> (\<lambda>x. x - a) ` S" |
|
14099 |
using e' by blast |
|
14100 |
show "independent ((\<lambda>x. e *\<^sub>R x) ` B)" |
|
14101 |
using \<open>independent B\<close> |
|
14102 |
apply (rule independent_injective_image, simp) |
|
14103 |
by (metis \<open>0 < e\<close> injective_scaleR less_irrefl) |
|
14104 |
qed |
|
14105 |
also have "... = aff_dim S" |
|
14106 |
using \<open>a \<in> S\<close> aff_dim_eq_dim hull_inc by force |
|
14107 |
finally show "aff_dim T \<le> aff_dim S" . |
|
14108 |
qed |
|
14109 |
qed |
|
14110 |
||
14111 |
lemma dim_openin: |
|
14112 |
fixes S :: "'a::euclidean_space set" |
|
14113 |
assumes ope: "openin (subtopology euclidean T) S" and "subspace T" "S \<noteq> {}"
|
|
14114 |
shows "dim S = dim T" |
|
14115 |
proof (rule order_antisym) |
|
14116 |
show "dim S \<le> dim T" |
|
14117 |
by (metis ope dim_subset openin_subset topspace_euclidean_subtopology) |
|
14118 |
next |
|
14119 |
have "dim T = aff_dim S" |
|
14120 |
using aff_dim_openin |
|
14121 |
by (metis aff_dim_subspace \<open>subspace T\<close> \<open>S \<noteq> {}\<close> ope subspace_affine)
|
|
14122 |
also have "... \<le> dim S" |
|
14123 |
by (metis aff_dim_subset aff_dim_subspace dim_span span_inc subspace_span) |
|
14124 |
finally show "dim T \<le> dim S" by simp |
|
14125 |
qed |
|
14126 |
||
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14127 |
subsection\<open>Parallel slices, etc.\<close> |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14128 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14129 |
text\<open> If we take a slice out of a set, we can do it perpendicularly, |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14130 |
with the normal vector to the slice parallel to the affine hull.\<close> |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14131 |
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14132 |
proposition affine_parallel_slice: |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14133 |
fixes S :: "'a :: euclidean_space set" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14134 |
assumes "affine S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14135 |
and "S \<inter> {x. a \<bullet> x \<le> b} \<noteq> {}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14136 |
and "~ (S \<subseteq> {x. a \<bullet> x \<le> b})"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14137 |
obtains a' b' where "a' \<noteq> 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14138 |
"S \<inter> {x. a' \<bullet> x \<le> b'} = S \<inter> {x. a \<bullet> x \<le> b}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14139 |
"S \<inter> {x. a' \<bullet> x = b'} = S \<inter> {x. a \<bullet> x = b}"
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14140 |
"\<And>w. w \<in> S \<Longrightarrow> (w + a') \<in> S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14141 |
proof (cases "S \<inter> {x. a \<bullet> x = b} = {}")
|
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14142 |
case True |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14143 |
then obtain u v where "u \<in> S" "v \<in> S" "a \<bullet> u \<le> b" "a \<bullet> v > b" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14144 |
using assms by (auto simp: not_le) |
| 63148 | 14145 |
define \<eta> where "\<eta> = u + ((b - a \<bullet> u) / (a \<bullet> v - a \<bullet> u)) *\<^sub>R (v - u)" |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14146 |
have "\<eta> \<in> S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14147 |
by (simp add: \<eta>_def \<open>u \<in> S\<close> \<open>v \<in> S\<close> \<open>affine S\<close> mem_affine_3_minus) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14148 |
moreover have "a \<bullet> \<eta> = b" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14149 |
using \<open>a \<bullet> u \<le> b\<close> \<open>b < a \<bullet> v\<close> |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14150 |
by (simp add: \<eta>_def algebra_simps) (simp add: field_simps) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14151 |
ultimately have False |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14152 |
using True by force |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14153 |
then show ?thesis .. |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14154 |
next |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14155 |
case False |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14156 |
then obtain z where "z \<in> S" and z: "a \<bullet> z = b" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14157 |
using assms by auto |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14158 |
with affine_diffs_subspace [OF \<open>affine S\<close>] |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14159 |
have sub: "subspace (op + (- z) ` S)" by blast |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14160 |
then have aff: "affine (op + (- z) ` S)" and span: "span (op + (- z) ` S) = (op + (- z) ` S)" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14161 |
by (auto simp: subspace_imp_affine) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14162 |
obtain a' a'' where a': "a' \<in> span (op + (- z) ` S)" and a: "a = a' + a''" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14163 |
and "\<And>w. w \<in> span (op + (- z) ` S) \<Longrightarrow> orthogonal a'' w" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14164 |
using orthogonal_subspace_decomp_exists [of "op + (- z) ` S" "a"] by metis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14165 |
then have "\<And>w. w \<in> S \<Longrightarrow> a'' \<bullet> (w-z) = 0" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14166 |
by (simp add: imageI orthogonal_def span) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14167 |
then have a'': "\<And>w. w \<in> S \<Longrightarrow> a'' \<bullet> w = (a - a') \<bullet> z" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14168 |
by (simp add: a inner_diff_right) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14169 |
then have ba'': "\<And>w. w \<in> S \<Longrightarrow> a'' \<bullet> w = b - a' \<bullet> z" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14170 |
by (simp add: inner_diff_left z) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14171 |
have "\<And>w. w \<in> op + (- z) ` S \<Longrightarrow> (w + a') \<in> op + (- z) ` S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14172 |
by (metis subspace_add a' span_eq sub) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14173 |
then have Sclo: "\<And>w. w \<in> S \<Longrightarrow> (w + a') \<in> S" |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14174 |
by fastforce |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14175 |
show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14176 |
proof (cases "a' = 0") |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14177 |
case True |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14178 |
with a assms True a'' diff_zero less_irrefl show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14179 |
by auto |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14180 |
next |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14181 |
case False |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14182 |
show ?thesis |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14183 |
apply (rule_tac a' = "a'" and b' = "a' \<bullet> z" in that) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14184 |
apply (auto simp: a ba'' inner_left_distrib False Sclo) |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14185 |
done |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14186 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14187 |
qed |
|
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63072
diff
changeset
|
14188 |
|
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14189 |
lemma diffs_affine_hull_span: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14190 |
assumes "a \<in> S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14191 |
shows "{x - a |x. x \<in> affine hull S} = span {x - a |x. x \<in> S}"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14192 |
proof - |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14193 |
have *: "((\<lambda>x. x - a) ` (S - {a})) = {x. x + a \<in> S} - {0}"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14194 |
by (auto simp: algebra_simps) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14195 |
show ?thesis |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14196 |
apply (simp add: affine_hull_span2 [OF assms] *) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14197 |
apply (auto simp: algebra_simps) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14198 |
done |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14199 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14200 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14201 |
lemma aff_dim_dim_affine_diffs: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14202 |
fixes S :: "'a :: euclidean_space set" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14203 |
assumes "affine S" "a \<in> S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14204 |
shows "aff_dim S = dim {x - a |x. x \<in> S}"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14205 |
proof - |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14206 |
obtain B where aff: "affine hull B = affine hull S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14207 |
and ind: "\<not> affine_dependent B" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14208 |
and card: "of_nat (card B) = aff_dim S + 1" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14209 |
using aff_dim_basis_exists by blast |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14210 |
then have "B \<noteq> {}" using assms
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14211 |
by (metis affine_hull_eq_empty ex_in_conv) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14212 |
then obtain c where "c \<in> B" by auto |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14213 |
then have "c \<in> S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14214 |
by (metis aff affine_hull_eq \<open>affine S\<close> hull_inc) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14215 |
have xy: "x - c = y - a \<longleftrightarrow> y = x + 1 *\<^sub>R (a - c)" for x y c and a::'a |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14216 |
by (auto simp: algebra_simps) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14217 |
have *: "{x - c |x. x \<in> S} = {x - a |x. x \<in> S}"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14218 |
apply safe |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14219 |
apply (simp_all only: xy) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14220 |
using mem_affine_3_minus [OF \<open>affine S\<close>] \<open>a \<in> S\<close> \<open>c \<in> S\<close> apply blast+ |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14221 |
done |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14222 |
have affS: "affine hull S = S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14223 |
by (simp add: \<open>affine S\<close>) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14224 |
have "aff_dim S = of_nat (card B) - 1" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14225 |
using card by simp |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14226 |
also have "... = dim {x - c |x. x \<in> B}"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14227 |
by (simp add: affine_independent_card_dim_diffs [OF ind \<open>c \<in> B\<close>]) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14228 |
also have "... = dim {x - c | x. x \<in> affine hull B}"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14229 |
by (simp add: diffs_affine_hull_span \<open>c \<in> B\<close>) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14230 |
also have "... = dim {x - a |x. x \<in> S}"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14231 |
by (simp add: affS aff *) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14232 |
finally show ?thesis . |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14233 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14234 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14235 |
lemma aff_dim_linear_image_le: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14236 |
assumes "linear f" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14237 |
shows "aff_dim(f ` S) \<le> aff_dim S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14238 |
proof - |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14239 |
have "aff_dim (f ` T) \<le> aff_dim T" if "affine T" for T |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14240 |
proof (cases "T = {}")
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14241 |
case True then show ?thesis by (simp add: aff_dim_geq) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14242 |
next |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14243 |
case False |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14244 |
then obtain a where "a \<in> T" by auto |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14245 |
have 1: "((\<lambda>x. x - f a) ` f ` T) = {x - f a |x. x \<in> f ` T}"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14246 |
by auto |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14247 |
have 2: "{x - f a| x. x \<in> f ` T} = f ` {x - a| x. x \<in> T}"
|
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
14248 |
by (force simp: linear_diff [OF assms]) |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14249 |
have "aff_dim (f ` T) = int (dim {x - f a |x. x \<in> f ` T})"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14250 |
by (simp add: \<open>a \<in> T\<close> hull_inc aff_dim_eq_dim [of "f a"] 1) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14251 |
also have "... = int (dim (f ` {x - a| x. x \<in> T}))"
|
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
14252 |
by (force simp: linear_diff [OF assms] 2) |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14253 |
also have "... \<le> int (dim {x - a| x. x \<in> T})"
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14254 |
by (simp add: dim_image_le [OF assms]) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14255 |
also have "... \<le> aff_dim T" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14256 |
by (simp add: aff_dim_dim_affine_diffs [symmetric] \<open>a \<in> T\<close> \<open>affine T\<close>) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14257 |
finally show ?thesis . |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14258 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14259 |
then |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14260 |
have "aff_dim (f ` (affine hull S)) \<le> aff_dim (affine hull S)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14261 |
using affine_affine_hull [of S] by blast |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14262 |
then show ?thesis |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14263 |
using affine_hull_linear_image assms linear_conv_bounded_linear by fastforce |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14264 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14265 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14266 |
lemma aff_dim_injective_linear_image [simp]: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14267 |
assumes "linear f" "inj f" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14268 |
shows "aff_dim (f ` S) = aff_dim S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14269 |
proof (rule antisym) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14270 |
show "aff_dim (f ` S) \<le> aff_dim S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14271 |
by (simp add: aff_dim_linear_image_le assms(1)) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14272 |
next |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14273 |
obtain g where "linear g" "g \<circ> f = id" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14274 |
using linear_injective_left_inverse assms by blast |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14275 |
then have "aff_dim S \<le> aff_dim(g ` f ` S)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14276 |
by (simp add: image_comp) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14277 |
also have "... \<le> aff_dim (f ` S)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14278 |
by (simp add: \<open>linear g\<close> aff_dim_linear_image_le) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14279 |
finally show "aff_dim S \<le> aff_dim (f ` S)" . |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14280 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14281 |
|
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
14282 |
|
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14283 |
text\<open>Choosing a subspace of a given dimension\<close> |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14284 |
proposition choose_subspace_of_subspace: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14285 |
fixes S :: "'n::euclidean_space set" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14286 |
assumes "n \<le> dim S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14287 |
obtains T where "subspace T" "T \<subseteq> span S" "dim T = n" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14288 |
proof - |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14289 |
have "\<exists>T. subspace T \<and> T \<subseteq> span S \<and> dim T = n" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14290 |
using assms |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14291 |
proof (induction n) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14292 |
case 0 then show ?case by force |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14293 |
next |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14294 |
case (Suc n) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14295 |
then obtain T where "subspace T" "T \<subseteq> span S" "dim T = n" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14296 |
by force |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14297 |
then show ?case |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14298 |
proof (cases "span S \<subseteq> span T") |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14299 |
case True |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14300 |
have "dim S = dim T" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14301 |
apply (rule span_eq_dim [OF subset_antisym [OF True]]) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14302 |
by (simp add: \<open>T \<subseteq> span S\<close> span_minimal subspace_span) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14303 |
then show ?thesis |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14304 |
using Suc.prems \<open>dim T = n\<close> by linarith |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14305 |
next |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14306 |
case False |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14307 |
then obtain y where y: "y \<in> S" "y \<notin> T" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14308 |
by (meson span_mono subsetI) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14309 |
then have "span (insert y T) \<subseteq> span S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14310 |
by (metis (no_types) \<open>T \<subseteq> span S\<close> subsetD insert_subset span_inc span_mono span_span) |
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
14311 |
with \<open>dim T = n\<close> \<open>subspace T\<close> y show ?thesis |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14312 |
apply (rule_tac x="span(insert y T)" in exI) |
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
14313 |
apply (auto simp: dim_insert) |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
14314 |
using span_eq by blast |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14315 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14316 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14317 |
with that show ?thesis by blast |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14318 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14319 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14320 |
lemma choose_affine_subset: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14321 |
assumes "affine S" "-1 \<le> d" and dle: "d \<le> aff_dim S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14322 |
obtains T where "affine T" "T \<subseteq> S" "aff_dim T = d" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14323 |
proof (cases "d = -1 \<or> S={}")
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14324 |
case True with assms show ?thesis |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14325 |
by (metis aff_dim_empty affine_empty bot.extremum that eq_iff) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14326 |
next |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14327 |
case False |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14328 |
with assms obtain a where "a \<in> S" "0 \<le> d" by auto |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14329 |
with assms have ss: "subspace (op + (- a) ` S)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14330 |
by (simp add: affine_diffs_subspace) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14331 |
have "nat d \<le> dim (op + (- a) ` S)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14332 |
by (metis aff_dim_subspace aff_dim_translation_eq dle nat_int nat_mono ss) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14333 |
then obtain T where "subspace T" and Tsb: "T \<subseteq> span (op + (- a) ` S)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14334 |
and Tdim: "dim T = nat d" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14335 |
using choose_subspace_of_subspace [of "nat d" "op + (- a) ` S"] by blast |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14336 |
then have "affine T" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14337 |
using subspace_affine by blast |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14338 |
then have "affine (op + a ` T)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14339 |
by (metis affine_hull_eq affine_hull_translation) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14340 |
moreover have "op + a ` T \<subseteq> S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14341 |
proof - |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14342 |
have "T \<subseteq> op + (- a) ` S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14343 |
by (metis (no_types) span_eq Tsb ss) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14344 |
then show "op + a ` T \<subseteq> S" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14345 |
using add_ac by auto |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14346 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14347 |
moreover have "aff_dim (op + a ` T) = d" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14348 |
by (simp add: aff_dim_subspace Tdim \<open>0 \<le> d\<close> \<open>subspace T\<close> aff_dim_translation_eq) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14349 |
ultimately show ?thesis |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14350 |
by (rule that) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14351 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63092
diff
changeset
|
14352 |
|
| 63301 | 14353 |
subsection\<open>Several Variants of Paracompactness\<close> |
14354 |
||
14355 |
proposition paracompact: |
|
14356 |
fixes S :: "'a :: euclidean_space set" |
|
14357 |
assumes "S \<subseteq> \<Union>\<C>" and opC: "\<And>T. T \<in> \<C> \<Longrightarrow> open T" |
|
14358 |
obtains \<C>' where "S \<subseteq> \<Union> \<C>'" |
|
14359 |
and "\<And>U. U \<in> \<C>' \<Longrightarrow> open U \<and> (\<exists>T. T \<in> \<C> \<and> U \<subseteq> T)" |
|
14360 |
and "\<And>x. x \<in> S |
|
14361 |
\<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> |
|
14362 |
finite {U. U \<in> \<C>' \<and> (U \<inter> V \<noteq> {})}"
|
|
14363 |
proof (cases "S = {}")
|
|
14364 |
case True with that show ?thesis by blast |
|
14365 |
next |
|
14366 |
case False |
|
14367 |
have "\<exists>T U. x \<in> U \<and> open U \<and> closure U \<subseteq> T \<and> T \<in> \<C>" if "x \<in> S" for x |
|
14368 |
proof - |
|
14369 |
obtain T where "x \<in> T" "T \<in> \<C>" "open T" |
|
14370 |
using assms \<open>x \<in> S\<close> by blast |
|
14371 |
then obtain e where "e > 0" "cball x e \<subseteq> T" |
|
14372 |
by (force simp: open_contains_cball) |
|
14373 |
then show ?thesis |
|
14374 |
apply (rule_tac x = T in exI) |
|
14375 |
apply (rule_tac x = "ball x e" in exI) |
|
14376 |
using \<open>T \<in> \<C>\<close> |
|
14377 |
apply (simp add: closure_minimal) |
|
14378 |
done |
|
14379 |
qed |
|
14380 |
then obtain F G where Gin: "x \<in> G x" and oG: "open (G x)" |
|
14381 |
and clos: "closure (G x) \<subseteq> F x" and Fin: "F x \<in> \<C>" |
|
14382 |
if "x \<in> S" for x |
|
14383 |
by metis |
|
14384 |
then obtain \<F> where "\<F> \<subseteq> G ` S" "countable \<F>" "\<Union>\<F> = UNION S G" |
|
14385 |
using Lindelof [of "G ` S"] by (metis image_iff) |
|
14386 |
then obtain K where K: "K \<subseteq> S" "countable K" and eq: "UNION K G = UNION S G" |
|
14387 |
by (metis countable_subset_image) |
|
14388 |
with False Gin have "K \<noteq> {}" by force
|
|
14389 |
then obtain a :: "nat \<Rightarrow> 'a" where "range a = K" |
|
14390 |
by (metis range_from_nat_into \<open>countable K\<close>) |
|
14391 |
then have odif: "\<And>n. open (F (a n) - \<Union>{closure (G (a m)) |m. m < n})"
|
|
14392 |
using \<open>K \<subseteq> S\<close> Fin opC by (fastforce simp add:) |
|
14393 |
let ?C = "range (\<lambda>n. F(a n) - \<Union>{closure(G(a m)) |m. m < n})"
|
|
14394 |
have enum_S: "\<exists>n. x \<in> F(a n) \<and> x \<in> G(a n)" if "x \<in> S" for x |
|
14395 |
proof - |
|
14396 |
have "\<exists>y \<in> K. x \<in> G y" using eq that Gin by fastforce |
|
14397 |
then show ?thesis |
|
14398 |
using clos K \<open>range a = K\<close> closure_subset by blast |
|
14399 |
qed |
|
14400 |
have 1: "S \<subseteq> Union ?C" |
|
14401 |
proof |
|
14402 |
fix x assume "x \<in> S" |
|
14403 |
define n where "n \<equiv> LEAST n. x \<in> F(a n)" |
|
14404 |
have n: "x \<in> F(a n)" |
|
14405 |
using enum_S [OF \<open>x \<in> S\<close>] by (force simp: n_def intro: LeastI) |
|
14406 |
have notn: "x \<notin> F(a m)" if "m < n" for m |
|
14407 |
using that not_less_Least by (force simp: n_def) |
|
14408 |
then have "x \<notin> \<Union>{closure (G (a m)) |m. m < n}"
|
|
14409 |
using n \<open>K \<subseteq> S\<close> \<open>range a = K\<close> clos notn by fastforce |
|
14410 |
with n show "x \<in> Union ?C" |
|
14411 |
by blast |
|
14412 |
qed |
|
14413 |
have 3: "\<exists>V. open V \<and> x \<in> V \<and> finite {U. U \<in> ?C \<and> (U \<inter> V \<noteq> {})}" if "x \<in> S" for x
|
|
14414 |
proof - |
|
14415 |
obtain n where n: "x \<in> F(a n)" "x \<in> G(a n)" |
|
14416 |
using \<open>x \<in> S\<close> enum_S by auto |
|
14417 |
have "{U \<in> ?C. U \<inter> G (a n) \<noteq> {}} \<subseteq> (\<lambda>n. F(a n) - \<Union>{closure(G(a m)) |m. m < n}) ` atMost n"
|
|
14418 |
proof clarsimp |
|
14419 |
fix k assume "(F (a k) - \<Union>{closure (G (a m)) |m. m < k}) \<inter> G (a n) \<noteq> {}"
|
|
14420 |
then have "k \<le> n" |
|
14421 |
by auto (metis closure_subset not_le subsetCE) |
|
14422 |
then show "F (a k) - \<Union>{closure (G (a m)) |m. m < k}
|
|
14423 |
\<in> (\<lambda>n. F (a n) - \<Union>{closure (G (a m)) |m. m < n}) ` {..n}"
|
|
14424 |
by force |
|
14425 |
qed |
|
14426 |
moreover have "finite ((\<lambda>n. F(a n) - \<Union>{closure(G(a m)) |m. m < n}) ` atMost n)"
|
|
14427 |
by force |
|
14428 |
ultimately have *: "finite {U \<in> ?C. U \<inter> G (a n) \<noteq> {}}"
|
|
14429 |
using finite_subset by blast |
|
14430 |
show ?thesis |
|
14431 |
apply (rule_tac x="G (a n)" in exI) |
|
14432 |
apply (intro conjI oG n *) |
|
14433 |
using \<open>K \<subseteq> S\<close> \<open>range a = K\<close> apply blast |
|
14434 |
done |
|
14435 |
qed |
|
14436 |
show ?thesis |
|
14437 |
apply (rule that [OF 1 _ 3]) |
|
14438 |
using Fin \<open>K \<subseteq> S\<close> \<open>range a = K\<close> apply (auto simp: odif) |
|
14439 |
done |
|
14440 |
qed |
|
14441 |
||
14442 |
corollary paracompact_closedin: |
|
14443 |
fixes S :: "'a :: euclidean_space set" |
|
14444 |
assumes cin: "closedin (subtopology euclidean U) S" |
|
14445 |
and oin: "\<And>T. T \<in> \<C> \<Longrightarrow> openin (subtopology euclidean U) T" |
|
14446 |
and "S \<subseteq> \<Union>\<C>" |
|
14447 |
obtains \<C>' where "S \<subseteq> \<Union> \<C>'" |
|
14448 |
and "\<And>V. V \<in> \<C>' \<Longrightarrow> openin (subtopology euclidean U) V \<and> (\<exists>T. T \<in> \<C> \<and> V \<subseteq> T)" |
|
14449 |
and "\<And>x. x \<in> U |
|
14450 |
\<Longrightarrow> \<exists>V. openin (subtopology euclidean U) V \<and> x \<in> V \<and> |
|
14451 |
finite {X. X \<in> \<C>' \<and> (X \<inter> V \<noteq> {})}"
|
|
14452 |
proof - |
|
14453 |
have "\<exists>Z. open Z \<and> (T = U \<inter> Z)" if "T \<in> \<C>" for T |
|
14454 |
using oin [OF that] by (auto simp: openin_open) |
|
14455 |
then obtain F where opF: "open (F T)" and intF: "U \<inter> F T = T" if "T \<in> \<C>" for T |
|
14456 |
by metis |
|
14457 |
obtain K where K: "closed K" "U \<inter> K = S" |
|
14458 |
using cin by (auto simp: closedin_closed) |
|
14459 |
have 1: "U \<subseteq> \<Union>insert (- K) (F ` \<C>)" |
|
14460 |
by clarsimp (metis Int_iff Union_iff \<open>U \<inter> K = S\<close> \<open>S \<subseteq> \<Union>\<C>\<close> subsetD intF) |
|
14461 |
have 2: "\<And>T. T \<in> insert (- K) (F ` \<C>) \<Longrightarrow> open T" |
|
14462 |
using \<open>closed K\<close> by (auto simp: opF) |
|
14463 |
obtain \<D> where "U \<subseteq> \<Union>\<D>" |
|
14464 |
and D1: "\<And>U. U \<in> \<D> \<Longrightarrow> open U \<and> (\<exists>T. T \<in> insert (- K) (F ` \<C>) \<and> U \<subseteq> T)" |
|
14465 |
and D2: "\<And>x. x \<in> U \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<D>. U \<inter> V \<noteq> {}}"
|
|
14466 |
using paracompact [OF 1 2] by auto |
|
14467 |
let ?C = "{U \<inter> V |V. V \<in> \<D> \<and> (V \<inter> K \<noteq> {})}"
|
|
14468 |
show ?thesis |
|
14469 |
proof (rule_tac \<C>' = "{U \<inter> V |V. V \<in> \<D> \<and> (V \<inter> K \<noteq> {})}" in that)
|
|
14470 |
show "S \<subseteq> \<Union>?C" |
|
14471 |
using \<open>U \<inter> K = S\<close> \<open>U \<subseteq> \<Union>\<D>\<close> K by (blast dest!: subsetD) |
|
14472 |
show "\<And>V. V \<in> ?C \<Longrightarrow> openin (subtopology euclidean U) V \<and> (\<exists>T. T \<in> \<C> \<and> V \<subseteq> T)" |
|
14473 |
using D1 intF by fastforce |
|
14474 |
have *: "{X. (\<exists>V. X = U \<inter> V \<and> V \<in> \<D> \<and> V \<inter> K \<noteq> {}) \<and> X \<inter> (U \<inter> V) \<noteq> {}} \<subseteq>
|
|
14475 |
(\<lambda>x. U \<inter> x) ` {U \<in> \<D>. U \<inter> V \<noteq> {}}" for V
|
|
14476 |
by blast |
|
14477 |
show "\<exists>V. openin (subtopology euclidean U) V \<and> x \<in> V \<and> finite {X \<in> ?C. X \<inter> V \<noteq> {}}"
|
|
14478 |
if "x \<in> U" for x |
|
14479 |
using D2 [OF that] |
|
14480 |
apply clarify |
|
14481 |
apply (rule_tac x="U \<inter> V" in exI) |
|
14482 |
apply (auto intro: that finite_subset [OF *]) |
|
14483 |
done |
|
14484 |
qed |
|
14485 |
qed |
|
14486 |
||
14487 |
corollary paracompact_closed: |
|
14488 |
fixes S :: "'a :: euclidean_space set" |
|
14489 |
assumes "closed S" |
|
14490 |
and opC: "\<And>T. T \<in> \<C> \<Longrightarrow> open T" |
|
14491 |
and "S \<subseteq> \<Union>\<C>" |
|
14492 |
obtains \<C>' where "S \<subseteq> \<Union>\<C>'" |
|
14493 |
and "\<And>U. U \<in> \<C>' \<Longrightarrow> open U \<and> (\<exists>T. T \<in> \<C> \<and> U \<subseteq> T)" |
|
14494 |
and "\<And>x. \<exists>V. open V \<and> x \<in> V \<and> |
|
14495 |
finite {U. U \<in> \<C>' \<and> (U \<inter> V \<noteq> {})}"
|
|
14496 |
using paracompact_closedin [of UNIV S \<C>] assms |
|
14497 |
by (auto simp: open_openin [symmetric] closed_closedin [symmetric]) |
|
14498 |
||
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14499 |
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14500 |
subsection\<open>Closed-graph characterization of continuity\<close> |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14501 |
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14502 |
lemma continuous_closed_graph_gen: |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14503 |
fixes T :: "'b::real_normed_vector set" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14504 |
assumes contf: "continuous_on S f" and fim: "f ` S \<subseteq> T" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14505 |
shows "closedin (subtopology euclidean (S \<times> T)) ((\<lambda>x. Pair x (f x)) ` S)" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14506 |
proof - |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14507 |
have eq: "((\<lambda>x. Pair x (f x)) ` S) = {z. z \<in> S \<times> T \<and> (f \<circ> fst)z - snd z \<in> {0}}"
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14508 |
using fim by auto |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14509 |
show ?thesis |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14510 |
apply (subst eq) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14511 |
apply (intro continuous_intros continuous_closedin_preimage continuous_on_subset [OF contf]) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14512 |
by auto |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14513 |
qed |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14514 |
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14515 |
lemma continuous_closed_graph_eq: |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14516 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14517 |
assumes "compact T" and fim: "f ` S \<subseteq> T" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14518 |
shows "continuous_on S f \<longleftrightarrow> |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14519 |
closedin (subtopology euclidean (S \<times> T)) ((\<lambda>x. Pair x (f x)) ` S)" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14520 |
(is "?lhs = ?rhs") |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14521 |
proof - |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14522 |
have "?lhs" if ?rhs |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14523 |
proof (clarsimp simp add: continuous_on_closed_gen [OF fim]) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14524 |
fix U |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14525 |
assume U: "closedin (subtopology euclidean T) U" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14526 |
have eq: "{x. x \<in> S \<and> f x \<in> U} = fst ` (((\<lambda>x. Pair x (f x)) ` S) \<inter> (S \<times> U))"
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14527 |
by (force simp: image_iff) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14528 |
show "closedin (subtopology euclidean S) {x \<in> S. f x \<in> U}"
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14529 |
by (simp add: U closedin_Int closedin_Times closed_map_fst [OF \<open>compact T\<close>] that eq) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14530 |
qed |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14531 |
with continuous_closed_graph_gen assms show ?thesis by blast |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14532 |
qed |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14533 |
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14534 |
lemma continuous_closed_graph: |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14535 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14536 |
assumes "closed S" and contf: "continuous_on S f" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14537 |
shows "closed ((\<lambda>x. Pair x (f x)) ` S)" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14538 |
apply (rule closedin_closed_trans) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14539 |
apply (rule continuous_closed_graph_gen [OF contf subset_UNIV]) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14540 |
by (simp add: \<open>closed S\<close> closed_Times) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14541 |
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14542 |
lemma continuous_from_closed_graph: |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14543 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14544 |
assumes "compact T" and fim: "f ` S \<subseteq> T" and clo: "closed ((\<lambda>x. Pair x (f x)) ` S)" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14545 |
shows "continuous_on S f" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14546 |
using fim clo |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14547 |
by (auto intro: closed_subset simp: continuous_closed_graph_eq [OF \<open>compact T\<close> fim]) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14548 |
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14549 |
lemma continuous_on_Un_local_open: |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14550 |
assumes opS: "openin (subtopology euclidean (S \<union> T)) S" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14551 |
and opT: "openin (subtopology euclidean (S \<union> T)) T" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14552 |
and contf: "continuous_on S f" and contg: "continuous_on T f" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14553 |
shows "continuous_on (S \<union> T) f" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14554 |
using pasting_lemma [of "{S,T}" "S \<union> T" "\<lambda>i. i" "\<lambda>i. f" f] contf contg opS opT by blast
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14555 |
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14556 |
lemma continuous_on_cases_local_open: |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14557 |
assumes opS: "openin (subtopology euclidean (S \<union> T)) S" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14558 |
and opT: "openin (subtopology euclidean (S \<union> T)) T" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14559 |
and contf: "continuous_on S f" and contg: "continuous_on T g" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14560 |
and fg: "\<And>x. x \<in> S \<and> ~P x \<or> x \<in> T \<and> P x \<Longrightarrow> f x = g x" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14561 |
shows "continuous_on (S \<union> T) (\<lambda>x. if P x then f x else g x)" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14562 |
proof - |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14563 |
have "\<And>x. x \<in> S \<Longrightarrow> (if P x then f x else g x) = f x" "\<And>x. x \<in> T \<Longrightarrow> (if P x then f x else g x) = g x" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14564 |
by (simp_all add: fg) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14565 |
then have "continuous_on S (\<lambda>x. if P x then f x else g x)" "continuous_on T (\<lambda>x. if P x then f x else g x)" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14566 |
by (simp_all add: contf contg cong: continuous_on_cong) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14567 |
then show ?thesis |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14568 |
by (rule continuous_on_Un_local_open [OF opS opT]) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14569 |
qed |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14570 |
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14571 |
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
14572 |
|
| 33175 | 14573 |
end |